X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=0418444e376e0ecd47ca07a83b8915b1aa8cbc57;hb=221472ea1597505d12677f5742e388125a15e2b9;hp=189f2f1524598296296b945f3a4841685f05acdb;hpb=b913b3ee973dfc9f1e6d7c0fd0c720d03699689e;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 189f2f152..0418444e3 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -13,6 +13,7 @@ (**************************************************************************) (* include "higher_order_defs/functions.ma". *) +include "hints_declaration.ma". include "basics/functions.ma". include "basics/eq.ma". @@ -22,21 +23,26 @@ ninductive nat : Type[0] ≝ interpretation "Natural numbers" 'N = nat. -default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind. +alias num (instance 0) = "nnatural number". -alias num (instance 0) = "natural number". +(* +nrecord pos : Type ≝ + {n:>nat; is_pos: n ≠ 0}. + +ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on +*) + +(* default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind. +*) 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. +//; nqed. ntheorem injective_S : injective nat nat S. -(* nwhd; #a; #b;#H;napplyS pred_Sn. *) -//. nqed. +//; nqed. (* ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. @@ -47,8 +53,7 @@ ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. ndefinition not_zero: nat → Prop ≝ λn: nat. match n with - [ O ⇒ False - | (S p) ⇒ True ]. + [ O ⇒ False | (S p) ⇒ True ]. ntheorem not_eq_O_S : ∀n:nat. O ≠ S n. #n; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. @@ -71,11 +76,11 @@ 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). -napply nat_elim2;#n; +ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). +napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; (* ??? /2/;#neqnm; /3/; *) + ##| #m; #Hind; ncases Hind; /3/; ##] nqed. @@ -88,32 +93,46 @@ nlet rec plus n m ≝ interpretation "natural plus" 'plus x y = (plus x y). -ntheorem plus_n_O: ∀n:nat. n = n+O. -#n; nelim n; /2/; nqed. +ntheorem plus_O_n: ∀n:nat. n = 0+n. +//; nqed. + +(* +ntheorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m. +//; nqed. +*) -ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n+(S m). +ntheorem plus_n_O: ∀n:nat. n = n+0. #n; nelim n; nnormalize; //; nqed. -ntheorem plus_n_SO : ∀n:nat. S n = n+(S O). -//; nqed. +ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m. +#n; nelim n; nnormalize; //; nqed. -ntheorem sym_plus: ∀n,m:nat. n+m = m+n. +(* +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. *) + +ntheorem symmetric_plus: symmetric ? plus. +#n; nelim n; nnormalize; //; nqed. ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. -(* ntheorem assoc_plus : \forall n,m,p:nat. (n+m)+p = n+(m+p) -\def associative_plus. *) +ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a. +//; nqed. ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). #n; nelim n; nnormalize; /3/; nqed. (* ntheorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m -\def injective_plus_r. *) +\def injective_plus_r. ntheorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m). -/2/; nqed. +/2/; nqed. *) (* ntheorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m \def injective_plus_l. *) @@ -130,103 +149,835 @@ interpretation "natural times" 'times x y = (times x y). ntheorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m. //; nqed. +ntheorem times_O_n: ∀n:nat. O = O*n. +//; nqed. + ntheorem times_n_O: ∀n:nat. O = n*O. #n; nelim n; //; nqed. ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). -#n; nelim n; nnormalize; /2/; nqed. +#n; nelim n; nnormalize; //; nqed. + +ntheorem symmetric_times : symmetric nat times. +#n; nelim n; nnormalize; //; nqed. + +(* variant sym_times : \forall n,m:nat. n*m = m*n \def +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 associative_times: associative nat times. +#n; nelim n; nnormalize; //; nqed. + +nlemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). +//; nqed. + +(* ci servono questi risultati? ntheorem times_O_to_O: ∀n,m:nat.n*m=O → n=O ∨ m=O. -napply nat_elim2; (* /2/ slow! *) - ##[#n; #H; @1; //; - ##|#n; #H; @2; //; (* ?? *) - ##|#n; #m; #H; #H1; napply False_ind;napply not_eq_O_S; - (* why two steps? *) /2/; - ##] -nqed. +napply nat_elim2; /2/; +#n; #m; #H; nnormalize; #H1; napply False_ind;napply not_eq_O_S; +//; nqed. ntheorem times_n_SO : ∀n:nat. n = n * S O. +#n; //; nqed. + +ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n. +nnormalize; //; nqed. + +nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)). //; nqed. +ntheorem or_eq_eq_S: \forall n.\exists m. +n = (S(S O))*m \lor n = S ((S(S O))*m). +#n; nelim n; + ##[@; /2/; + ##|#a; #H; nelim H; #b;#or;nelim or;#aeq; + ##[@ b; @ 2; //; + ##|@ (S b); @ 1; /2/; + ##] +nqed. +*) + +(******************** ordering relations ************************) + +ninductive le (n:nat) : nat → Prop ≝ + | le_n : le n n + | le_S : ∀ m:nat. le n m → le n (S m). + +interpretation "natural 'less or equal to'" 'leq x y = (le x y). + +interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)). + +ndefinition lt: nat → nat → Prop ≝ +λn,m:nat. S n ≤ m. + +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. + +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 H2. - apply sym_eq. - apply times_SSO - ] +theorem le_pred_to_le: + ∀n,m. O < m → pred n ≤ pred m → n ≤ m. +intros 2; +elim n; +[ apply le_O_n +| simplify in H2; + rewrite > (S_pred m); + [ apply le_S_S; + assumption + | assumption ] +]. +qed. + +*) + +(* le to lt or eq *) +ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. +#n; #m; #lenm; nelim lenm; /3/; nqed. + +(* not eq *) +ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. +/2/; nqed. + +(*not lt +ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b. +intros. +unfold Not. +intros. +rewrite > H in H1. +apply (lt_to_not_eq b b) +[ assumption +| reflexivity +] +qed. + +theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n. +intros; +unfold Not; +intro; +unfold lt in H; +unfold lt in H1; +generalize in match (le_S_S ? ? H); +intro; +generalize in match (transitive_le ? ? ? H2 H1); +intro; +apply (not_le_Sn_n ? H3). +qed. *) + +ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n H.apply times_n_Sm.] +theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m +\def antisymmetric_le. + +theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m. +intros; +unfold lt in H1; +generalize in match (le_S_S_to_le ? ? H1); +intro; +apply antisym_le; +assumption. qed. +*) -variant sym_times : \forall n,m:nat. n*m = m*n \def -symmetric_times. +(* well founded induction principles *) + +ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop. +(∀m.(∀p. p < m → P p) → P m) → P n. +#n; #P; #H; +ncut (∀q:nat. q ≤ n → P q);/2/; +nelim n; + ##[#q; #HleO; (* applica male *) + napply (le_n_O_elim ? HleO); + napply H; #p; #ltpO; + napply False_ind; /2/; + ##|#p; #Hind; #q; #HleS; + napply H; #a; #lta; napply Hind; + napply le_S_S_to_le;/2/; + ##] +nqed. -theorem distributive_times_plus : distributive nat times plus. -unfold distributive. -intros.elim x. -simplify.reflexivity. -simplify. -demodulate all. +(* some properties of functions *) (* -rewrite > H. rewrite > assoc_plus.rewrite > assoc_plus. -apply eq_f.rewrite < assoc_plus. rewrite < (sym_plus ? z). -rewrite > assoc_plus.reflexivity. *) +definition increasing \def \lambda f:nat \to nat. +\forall n:nat. f n < f (S n). + +theorem increasing_to_monotonic: \forall f:nat \to nat. +increasing f \to monotonic nat lt f. +unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H. +apply (trans_le ? (f n1)). +assumption.apply (trans_le ? (S (f n1))). +apply le_n_Sn. +apply H. +qed. + +theorem le_n_fn: \forall f:nat \to nat. (increasing f) +\to \forall n:nat. n \le (f n). +intros.elim n. +apply le_O_n. +apply (trans_le ? (S (f n1))). +apply le_S_S.apply H1. +simplify in H. unfold increasing in H.unfold lt in H.apply H. qed. -variant distr_times_plus: \forall n,m,p:nat. n*(m+p) = n*m + n*p -\def distributive_times_plus. +theorem increasing_to_le: \forall f:nat \to nat. (increasing f) +\to \forall m:nat. \exists i. m \le (f i). +intros.elim m. +apply (ex_intro ? ? O).apply le_O_n. +elim H1. +apply (ex_intro ? ? (S a)). +apply (trans_le ? (S (f a))). +apply le_S_S.assumption. +simplify in H.unfold increasing in H.unfold lt in H. +apply H. +qed. -theorem associative_times: associative nat times. -unfold associative. +theorem increasing_to_le2: \forall f:nat \to nat. (increasing f) +\to \forall m:nat. (f O) \le m \to +\exists i. (f i) \le m \land m <(f (S i)). +intros.elim H1. +apply (ex_intro ? ? O). +split.apply le_n.apply H. +elim H3.elim H4. +cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))). +elim Hcut. +apply (ex_intro ? ? a). +split.apply le_S. assumption.assumption. +apply (ex_intro ? ? (S a)). +split.rewrite < H7.apply le_n. +rewrite > H7. +apply H. +apply le_to_or_lt_eq.apply H6. +qed. +*) + +(*********************** monotonicity ***************************) +ntheorem monotonic_le_plus_r: +∀n:nat.monotonic nat le (λm.n + m). +#n; #a; #b; nelim n; nnormalize; //; +#m; #H; #leab;napply le_S_S; /2/; nqed. + +(* +ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m +≝ monotonic_le_plus_r. *) + +ntheorem monotonic_le_plus_l: +∀m:nat.monotonic nat le (λn.n + m). +/2/; nqed. + +(* +ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p +\def monotonic_le_plus_l. *) + +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; +/2/; nqed. + +ntheorem le_plus_n :∀n,m:nat. m ≤ n + m. +/2/; nqed. + +ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n. +/2/; nqed. + +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. + +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 x. simplify.apply refl_eq. +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. -demodulate all. -(* -rewrite < sym_times. ->>>>>>> .r9770 -rewrite > distr_times_plus. -rewrite < sym_times. -rewrite < (sym_times (times n y) z). -rewrite < H.apply refl_eq. -*) +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. -variant assoc_times: \forall n,m,p:nat. (n*m)*p = n*(m*p) \def -associative_times. +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. -lemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). -intros. -demodulate. reflexivity. -(* autobatch paramodulation. *) +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. -*) \ No newline at end of file +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.