lean2/tests/lean/run/calc_auto_trans_eq.lean
Leonardo de Moura 4ec0e1b07c feat(frontends/lean): improve calc mode
Now, it automatically supports transitivity of the form
    (R a b) -> (b = c) -> R a c
    (a = b) -> (R b c) -> R a c

closes #507
2015-04-04 08:58:35 -07:00

18 lines
533 B
Text

import data.list
constant R : Π {A : Type}, A → A → Prop
infix `~`:50 := R
example {A : Type} {a b c d : list nat} (H₁ : a ~ b) (H₂ : b = c) (H₃ : c = d) : a ~ d :=
calc a ~ b : H₁
... = c : H₂
... = d : H₃
example {A : Type} {a b c d : list nat} (H₁ : a = b) (H₂ : b = c) (H₃ : c ~ d) : a ~ d :=
calc a = b : H₁
... = c : H₂
... ~ d : H₃
example {A : Type} {a b c d : list nat} (H₁ : a = b) (H₂ : b ~ c) (H₃ : c = d) : a ~ d :=
calc a = b : H₁
... ~ c : H₂
... = d : H₃