X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=120a0bb16199c888d16866c495c1fbc7a2f1579e;hb=603f2f6b596d8632b9bd53c73ae0b9c3575231e0;hp=8efe32043339fc033900d6495826b2b96d6d17fe;hpb=a356482acf3bf92eaa34d2b38ee2e81761f58520;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 8efe32043..120a0bb16 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -76,9 +76,9 @@ ntheorem nat_elim2 : #R; #ROn; #RSO; #RSS; #n; nelim n;//; #n0; #Rn0m; #m; ncases m;/2/; nqed. -ntheorem decidable_eq_nat : \forall n,m:nat.decidable (n=m). +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,44 +263,101 @@ 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. -ntheorem le_S_S_to_le : ∀n,m:nat. S n ≤ S m → n ≤ m. +ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m. +(* XXX *) nletin hint ≝ monotonic. /2/; nqed. + +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. /2/; nqed. +ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m. +#n; #m; #Hlt; nelim Hlt;//; nqed. + +(* lt vs. le *) +ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O. +#n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed. + +ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m. +/3/; nqed. + +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). (* -nchange with (pred (S n) ≤ pred (S m)); -nelim leSS; apply le_n.apply (trans_le ? (pred n1)).assumption. -apply le_pred_n. -qed. +ntheorem decidable_le: ∀n,m. decidable (n≤m). +napply nat_elim2; #n; /3/; +#m; #dec; ncases dec;/4/; nqed. *) -theorem lt_S_S_to_lt: \forall n,m. - S n < S m \to n < m. -intros. apply le_S_S_to_le. assumption. -qed. +ntheorem decidable_lt: ∀n,m. decidable (n < m). +#n; #m; napply decidable_le ; nqed. -theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. -intros; -unfold lt in H; -apply (le_S_S ? ? H). -qed. +ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n. +#n; nelim n; /3/; nqed. -theorem leS_to_not_zero : \forall n,m:nat. S n \leq m \to not_zero m. -intros.elim H.exact I.exact I. -qed. +ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m. +/2/; nqed. -(* not le *) -theorem not_le_Sn_O: \forall n:nat. S n \nleq O. -intros.unfold Not.simplify.intros.apply (leS_to_not_zero ? ? H). -qed. +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;/4/; + ##] +nqed. -theorem not_le_Sn_n: \forall n:nat. S n \nleq n. -intros.elim n.apply not_le_Sn_O.unfold Not.simplify.intros.cut (S n1 \leq n1). -apply H.assumption. -apply le_S_S_to_le.assumption. -qed. +ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n. +#n; #m; #Hltnm; nelim Hltnm;/3/; nqed. + +ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n. +#n; #m; #Hnlt; napply lt_S_to_le; +(* something strange here: /2/ fails: + we need an extra depths for unfolding not *) +napply not_le_to_lt; napply Hnlt; nqed. + +ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n. +/2/; nqed. + +(* lt and le trans *) +ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p. +#n; #m; #p; #H; #H1; nelim H1; /2/; nqed. + +ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. +#n; #m; #p; #H; nelim H; /3/; nqed. + +ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m. +/2/; nqed. + +ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m. +/2/; nqed. + +(* +theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat. +(S O) \lt n \to O \lt (pred n). +intros. +apply (ltn_to_ltO (pred (S O)) (pred n) ?). + apply (lt_pred (S O) n); + [ apply (lt_O_S O) + | assumption + ] +qed. *) + +ntheorem lt_O_n_elim: ∀n:nat. O < n → + ∀P:nat → Prop.(∀m:nat.P (S m)) → P n. +#n; nelim n; //; #abs; napply False_ind; /2/; nqed. + +(* theorem lt_pred: \forall n,m. O < n \to n < m \to pred n < pred m. apply nat_elim2 @@ -325,48 +386,18 @@ elim n; ]. qed. -theorem le_to_le_pred: - ∀n,m. n ≤ m → pred n ≤ pred m. -intros 2; -elim n; -[ simplify; - apply le_O_n -| simplify; - elim m in H1 ⊢ %; - [ elim (not_le_Sn_O ? H1) - | simplify; - apply le_S_S_to_le; - assumption - ] -]. -qed. +*) (* le to lt or eq *) -theorem le_to_or_lt_eq : \forall n,m:nat. -n \leq m \to n < m \lor n = m. -intros.elim H. -right.reflexivity. -left.unfold lt.apply le_S_S.assumption. -qed. - -theorem Not_lt_n_n: ∀n. n ≮ n. -intro; -unfold Not; -intro; -unfold lt in H; -apply (not_le_Sn_n ? H). -qed. +ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. +#n; #m; #lenm; nelim lenm; /3/; nqed. (* not eq *) -theorem lt_to_not_eq : \forall n,m:nat. n H2 in H1. + rewrite > (S_pred a) in H1 + [ apply False_ind. + apply (eq_to_not_lt O ((S (pred a))*(S m))) + [ apply sym_eq. + assumption + | apply lt_O_times_S_S + ] + | assumption + ] +] +qed. + +theorem O_lt_times_to_O_lt: \forall a,c:nat. +O \lt (a * c) \to O \lt a. +intros. +apply (nat_case1 a) +[ intros. + rewrite > H1 in H. + simplify in H. + assumption +| intros. + apply lt_O_S +] +qed. + +lemma lt_times_to_lt_O: \forall i,n,m:nat. i < n*m \to O < m. +intros. +elim (le_to_or_lt_eq O ? (le_O_n m)) + [assumption + |apply False_ind. + rewrite < H1 in H. + rewrite < times_n_O in H. + apply (not_le_Sn_O ? H) + ] +qed. *) + +(* +ntheorem monotonic_lt_times_r: +∀n:nat.monotonic nat lt (λm.(S n)*m). +/2/; +simplify. +intros.elim n. +simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption. +apply lt_plus.assumption.assumption. +qed. *) + +ntheorem monotonic_lt_times_l: + ∀c:nat. O < c → monotonic nat lt (λt.(t*c)). +#c; #posc; #n; #m; #ltnm; +nelim ltnm; nnormalize; + ##[napplyS monotonic_lt_plus_l;//; + ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//; + ##] +nqed. + +ntheorem monotonic_lt_times_r: + ∀c:nat. O < c → monotonic nat lt (λt.(c*t)). +(* /2/ lentissimo *) +#c; #posc; #n; #m; #ltnm; +(* why?? napplyS (monotonic_lt_times_l c posc n m ltnm); *) +nrewrite > (symmetric_times c n); +nrewrite > (symmetric_times c m); +napply monotonic_lt_times_l;//; +nqed. + +ntheorem lt_to_le_to_lt_times: +∀n,m,p,q:nat. n < m → p ≤ q → O < q → n*p < m*q. +#n; #m; #p; #q; #ltnm; #lepq; #posq; +napply (le_to_lt_to_lt ? (n*q)); + ##[napply monotonic_le_times_r;//; + ##|napply monotonic_lt_times_l;//; + ##] +nqed. + +ntheorem lt_times:∀n,m,p,q:nat. n nat_compare_n_n.reflexivity. +intro.apply nat_compare_elim.intro. +absurd (p minus_Sn_m. +apply le_S.assumption. +apply lt_to_le.assumption. +qed. + +theorem minus_le_S_minus_S: \forall n,m:nat. m-n \leq S (m-(S n)). +intros. +apply (nat_elim2 (\lambda n,m.m-n \leq S (m-(S n)))). +intro.elim n1.simplify.apply le_n_Sn. +simplify.rewrite < minus_n_O.apply le_n. +intros.simplify.apply le_n_Sn. +intros.simplify.apply H. +qed. + +theorem lt_minus_S_n_to_le_minus_n : \forall n,m,p:nat. m-(S n) < p \to m-n \leq p. +intros 3.intro. +(* autobatch *) +(* end auto($Revision: 9739 $) proof: TIME=1.33 SIZE=100 DEPTH=100 *) +apply (trans_le (m-n) (S (m-(S n))) p). +apply minus_le_S_minus_S. +assumption. +qed. + +theorem le_minus_m: \forall n,m:nat. n-m \leq n. +intros.apply (nat_elim2 (\lambda m,n. n-m \leq n)). +intros.rewrite < minus_n_O.apply le_n. +intros.simplify.apply le_n. +intros.simplify.apply le_S.assumption. +qed. + +theorem lt_minus_m: \forall n,m:nat. O < n \to O < m \to n-m \lt n. +intros.apply (lt_O_n_elim n H).intro. +apply (lt_O_n_elim m H1).intro. +simplify.unfold lt.apply le_S_S.apply le_minus_m. +qed. + +theorem minus_le_O_to_le: \forall n,m:nat. n-m \leq O \to n \leq m. +intros 2. +apply (nat_elim2 (\lambda n,m:nat.n-m \leq O \to n \leq m)). +intros.apply le_O_n. +simplify.intros. assumption. +simplify.intros.apply le_S_S.apply H.assumption. +qed. +*) + +(* monotonicity and galois *) + +ntheorem monotonic_le_minus_l: +∀p,q,n:nat. q ≤ p → q-n ≤ p-n. +napply nat_elim2; #p; #q; + ##[#lePO; napply (le_n_O_elim ? lePO);//; + ##|//; + ##|#Hind; #n; ncases n; + ##[//; + ##|#a; #leSS; napply Hind; /2/; + ##] + ##] +nqed. + +ntheorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m. +#n; #m; #p; #lep; +napply transitive_le; + ##[##|napply le_plus_minus_m_m + ##|napply monotonic_le_plus_l;//; + ##] +nqed. + +ntheorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. +#n; #m; #p; #lep; +(* bello *) +napplyS monotonic_le_minus_l;//; +nqed. + +ntheorem monotonic_le_minus_r: +∀p,q,n:nat. q ≤ p → n-p ≤ n-q. +#p; #q; #n; #lepq; +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; +nelim (true_or_false (eqb n m)); //; +#Heq; napply False_ind; napply noteq;/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; #leSS; /3/; + ##] + ##] +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/; + ##|/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 lt_to_leb_false: ∀n,m. m < n → leb n m = false. +#n; #m; napply leb_elim; //; +#H; #H1; napply False_ind; /2/; +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. *)