From fb0a4ba033da3ea305e0add8c2ec4573e95fa3d2 Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Thu, 18 Mar 2010 11:35:44 +0000 Subject: [PATCH] Nuova versione di not. --- .../matita/nlibrary/arithmetics/nat.ma | 198 ++++++++++-------- 1 file changed, 109 insertions(+), 89 deletions(-) diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 25745e370..fe9ef8267 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -15,7 +15,11 @@ (* include "higher_order_defs/functions.ma". *) include "hints_declaration.ma". include "basics/functions.ma". -include "basics/eq.ma". +include "basics/eq.ma". + +ntheorem foo: ∀A:Type.∀a,b:A.∀f:A→A.∀g:A→A→A. +(∀x,y.f (g x y) = x) → ∀x. g (f a) x = b → f a = f b. +//; nqed. ninductive nat : Type[0] ≝ | O : nat @@ -49,18 +53,18 @@ ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. //. nqed. *) ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2/; nqed. +/3/; nqed. ndefinition not_zero: nat → Prop ≝ λn: nat. match n with [ O ⇒ False | (S p) ⇒ True ]. ntheorem not_eq_O_S : ∀n:nat. O ≠ S n. -#n; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. +#n; napply nmk; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. nqed. -ntheorem not_eq_n_Sn : ∀n:nat. n ≠ S n. -#n; nelim n; /2/; nqed. +ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n. +#n; nelim n;/2/; nqed. ntheorem nat_case: ∀n:nat.∀P:nat → Prop. @@ -78,7 +82,7 @@ ntheorem nat_elim2 : ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). napply nat_elim2; #n; - ##[ ncases n; /3/; + ##[ ncases n; /2/; ##| /3/; ##| #m; #Hind; ncases Hind; /3/; ##] @@ -112,9 +116,10 @@ ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m. #n; nelim n; nnormalize; //; nqed. *) -(* deleterio -ntheorem plus_n_SO : ∀n:nat. S n = n+S O. -//; nqed. *) +(* deleterio? +ntheorem plus_n_1 : ∀n:nat. S n = n+1. +//; nqed. +*) ntheorem symmetric_plus: symmetric ? plus. #n; nelim n; nnormalize; //; nqed. @@ -159,7 +164,7 @@ ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). #n; nelim n; nnormalize; //; nqed. ntheorem symmetric_times : symmetric nat times. -#n; nelim n; nnormalize; //; nqed. +#n; nelim n; nnormalize; //; nqed. (* variant sym_times : \forall n,m:nat. n*m = m*n \def symmetric_times. *) @@ -167,9 +172,9 @@ symmetric_times. *) ntheorem distributive_times_plus : distributive nat times plus. #n; nelim n; nnormalize; //; nqed. -ntheorem distributive_times_plus_r: -\forall a,b,c:nat. (b+c)*a = b*a + c*a. -//; nqed. +ntheorem distributive_times_plus_r : + ∀a,b,c:nat. (b+c)*a = b*a + c*a. +//; nqed. ntheorem associative_times: associative nat times. #n; nelim n; nnormalize; //; nqed. @@ -223,13 +228,13 @@ interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). (* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m). //; nqed. *) -ndefinition ge: nat \to nat \to Prop \def -\lambda n,m:nat.m \leq n. +ndefinition ge: nat → nat → Prop ≝ +λn,m:nat.m ≤ n. interpretation "natural 'greater or equal to'" 'geq x y = (ge x y). -ndefinition gt: nat \to nat \to Prop \def -\lambda n,m:nat.m (symmetric_times c n); -nrewrite > (symmetric_times c m); -napply monotonic_lt_times_l;//; -nqed. +/2/; nqed. ntheorem lt_to_le_to_lt_times: ∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q. @@ -756,19 +773,17 @@ napply lt_to_le_to_lt_times;/2/; nqed. ntheorem lt_times_n_to_lt_l: -∀n,p,q:nat. O < n → p*n < q*n → p < q. -#n; #p; #q; #posn; #Hlt; +∀n,p,q:nat. p*n < q*n → p < q. +#n; #p; #q; #Hlt; nelim (decidable_lt p q);//; -#nltpq;napply False_ind; -napply (lt_to_not_le ? ? Hlt); -napply monotonic_le_times_l;/3/; +#nltpq; napply False_ind; +napply (absurd ? ? (lt_to_not_le ? ? Hlt)); +napplyS monotonic_le_times_r;/2/; nqed. ntheorem lt_times_n_to_lt_r: -∀n,p,q:nat. O < n → n*p < n*q → p < q. -#n; #p; #q; #posn; #Hlt; -napply (lt_times_n_to_lt_l ??? posn);//; -nqed. +∀n,p,q:nat. n*p < n*q → p < q. +/2/; nqed. (* theorem nat_compare_times_l : \forall n,p,q:nat. @@ -839,14 +854,14 @@ ntheorem minus_n_n: ∀n:nat.O=n-n. #n; nelim n; //; nqed. ntheorem minus_Sn_n: ∀n:nat. S O = (S n)-n. -#n; nelim n; //; nqed. +#n; nelim n; nnormalize; //; nqed. ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m). (* qualcosa da capire qui #n; #m; #lenm; nelim lenm; napplyS refl_eq. *) napply nat_elim2; ##[// - ##|#n; #abs; napply False_ind; /2/. + ##|#n; #abs; napply False_ind; /2/ ##|#n; #m; #Hind; #c; napplyS Hind; /2/; ##] nqed. @@ -859,7 +874,7 @@ napplyS (not_eq_to_le_to_lt n (S m) H H1); nqed. ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m). -napply nat_elim2; //; nqed. +napply nat_elim2; nnormalize; //; nqed. ntheorem plus_minus: ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m. @@ -874,7 +889,7 @@ ntheorem minus_plus_m_m: ∀n,m:nat.n = (n+m)-m. #n; #m; napplyS (plus_minus m m n); //; nqed. ntheorem plus_minus_m_m: ∀n,m:nat. -m \leq n \to n = (n-m)+m. + m ≤ n → n = (n-m)+m. #n; #m; #lemn; napplyS symmetric_eq; napplyS (plus_minus m n m); //; nqed. @@ -882,7 +897,7 @@ ntheorem le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m. #n; nelim n; ##[// ##|#a; #Hind; #m; ncases m;//; - nnormalize; #n;napplyS le_S_S;// + nnormalize; #n;/2/; ##] nqed. @@ -901,7 +916,7 @@ nqed. ntheorem minus_pred_pred : ∀n,m:nat. O < n → O < m → pred n - pred m = n - m. #n; #m; #posn; #posm; -napply (lt_O_n_elim n posn); +napply (lt_O_n_elim n posn); napply (lt_O_n_elim m posm);//. nqed. @@ -991,6 +1006,7 @@ ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. #n; #m; #p; #lep; (* bello *) napplyS monotonic_le_minus_l;//; +(* /2/; *) nqed. ntheorem monotonic_le_minus_r: @@ -1033,14 +1049,15 @@ simplify.unfold Not.intro.apply H1.apply inj_S.assumption. qed. *) -ntheorem eqb_elim : ∀ n,m:nat.∀ P:bool → Prop. +naxiom eqb_elim : ∀ n,m:nat.∀ P:bool → Prop. (n=m → (P true)) → (n ≠ m → (P false)) → (P (eqb n m)). +(* napply nat_elim2; ##[#n; ncases n; nnormalize; /3/; - ##|nnormalize; /3/; + ##|nnormalize; /3/; ##|nnormalize; /4/; ##] -nqed. +nqed.*) ntheorem eqb_n_n: ∀n. eqb n n = true. #n; nelim n; nnormalize; //. @@ -1062,8 +1079,8 @@ ntheorem eq_to_eqb_true: ∀n,m:nat. ntheorem not_eq_to_eqb_false: ∀n,m:nat. n ≠ m → eqb n m = false. #n; #m; #noteq; -nelim (true_or_false (eqb n m)); //; -#Heq; napply False_ind; napply noteq;/2/; +napply eqb_elim;//; +#Heq; napply False_ind; /2/; nqed. nlet rec leb n m ≝ @@ -1078,10 +1095,10 @@ ntheorem leb_elim: ∀n,m:nat. ∀P:bool → Prop. (n ≤ m → P true) → (n ≰ m → P false) → P (leb n m). napply nat_elim2; nnormalize; ##[/2/ - ##| /3/; + ##|/3/; ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind; ##[#lenm; napply Pt; napply le_S_S;//; - ##|#nlenm; napply Pf; #leSS; /3/; + ##|#nlenm; napply Pf; /2/; ##] ##] nqed. @@ -1097,7 +1114,7 @@ ntheorem leb_false_to_not_le:∀n,m. leb n m = false → n ≰ m. #n; #m; napply leb_elim; ##[#_; #abs; napply False_ind; /2/; - ##|/2/; + ##|//; ##] nqed. @@ -1106,11 +1123,14 @@ ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true. #H; #H1; napply False_ind; /2/; nqed. -ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false. +ntheorem not_le_to_leb_false: ∀n,m. n ≰ m → leb n m = false. #n; #m; napply leb_elim; //; #H; #H1; napply False_ind; /2/; nqed. +ntheorem lt_to_leb_false: ∀n,m. m < n → leb n m = false. +/3/; nqed. + (* serve anche ltb? ndefinition ltb ≝λn,m. leb (S n) m. -- 2.39.2