lean2/tests/lean/run/calc_heq_symm.lean
Leonardo de Moura 94cf10284a fix(frontends/lean/calc_proof_elaborator): bug when inserting symmetry proofs for heq, fixes #286
The problem was that heq type is
    Pi {A : Type} (a : A) {B : Type} (b : B), Prop

The calc_proof_elaborator was assuming that (a : A) (b : B) were the
last two arguments in any relation supported by calc.

The fix is to remove this assumption.
2014-11-01 07:30:04 -07:00

10 lines
292 B
Text

import logic
theorem tst {A B C D : Type} {a₁ a₂ : A} {b : B} {c : C} {d : D}
(H₀ : a₁ = a₂) (H₁ : a₂ == b) (H₂ : b == c) (H₃ : c == d) : d == a₁ :=
calc d == c : H₃
... == b : H₂
... == a₂ : H₁
... = a₁ : H₀
print definition tst