fix(tests): to reflect recent changes in the standard library

This commit is contained in:
Leonardo de Moura 2015-07-06 15:05:01 -07:00
parent 0828ca775c
commit 7e0844a9e6
15 changed files with 28 additions and 13 deletions

View file

@ -753,6 +753,7 @@ of two even numbers is an even number.
import data.nat
open nat
namespace hide
definition even (a : nat) := ∃ b, a = 2*b
theorem EvenPlusEven {a b : nat} (H1 : even a) (H2 : even b) : even (a + b) :=
exists.elim H1 (fun (w1 : nat) (Hw1 : a = 2*w1),
@ -761,7 +762,7 @@ of two even numbers is an even number.
(calc a + b = 2*w1 + b : {Hw1}
... = 2*w1 + 2*w2 : {Hw2}
... = 2*(w1 + w2) : eq.symm !mul.left_distrib)))
end hide
#+END_SRC
The example above also uses [[./calc.org][calculational proofs]] to show that =a + b = 2*(w1 + w2)=.
@ -773,6 +774,8 @@ With this macro we can write the example above in a more natural way
#+BEGIN_SRC lean
import data.nat
open nat
namespace hide
definition even (a : nat) := ∃ b, a = 2*b
theorem EvenPlusEven {a b : nat} (H1 : even a) (H2 : even b) : even (a + b) :=
obtain (w1 : nat) (Hw1 : a = 2*w1), from H1,
@ -781,4 +784,5 @@ With this macro we can write the example above in a more natural way
(calc a + b = 2*w1 + b : {Hw1}
... = 2*w1 + 2*w2 : {Hw2}
... = 2*(w1 + w2) : eq.symm !mul.left_distrib)
end hide
#+END_SRC

View file

@ -1,4 +1,4 @@
import algebra.function
--
open function

View file

@ -6,6 +6,8 @@
bool.bor_tt|∀ (a : bool), eq (bool.bor a bool.tt) bool.tt
bool.band_tt|∀ (a : bool), eq (bool.band a bool.tt) a
bool.tt|bool
bool.eq_tt_of_bnot_eq_ff|eq (bool.bnot ?a) bool.ff → eq ?a bool.tt
bool.eq_ff_of_bnot_eq_tt|eq (bool.bnot ?a) bool.tt → eq ?a bool.ff
tactic.lettac|tactic.identifier → tactic.expr → tactic
bool.absurd_of_eq_ff_of_eq_tt|eq ?a bool.ff → eq ?a bool.tt → ?B
bool.eq_tt_of_ne_ff|ne ?a bool.ff → eq ?a bool.tt
@ -23,6 +25,8 @@ tt_bor|∀ (a : bool), eq (bor tt a) tt
tt_band|∀ (a : bool), eq (band tt a) a
bor_tt|∀ (a : bool), eq (bor a tt) tt
band_tt|∀ (a : bool), eq (band a tt) a
eq_tt_of_bnot_eq_ff|eq (bnot ?a) ff → eq ?a tt
eq_ff_of_bnot_eq_tt|eq (bnot ?a) tt → eq ?a ff
tactic.lettac|tactic.identifier → tactic.expr → tactic
absurd_of_eq_ff_of_eq_tt|eq ?a ff → eq ?a tt → ?B
eq_tt_of_ne_ff|ne ?a ff → eq ?a tt

View file

@ -3,7 +3,6 @@
--- Released under Apache 2.0 license as described in the file LICENSE.
--- Author: Jeremy Avigad
----------------------------------------------------------------------------------------------------
import algebra.function
open function
namespace congr

View file

@ -4,7 +4,7 @@
--- Author: Jeremy Avigad
----------------------------------------------------------------------------------------------------
import logic algebra.function
import logic
open eq
open function

View file

@ -1,4 +1,4 @@
import logic algebra.function
import logic
open function bool

View file

@ -9,7 +9,7 @@
import logic.eq
import data.unit data.sigma data.prod
import algebra.function algebra.binary
import algebra.binary
open eq

View file

@ -9,7 +9,7 @@
import logic.eq
import data.unit data.sigma data.prod
import algebra.function algebra.binary
import algebra.binary
open eq

View file

@ -9,7 +9,7 @@
import logic.eq
import data.unit data.sigma data.prod
import algebra.function algebra.binary
import algebra.binary
open eq

View file

@ -1,4 +1,4 @@
import data.fintype algebra.function
import data.fintype
open function
section test

View file

@ -1,4 +1,3 @@
import algebra.function
open bool nat
open function

View file

@ -53,6 +53,7 @@ obtain x y [[pxy qxy] qyx], from ex,
open nat
namespace hidden
definition even (a : nat) := ∃ x, a = 2*x
example (a b : nat) (H₁ : even a) (H₂ : even b) : even (a+b) :=
@ -72,3 +73,4 @@ have aux : m * (c₁ - c₂) = n₂, from calc
... = n₁ + n₂ - n₁ : Hc₂
... = n₂ : add_sub_cancel_left,
dvd.intro aux
end hidden

View file

@ -1,6 +1,8 @@
import data.nat
open nat
namespace foo
inductive Parity : nat → Type :=
| even : ∀ n : nat, Parity (2 * n)
| odd : ∀ n : nat, Parity (2 * n + 1)
@ -14,12 +16,12 @@ definition parity : Π (n : nat), Parity n
have aux : Parity n, from parity n,
cases aux with [k, k],
begin
apply (odd k)
apply (Parity.odd k)
end,
begin
change (Parity (2*k + 2*1)),
rewrite -mul.left_distrib,
apply (even (k+1))
apply (Parity.even (k+1))
end
end
@ -30,3 +32,5 @@ match ⟨n, parity n⟩ with
| ⟨⌞2 * k⌟, even k⟩ := k
| ⟨⌞2 * k + 1⌟, odd k⟩ := k
end
end foo

View file

@ -11,6 +11,7 @@ end
example (x y : ) : (x + y) * (x + y) = x * x + y * x + x * y + y * y :=
by rewrite [*mul.left_distrib, *mul.right_distrib, -add.assoc]
namespace tst
definition even (a : nat) := ∃b, a = 2*b
theorem even_plus_even {a b : nat} (H1 : even a) (H2 : even b) : even (a + b) :=
@ -25,3 +26,5 @@ theorem T2 (a b c : nat) (H1 : a = b) (H2 : b = c + 1) : a ≠ 0 :=
calc
a = succ c : by rewrite [H1, H2, add_one]
... ≠ 0 : succ_ne_zero c
end tst

View file

@ -1,4 +1,4 @@
import algebra.function
--
open function