X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=3152b0dbe3503aeee43bbd0f872b847f87979239;hb=eb4144a401147a44a9620169eb6dafeb8f5a2c17;hp=b6683b267801bac94d220ad5ca573ac644eaf1a5;hpb=8899a3f240f62633f4df58b2ee358fa285a82d1d;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index b6683b267..3152b0dbe 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -310,7 +310,8 @@ ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n. 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 *) +(* 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. @@ -528,21 +529,23 @@ apply le_to_or_lt_eq.apply H6. qed. *) -(******************* monotonicity ******************************) +(*********************** 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. +≝ 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. +\def monotonic_le_plus_l. *) ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 \to m1 ≤ m2 → n1 + m1 ≤ n2 + m2. @@ -559,88 +562,234 @@ 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. - -(* times -theorem monotonic_le_times_r: -\forall n:nat.monotonic nat le (\lambda m. n * m). -simplify.intros.elim n. -simplify.apply le_O_n. -simplify.apply le_plus. -assumption. -assumption. -qed. +#a; nelim a; /3/; nqed. -theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m -\def monotonic_le_times_r. +ntheorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m. +/2/; nqed. -theorem monotonic_le_times_l: -\forall m:nat.monotonic nat le (\lambda n.n*m). -simplify.intros. -rewrite < sym_times.rewrite < (sym_times m). -apply le_times_r.assumption. -qed. +(* times *) +ntheorem monotonic_le_times_r: +∀n:nat.monotonic nat le (λm. n * m). +#n; #x; #y; #lexy; nelim n; nnormalize;//;(* lento /2/;*) +#a; #lea; napply le_plus; //; +nqed. + +(* +ntheorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m +\def monotonic_le_times_r. *) + +ntheorem monotonic_le_times_l: +∀m:nat.monotonic nat le (λn.n*m). +/2/; nqed. +(* theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p -\def monotonic_le_times_l. +\def monotonic_le_times_l. *) + +ntheorem le_times: ∀n1,n2,m1,m2:nat. +n1 ≤ n2 → m1 ≤ m2 → n1*m1 ≤ n2*m2. +#n1; #n2; #m1; #m2; #len; #lem; +napply transitive_le; (* /2/ slow *) + ##[ ##| napply monotonic_le_times_l;//; + ##| napply monotonic_le_times_r;//; + ##] +nqed. + +ntheorem lt_times_n: ∀n,m:nat. O < n → m ≤ n*m. +(* bello *) +/2/; nqed. + +ntheorem le_times_to_le: +∀a,n,m. O < a → a * n ≤ a * m → n ≤ m. +#a; napply nat_elim2; nnormalize; + ##[//; + ##|#n; #H1; #H2; napply False_ind; + ngeneralize in match H2; + napply lt_to_not_le; + napply (transitive_le ? (S n));/2/; + ##|#n; #m; #H; #lta; #le; + napply le_S_S; napply H; /2/; + ##] +nqed. + +ntheorem le_S_times_2: ∀n,m.O < m → n ≤ m → n < 2*m. +#n; #m; #posm; #lenm; (* interessante *) +nnormalize; napplyS (le_plus n); //; nqed. + +(************************** minus ******************************) + +nlet rec minus n m ≝ + match n with + [ O ⇒ O + | S p ⇒ + match m with + [ O ⇒ S p + | S q ⇒ minus p q ]]. + +interpretation "natural minus" 'minus x y = (minus x y). + +ntheorem minus_S_S: ∀n,m:nat.S n - S m = n -m. +//; nqed. + +ntheorem minus_O_n: ∀n:nat.O=O-n. +#n; ncases n; //; nqed. + +ntheorem minus_n_O: ∀n:nat.n=n-O. +#n; ncases n; //; nqed. + +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. + +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/; + ##] +nqed. + +ntheorem eq_minus_S_pred: ∀n,m. n - (S m) = pred(n -m). +napply nat_elim2; //; 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/; + ##|nnormalize;/3/; + ##] +nqed. -theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2 -\to n1*m1 \le n2*m2. +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; +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;/2/; + ##] +nqed. + +ntheorem minus_to_plus :∀n,m,p:nat. + m ≤ n → n-m = p → n = m+p. +#n; #m; #p; #lemn; #eqp; napplyS plus_minus_m_m; //; +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; +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 m posm);//. +nqed. + +(* +theorem eq_minus_n_m_O: \forall n,m:nat. +n \leq m \to n-m = O. +intros 2. +apply (nat_elim2 (\lambda n,m.n \leq m \to n-m = O)). +intros.simplify.reflexivity. +intros.apply False_ind. +apply not_le_Sn_O; +[2: apply H | skip]. intros. -apply (trans_le ? (n2*m1)). -apply le_times_l.assumption. -apply le_times_r.assumption. +simplify.apply H.apply le_S_S_to_le. apply H1. qed. -theorem le_times_n: \forall n,m:nat.(S O) \le n \to m \le n*m. -intros.elim H.simplify. -elim (plus_n_O ?).apply le_n. -simplify.rewrite < sym_plus.apply le_plus_n. +theorem le_SO_minus: \forall n,m:nat.S n \leq m \to S O \leq m-n. +intros.elim H.elim (minus_Sn_n n).apply le_n. +rewrite > minus_Sn_m. +apply le_S.assumption. +apply lt_to_le.assumption. qed. -theorem le_times_to_le: -\forall a,n,m. S O \le a \to a * n \le a * m \to n \le m. -intro. -apply nat_elim2;intros - [apply le_O_n - |apply False_ind. - rewrite < times_n_O in H1. - generalize in match H1. - apply (lt_O_n_elim ? H). - intros. - simplify in H2. - apply (le_to_not_lt ? ? H2). - apply lt_O_S - |apply le_S_S. - apply H - [assumption - |rewrite < times_n_Sm in H2. - rewrite < times_n_Sm in H2. - apply (le_plus_to_le a). - assumption - ] - ] +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 le_S_times_SSO: \forall n,m.O < m \to -n \le m \to S n \le (S(S O))*m. -intros. -simplify. -rewrite > plus_n_O. -simplify.rewrite > plus_n_Sm. -apply le_plus - [assumption - |rewrite < plus_n_O. - assumption - ] +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. -(*0 and times *) -theorem O_lt_const_to_le_times_const: \forall a,c:nat. -O \lt c \to a \le a*c. -intros. -rewrite > (times_n_SO a) in \vdash (? % ?). -apply le_times -[ apply le_n -| 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.