lean2/tests/lean/run/calc.lean
Leonardo de Moura cd17618f4a refactor(library): replace 'calc_trans', 'calc_symm', 'calc_refl' and 'calc_subst' commands with attributes '[symm]', '[refl]', '[trans]' and '[subst]'
These attributes are used by the calc command.
They will also be used by tactics such as 'reflexivity', 'symmetry' and
'transitivity'.

See issue #500
2015-05-02 15:15:35 -07:00

17 lines
379 B
Text

import data.num
namespace foo
constant le : num → num → Prop
axiom le_trans {a b c : num} : le a b → le b c → le a c
attribute le_trans [trans]
infix `≤` := le
end foo
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 foo