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=3152b0dbe3503aeee43bbd0f872b847f87979239;hpb=a79bf6edc13daaea8135ca71fdc92e02e229f030;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 3152b0dbe..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,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. 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. @@ -428,9 +439,9 @@ ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m → (* le and eq *) ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m. -napply nat_elim2; /3/; nqed. +napply nat_elim2; /4/; nqed. -ntheorem lt_O_S : \forall n:nat. O < S n. +ntheorem lt_O_S : ∀n:nat. O < S n. /2/; 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,11 +573,45 @@ 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. + +(* +variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def +monotonic_lt_plus_r. *) + +ntheorem monotonic_lt_plus_l: +∀n:nat.monotonic nat lt (λm.m+n). +/2/;nqed. + +(* +variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def +monotonic_lt_plus_l. *) + +ntheorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q. +#n; #m; #p; #q; #ltnm; #ltpq; +napply (transitive_lt ? (n+q));/2/; nqed. + +ntheorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p 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