X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=d4ba1135e241626f551b5bca4b55c0a43f7d01e5;hb=a4cd6a8a1d6a008518c12daca794b9e811c1dee5;hp=0418444e376e0ecd47ca07a83b8915b1aa8cbc57;hpb=221472ea1597505d12677f5742e388125a15e2b9;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 0418444e3..d4ba1135e 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -12,12 +12,11 @@ (* *) (**************************************************************************) -(* include "higher_order_defs/functions.ma". *) include "hints_declaration.ma". include "basics/functions.ma". -include "basics/eq.ma". +include "basics/eq.ma". -ninductive nat : Type[0] ≝ +ninductive nat : Type ≝ | O : nat | S : nat → nat. @@ -36,7 +35,7 @@ ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on *) ndefinition pred ≝ - λn. match n with [ O ⇒ O | (S p) ⇒ p]. + λn. match n with [ O ⇒ O | S p ⇒ p]. ntheorem pred_Sn : ∀n. n = pred (S n). //; nqed. @@ -49,18 +48,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. @@ -80,7 +79,7 @@ ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; + ##| #m; #Hind; ncases Hind;/3/; ##] nqed. @@ -112,9 +111,10 @@ ntheorem plus_Sn_m1: ∀n,m:nat. S m + n = n + S m. #n; nelim n; nnormalize; //; nqed. *) -(* -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. @@ -123,7 +123,7 @@ ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a. -//; nqed. +//; nqed. ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). #n; nelim n; nnormalize; /3/; nqed. @@ -159,7 +159,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 +167,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. @@ -220,13 +220,16 @@ interpretation "natural 'less than'" 'lt x y = (lt x y). interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). -ndefinition ge: nat \to nat \to Prop \def -\lambda n,m:nat.m \leq n. +(* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m). +//; nqed. *) + +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. @@ -744,20 +767,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. -napply not_lt_to_le; //; +#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. @@ -828,26 +848,33 @@ 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/; - ##|/3/; + ##|#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. +napply nat_elim2; nnormalize; //; nqed. 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. @@ -856,15 +883,15 @@ 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. -#n; #m; #lemn; napplyS symmetric_eq; + m ≤ n → n = (n-m)+m. +#n; #m; #lemn; napplyS sym_eq; napplyS (plus_minus m n m); //; nqed. 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. @@ -876,14 +903,14 @@ nqed. ntheorem plus_to_minus :∀n,m,p:nat.n = m+p → n-m = p. (* /4/ done in 43.5 *) #n; #m; #p; #eqp; -napply symmetric_eq; +napply sym_eq; napplyS (minus_plus_m_m p m); 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. @@ -973,6 +1000,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: @@ -981,3 +1009,143 @@ ntheorem monotonic_le_minus_r: napply le_plus_to_minus; napply (transitive_le ??? (le_plus_minus_m_m ? q));/2/; nqed. + +(*********************** boolean arithmetics ********************) +include "basics/bool.ma". + +nlet rec eqb n m ≝ +match n with + [ O ⇒ match m with [ O ⇒ true | S q ⇒ false] + | S p ⇒ match m with [ O ⇒ false | S q ⇒ eqb p q] + ]. + +(* +ntheorem eqb_to_Prop: ∀n,m:nat. +match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m]. +intros. +apply (nat_elim2 +(\lambda n,m:nat.match (eqb n m) with +[ true \Rightarrow n = m +| false \Rightarrow n \neq m])). +intro.elim n1. +simplify.reflexivity. +simplify.apply not_eq_O_S. +intro. +simplify.unfold Not. +intro. apply (not_eq_O_S n1).apply sym_eq.assumption. +intros.simplify. +generalize in match H. +elim ((eqb n1 m1)). +simplify.apply eq_f.apply H1. +simplify.unfold Not.intro.apply H1.apply inj_S.assumption. +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. +#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. +#n; #m; napply (eqb_elim n m);/2/; +nqed. + +ntheorem eq_to_eqb_true: ∀n,m:nat. + n = m → eqb n m = true. +//; nqed. + +ntheorem not_eq_to_eqb_false: ∀n,m:nat. + n ≠ m → eqb n m = false. +#n; #m; #noteq; +napply eqb_elim;//; +#Heq; napply False_ind; /2/; +nqed. + +nlet rec leb n m ≝ +match n with + [ O ⇒ true + | (S p) ⇒ + match m with + [ O ⇒ false + | (S q) ⇒ leb p q]]. + +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/; + ##|#n; #m; #Hind; #P; #Pt; #Pf; napply Hind; + ##[#lenm; napply Pt; napply le_S_S;//; + ##|#nlenm; napply Pf; /2/; + ##] + ##] +nqed. + +ntheorem leb_true_to_le:∀n,m.leb n m = true → n ≤ m. +#n; #m; napply leb_elim; + ##[//; + ##|#_; #abs; napply False_ind; /2/; + ##] +nqed. + +ntheorem leb_false_to_not_le:∀n,m. + leb n m = false → n ≰ m. +#n; #m; napply leb_elim; + ##[#_; #abs; napply False_ind; /2/; + ##|//; + ##] +nqed. + +ntheorem le_to_leb_true: ∀n,m. n ≤ m → leb n m = true. +#n; #m; napply leb_elim; //; +#H; #H1; napply False_ind; /2/; +nqed. + +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. + +ntheorem ltb_elim: ∀n,m:nat. ∀P:bool → Prop. +(n < m → P true) → (n ≮ m → P false) → P (ltb n m). +#n; #m; #P; #Hlt; #Hnlt; +napply leb_elim; /3/; nqed. + +ntheorem ltb_true_to_lt:∀n,m.ltb n m = true → n < m. +#n; #m; #Hltb; napply leb_true_to_le; nassumption; +nqed. + +ntheorem ltb_false_to_not_lt:∀n,m. + ltb n m = false → n ≮ m. +#n; #m; #Hltb; napply leb_false_to_not_le; nassumption; +nqed. + +ntheorem lt_to_ltb_true: ∀n,m. n < m → ltb n m = true. +#n; #m; #Hltb; napply le_to_leb_true; nassumption; +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. *) +