lean2/tests/lean/run/match_tac4.lean
Leonardo de Moura 4e35afedcc feat(frontends/lean): rename 'wait' to 'reveal'
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2015-05-08 20:54:16 -07:00

35 lines
759 B
Text

notation `⟪`:max t:(foldr `,` (e r, and.intro e r)) `⟫`:0 := t
check ⟪ trivial, trivial, trivial ⟫
theorem tst (a b c d : Prop) : a ∧ b ∧ c ∧ d ↔ d ∧ c ∧ b ∧ a :=
begin
apply iff.intro,
begin
intro H,
match H with
| ⟪ H₁, H₂, H₃, H₄ ⟫ := ⟪ H₄, H₃, H₂, H₁ ⟫
end
end,
begin
intro H,
match H with
| ⟪ H₁, H₂, H₃, H₄ ⟫ :=
begin
repeat (apply and.intro | assumption)
end
end
end
end
reveal tst
print definition tst
theorem tst2 (a b c d : Prop) : a ∧ b ∧ c ∧ d ↔ d ∧ c ∧ b ∧ a :=
begin
apply iff.intro,
repeat (intro H; repeat (cases H with [H', H] | apply and.intro | assumption))
end
reveal tst2
print definition tst2