lean2/tests/lean/run/intros.lean
Leonardo de Moura e95c7c5f70 refactor(library/tactic/rename_tactic): use new 'tactic.expr' to implement 'intro/intros' tactic
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-10-22 17:29:50 -07:00

33 lines
593 B
Text

import logic tools.tactic
open tactic
theorem tst1 (a b : Prop) : a → b → b :=
by intro Ha; intro Hb; apply Hb
theorem tst2 (a b : Prop) : a → b → a ∧ b :=
by intro Ha; intro Hb; apply and.intro; apply Hb; apply Ha
theorem tst3 (a b : Prop) : a → b → a ∧ b :=
begin
intro Ha,
intro Hb,
apply and.intro,
apply Hb,
apply Ha
end
theorem tst4 (a b : Prop) : a → b → a ∧ b :=
begin
intros (Ha, Hb),
apply and.intro,
apply Hb,
apply Ha
end
theorem tst5 (a b : Prop) : a → b → a ∧ b :=
begin
intros,
apply and.intro,
eassumption,
eassumption
end