lean2/tests/lean/run/congr_tac.lean
Leonardo de Moura e379034b95 feat(library/tactic): improve 'assumption' tactic
- It uses the unifier in "conservative" mode
- It only affects the current goal

closes #570
2015-05-02 17:33:54 -07:00

34 lines
716 B
Text

import data.list
example (f : nat → nat → nat) (a b c : nat) : b = c → f a b = f a c :=
begin
intro bc,
congruence,
assumption
end
example (f g : nat → nat → nat) (a b c : nat) : f = g → b = c → f a b = g a c :=
begin
intro fg bc,
congruence,
exact fg,
exact bc
end
example (f g : nat → nat → nat) (a b c : nat) : f = g → b = c → f a b = g a c :=
by intros; congruence; repeat assumption
open list
example (a b : nat) : a = b → [a] ++ [b] = [b] ++ [a] :=
begin
intro ab,
congruence,
{congruence,
exact ab},
{congruence,
exact (eq.symm ab)}
end
example (a b : nat) : a = b → [a] ++ [b] = [b] ++ [a] :=
by intros; repeat (congruence | assumption | symmetry)