lean2/tests/lean/run/calc.lean
Leonardo de Moura 0f44e3c9f4 fix(frontends/lean): calc configuration commands, add check_constant_next auxiliary method
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-07-15 01:19:47 +01:00

17 lines
376 B
Text

import standard
using num
namespace foo
variable le : num → num → Bool
axiom le_trans {a b c : num} : le a b → le b c → le a c
calc_trans le_trans
infix `≤`:50 := le
end
namespace foo
theorem T {a b c d : num} : a ≤ b → b ≤ c → c ≤ d → a ≤ d
:= assume H1 H2 H3,
calc a ≤ b : H1
... ≤ c : H2
... ≤ d : H3
end