lean2/tests/lean/run/esimp1.lean
Leonardo de Moura a9515ac7a4 feat(library/tactic/rewrite_tactic): try to fold nested recursive applications after unfolding a recursive function
See issue #692.
The implementation still has some rough spots.
It is not clear what the right semantic is.
Moreover, the folds in e_closure could not be eliminated automatically.
2015-07-08 21:19:18 -04:00

30 lines
572 B
Text

open nat
example (x y : nat) (H : (fun (a : nat), sigma.pr1 ⟨a, y⟩) x = 0) : x = 0 :=
begin
esimp at H,
exact H
end
definition foo [irreducible] (a : nat) := a
example (x y : nat) (H : (fun (a : nat), sigma.pr1 ⟨foo a, y⟩) x = 0) : x = 0 :=
begin
esimp at H,
esimp ↑foo at H,
exact H
end
example (x y : nat) (H : x = 0) : (fun (a : nat), sigma.pr1 ⟨foo a, y⟩) x = 0 :=
begin
esimp,
esimp ↑foo,
exact H
end
example (x y : nat) (H : x = 0) : (fun (a : nat), sigma.pr1 ⟨foo a, y⟩) x = 0 :=
begin
esimp,
unfold foo,
exact H
end