lean2/tests/lean/run/unfold_test.lean
Leonardo de Moura f8d472c9f1 feat(frontends/lean/parse_rewrite_tactic): change the semantics of rewrite[↑f] when f is recursive
After this commit it behaves like 'unfold f'.
That is, it will unfold f even if it fails to fold recursive
applications. Now, only 'esimp[f]' will not unfold f-applications when
it cannot fold the recursive applications.

This commit also closes #692. It is part of a series of commits that
addresses this issue.

closes #692
2015-07-12 13:20:21 -04:00

39 lines
1.4 KiB
Text
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

import data.list data.vector
variables {A B : Type}
section
open list
theorem last_concat {x : A} : ∀ {l : list A} (h : concat x l ≠ []), last (concat x l) h = x
| [] h := rfl
| [a] h := rfl
| (a₁::a₂::l) h :=
by xrewrite [↑concat, ↑concat, last_cons_cons, ↓concat x (a₂::l), last_concat]
theorem reverse_append : ∀ (s t : list A), reverse (s ++ t) = (reverse t) ++ (reverse s)
| [] t2 :=
by esimp [append, reverse]; rewrite append_nil_right
| (a2 :: s2) t2 :=
by rewrite [↑append, ↑reverse, reverse_append, concat_eq_append, append.assoc, -concat_eq_append]
end
section
open vector nat prod
theorem unzip_zip : ∀ {n : nat} (v₁ : vector A n) (v₂ : vector B n), unzip (zip v₁ v₂) = (v₁, v₂)
| 0 [] [] := rfl
| (n+1) (a::va) (b::vb) := by rewrite [↑zip, ↑unzip, unzip_zip]
theorem zip_unzip : ∀ {n : nat} (v : vector (A × B) n), zip (pr₁ (unzip v)) (pr₂ (unzip v)) = v
| 0 [] := rfl
| (n+1) ((a, b) :: v) := by rewrite [↑unzip,↑zip,zip_unzip]
theorem reverse_concat : Π {n : nat} (xs : vector A n) (a : A), reverse (concat xs a) = a :: reverse xs
| 0 [] a := rfl
| (n+1) (x :: xs) a := by xrewrite [↑concat,↑reverse,reverse_concat]
theorem reverse_reverse : Π {n : nat} (xs : vector A n), reverse (reverse xs) = xs
| 0 [] := rfl
| (succ n) (x :: xs) := by rewrite [↑reverse at {1}, reverse_concat, reverse_reverse]
end