lean2/tests/lean/run/541a.lean
Leonardo de Moura c5fb3ec6d0 fix(library/definitional/equations): fixes #541
This commit allows recursive applications to have less or more arguments
than the equation left-hand-side.
We add two tests
   - 541a.lean  recursive call with more arguments
   - 542b.lean  recursive call with less arguments
2015-05-10 20:37:44 -07:00

25 lines
879 B
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.nat
open nat list eq.ops
theorem nat.le_of_eq {x y : } (H : x = y) : x ≤ y := H ▸ !le.refl
section
variable {Q : Type}
definition f : list Q → -- default if l is empty, else max l
| [] := 0
| (h :: t) := f t + 1
theorem f_foo : ∀{l : list Q}, ∀{q : Q}, q ∈ l → f l ≥ 1
| [] := take q, assume Hq, absurd Hq !not_mem_nil
| [h] := take q, assume Hq, nat.le_of_eq !rfl
| (h :: (h' :: t)) := take q, assume Hq,
have Hor : q = h q ∈ (h' :: t), from iff.mp !mem_cons_iff Hq,
have H : f (h' :: t) ≥ 1, from f_foo (mem_cons h' t),
have H1 : 1 + 1 ≤ f (h' :: t) + 1, from nat.add_le_add_right H 1,
calc
f (h :: h' :: t) = f (h' :: t) + 1 : rfl
... ≥ 1 + 1 : H1
... = 1 : sorry
end