X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=25745e370c19d3ed3072e525630234461c6a5725;hb=edccb29109d07b54b48230a280f4351ed042dd9f;hp=d1e401ae5f87de195cfdc4edbdd3f30d532b9470;hpb=71690ba2c49a618d36e5eafcc16726e1ba80f038;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index d1e401ae5..25745e370 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -78,7 +78,7 @@ ntheorem nat_elim2 : ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). napply nat_elim2; #n; - ##[ ncases n; /2/; + ##[ ncases n; /3/; ##| /3/; ##| #m; #Hind; ncases Hind; /3/; ##] @@ -112,7 +112,7 @@ 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. *) @@ -220,6 +220,9 @@ interpretation "natural 'less than'" 'lt x y = (lt x y). 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. @@ -240,8 +243,9 @@ nqed. ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p \def transitive_le. *) -ntheorem transitive_lt: transitive nat lt. -#a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed. + +naxiom transitive_lt: transitive nat lt. +(* #a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.*) (* theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p @@ -259,13 +263,18 @@ ntheorem le_n_Sn : ∀n:nat. n ≤ S n. ntheorem le_pred_n : ∀n:nat. pred n ≤ n. #n; nelim n; //; nqed. +(* XXX global problem *) +nlemma my_trans_le : ∀x,y,z:nat.x ≤ y → y ≤ z → x ≤ z. +napply transitive_le. +nqed. + ntheorem monotonic_pred: monotonic ? le pred. -#n; #m; #lenm; nelim lenm; //; /2/; nqed. +#n; #m; #lenm; nelim lenm; /2/; nqed. ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m. -/2/; nqed. +(* XXX *) nletin hint ≝ monotonic. /2/; nqed. -ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m. +ntheorem lt_S_S_to_lt: ∀n,m. S n < S m → n < m. /2/; nqed. ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. @@ -284,15 +293,17 @@ ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m. ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m. /3/; nqed. +naxiom decidable_le: ∀n,m. decidable (n≤m). +(* ntheorem decidable_le: ∀n,m. decidable (n≤m). -napply nat_elim2; #n; /2/; -#m; #dec; ncases dec;/3/; nqed. +napply nat_elim2; #n; /3/; +#m; #dec; ncases dec;/4/; nqed. *) ntheorem decidable_lt: ∀n,m. decidable (n < m). #n; #m; napply decidable_le ; nqed. ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n. -#n; nelim n; /2/; nqed. +#n; nelim n; /3/; nqed. ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m. /2/; nqed. @@ -301,7 +312,7 @@ ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n. napply nat_elim2; #n; ##[#abs; napply False_ind;/2/; ##|/2/; - ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/3/; + ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/4/; ##] nqed. @@ -549,7 +560,7 @@ ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 \to m1 ≤ m2 → n1 + m1 ≤ n2 + m2. -#n1; #n2; #m1; #m2; #len; #lem; napply transitive_le; +#n1; #n2; #m1; #m2; #len; #lem; napply (transitive_le ? (n1+m2)); /2/; nqed. ntheorem le_plus_n :∀n,m:nat. m ≤ n + m. @@ -562,15 +573,16 @@ ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n. //; nqed. ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m. -#a; nelim a; /3/; nqed. +#a; nelim a; nnormalize; /3/; nqed. ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m. /2/; nqed. (* plus & lt *) + ntheorem monotonic_lt_plus_r: ∀n:nat.monotonic nat lt (λm.n+m). -/2/; nqed. +/2/; nqed. (* variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def @@ -629,7 +641,7 @@ napply transitive_le; (* /2/ slow *) nqed. ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m. -(* bello *) +#n; #m; #H; napplyS monotonic_le_times_l; /2/; nqed. ntheorem le_times_to_le: @@ -645,9 +657,9 @@ ntheorem le_times_to_le: ##] nqed. -ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → n < 2*m. +ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m. #n; #m; #posm; #lenm; (* interessante *) -nnormalize; napplyS (le_plus n); //; nqed. +napplyS (le_plus n); //; nqed. (* times & lt *) (* @@ -749,8 +761,7 @@ ntheorem lt_times_n_to_lt_l: nelim (decidable_lt p q);//; #nltpq;napply False_ind; napply (lt_to_not_le ? ? Hlt); -napply monotonic_le_times_l. -napply not_lt_to_le; //; +napply monotonic_le_times_l;/3/; nqed. ntheorem lt_times_n_to_lt_r: @@ -835,11 +846,18 @@ ntheorem minus_Sn_m: ∀m,n:nat. m ≤ n → S n -m = S (n-m). #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. +ntheorem not_eq_to_le_to_le_minus: + ∀n,m.n ≠ m → n ≤ m → n ≤ m - 1. +#n; #m; ncases m;//; #m; nnormalize; +#H; #H1; napply le_S_S_to_le; +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. @@ -847,7 +865,7 @@ ntheorem plus_minus: ∀m,n,p:nat. m ≤ n → (n-m)+p = (n+p)-m. napply nat_elim2; ##[// - ##|#n; #p; #abs; napply False_ind;/2/; + ##|#n; #p; #abs; napply False_ind; /2/; ##|nnormalize;/3/; ##] nqed. @@ -1013,43 +1031,28 @@ elim ((eqb n1 m1)). simplify.apply eq_f.apply H1. simplify.unfold Not.intro.apply H1.apply inj_S.assumption. qed. - -theorem eqb_elim : \forall n,m:nat.\forall P:bool \to Prop. -(n=m \to (P true)) \to (n \neq m \to (P false)) \to (P (eqb n m)). -intros. -cut -(match (eqb n m) with -[ true \Rightarrow n = m -| false \Rightarrow n \neq m] \to (P (eqb n m))). -apply Hcut.apply eqb_to_Prop. -elim (eqb n m). -apply ((H H2)). -apply ((H1 H2)). -qed. - *) +ntheorem 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; /4/; + ##] +nqed. + ntheorem eqb_n_n: ∀n. eqb n n = true. #n; nelim n; nnormalize; //. nqed. ntheorem eqb_true_to_eq: ∀n,m:nat. eqb n m = true → n = m. -napply nat_elim2; - ##[#n; ncases n; nnormalize; //; - #m; #abs; napply False_ind;/2/; - ##|nnormalize; #m; #abs; napply False_ind;/2/; - ##|nnormalize; - #n; #m; #Hind; #eqnm; napply eq_f; napply Hind; //; - ##] +#n; #m; napply (eqb_elim n m);//; +#_; #abs; napply False_ind; /2/; nqed. -ntheorem eqb_false_to_not_eq: ∀n,m:nat. - eqb n m = false → n ≠ m. -napply nat_elim2; - ##[#n; ncases n; nnormalize; /2/; - ##|/2/; - ##|nnormalize;/2/; - ##] +ntheorem eqb_false_to_not_eq: ∀n,m:nat. eqb n m = false → n ≠ m. +#n; #m; napply (eqb_elim n m);/2/; nqed. ntheorem eq_to_eqb_true: ∀n,m:nat. @@ -1075,7 +1078,7 @@ 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/; @@ -1094,7 +1097,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. @@ -1132,3 +1135,99 @@ nqed. ntheorem le_to_ltb_false: ∀n,m. m \le n → ltb n m = false. #n; #m; #Hltb; napply lt_to_leb_false; /2/; nqed. *) + +ninductive compare : Type[0] ≝ +| LT : compare +| EQ : compare +| GT : compare. + +ndefinition compare_invert: compare → compare ≝ + λc.match c with + [ LT ⇒ GT + | EQ ⇒ EQ + | GT ⇒ LT ]. + +nlet rec nat_compare n m: compare ≝ +match n with +[ O ⇒ match m with + [ O ⇒ EQ + | (S q) ⇒ LT ] +| S p ⇒ match m with + [ O ⇒ GT + | S q ⇒ nat_compare p q]]. + +ntheorem nat_compare_n_n: ∀n. nat_compare n n = EQ. +#n;nelim n +##[// +##|#m;#IH;nnormalize;//] +nqed. + +ntheorem nat_compare_S_S: ∀n,m:nat.nat_compare n m = nat_compare (S n) (S m). +//; +nqed. + +ntheorem nat_compare_pred_pred: + ∀n,m.O < n → O < m → nat_compare n m = nat_compare (pred n) (pred m). +#n;#m;#Hn;#Hm; +napply (lt_O_n_elim n Hn); +napply (lt_O_n_elim m Hm); +#p;#q;//; +nqed. + +ntheorem nat_compare_to_Prop: + ∀n,m.match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n = m + | GT ⇒ m < n ]. +#n;#m; +napply (nat_elim2 (λn,m.match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n = m + | GT ⇒ m < n ]) ?????) (* FIXME: don't want to put all these ?, especially when … does not work! *) +##[##1,2:#n1;ncases n1;//; +##|#n1;#m1;nnormalize;ncases (nat_compare n1 m1); + ##[##1,3:nnormalize;#IH;napply le_S_S;//; + ##|nnormalize;#IH;nrewrite > IH;//] +nqed. + +ntheorem nat_compare_n_m_m_n: + ∀n,m:nat.nat_compare n m = compare_invert (nat_compare m n). +#n;#m; +napply (nat_elim2 (λn,m. nat_compare n m = compare_invert (nat_compare m n))) +##[##1,2:#n1;ncases n1;//; +##|#n1;#m1;#IH;nnormalize;napply IH] +nqed. + +ntheorem nat_compare_elim : + ∀n,m. ∀P:compare → Prop. + (n < m → P LT) → (n=m → P EQ) → (m < n → P GT) → P (nat_compare n m). +#n;#m;#P;#Hlt;#Heq;#Hgt; +ncut (match (nat_compare n m) with + [ LT ⇒ n < m + | EQ ⇒ n=m + | GT ⇒ m < n] → + P (nat_compare n m)) +##[ncases (nat_compare n m); + ##[napply Hlt + ##|napply Heq + ##|napply Hgt] +##|#Hcut;napply Hcut;//; +nqed. + +ninductive cmp_cases (n,m:nat) : CProp[0] ≝ + | cmp_le : n ≤ m → cmp_cases n m + | cmp_gt : m < n → cmp_cases n m. + +ntheorem lt_to_le : ∀n,m:nat. n < m → n ≤ m. +#n;#m;#H;nelim H +##[// +##|/2/] +nqed. + +nlemma cmp_nat: ∀n,m.cmp_cases n m. +#n;#m; nlapply (nat_compare_to_Prop n m); +ncases (nat_compare n m);#H +##[@;napply lt_to_le;// +##|@;// +##|@2;//] +nqed.