X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Farithmetics%2Fnat.ma;h=649108f47c27110a30a770489ce0d9f82bd221d3;hb=7163f2ff5eb4960162d2fcf135f031ae2a9f8b56;hp=dc0af3a6b14663ef1e6d4e269a8b9031129894f4;hpb=16a95f57b09ae92ea24ab2addd02c1d0be80f109;p=helm.git diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index dc0af3a6b..649108f47 100644 --- a/matita/matita/lib/arithmetics/nat.ma +++ b/matita/matita/lib/arithmetics/nat.ma @@ -11,6 +11,10 @@ include "basics/relations.ma". +(* Definitions **************************************************************) + +(* natural numbers *) + inductive nat : Type[0] ≝ | O : nat | S : nat → nat. @@ -22,27 +26,62 @@ alias num (instance 0) = "natural number". definition pred ≝ λn. match n with [ O ⇒ O | S p ⇒ p]. -theorem pred_Sn : ∀n. n = pred (S n). -// qed. +definition not_zero: nat → Prop ≝ + λn: nat. match n with [ O ⇒ False | (S p) ⇒ True ]. -theorem injective_S : injective nat nat S. -// qed. +(* order relations *) -(* -theorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. -//. qed. *) +inductive le (n:nat) : nat → Prop ≝ + | le_n : le n n + | le_S : ∀ m:nat. le n m → le n (S m). -theorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2/ qed. +interpretation "natural 'less or equal to'" 'leq x y = (le x y). -definition not_zero: nat → Prop ≝ - λn: nat. match n with [ O ⇒ False | (S p) ⇒ True ]. +interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)). -theorem not_eq_O_S : ∀n:nat. O ≠ S n. -#n @nmk #eqOS (change with (not_zero O)) >eqOS // qed. +definition lt: nat → nat → Prop ≝ λn,m. S n ≤ m. -theorem not_eq_n_Sn: ∀n:nat. n ≠ S n. -#n (elim n) /2/ qed. +interpretation "natural 'less than'" 'lt x y = (lt x y). + +interpretation "natural 'not less than'" 'nless x y = (Not (lt x y)). + +definition ge: nat → nat → Prop ≝ λn,m.m ≤ n. + +interpretation "natural 'greater or equal to'" 'geq x y = (ge x y). + +definition gt: nat → nat → Prop ≝ λn,m.meqOS // qed. -definition gt: nat → nat → Prop ≝ λn,m.m (S_pred m) - [ apply le_S_S - assumption - | assumption - ] -]. -qed. - -*) - -(* le to lt or eq *) -theorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m. -#n #m #lenm (elim lenm) /3/ qed. - -(* not eq *) -theorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m. -#n #m #H @not_to_not /2/ qed. - -(*not lt -theorem 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. *) - -theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n H7. -apply H. -apply le_to_or_lt_eq.apply H6. -qed. -*) - -(*********************** monotonicity ***************************) theorem monotonic_le_plus_r: ∀n:nat.monotonic nat le (λm.n + m). #n #a #b (elim n) normalize // #m #H #leab @le_S_S /2/ qed. -(* -theorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m -≝ monotonic_le_plus_r. *) - theorem monotonic_le_plus_l: ∀m:nat.monotonic nat le (λn.n + m). /2/ qed. -(* -theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p -\def monotonic_le_plus_l. *) - theorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2 → m1 ≤ m2 → n1 + m1 ≤ n2 + m2. #n1 #n2 #m1 #m2 #len #lem @(transitive_le ? (n1+m2)) /2/ qed. theorem le_plus_n :∀n,m:nat. m ≤ n + m. -/2/ qed. - -lemma le_plus_a: ∀a,n,m. n ≤ m → n ≤ a + m. -/2/ qed. - -lemma le_plus_b: ∀b,n,m. n + b ≤ m → n ≤ m. -/2/ qed. - -theorem le_plus_n_r :∀n,m:nat. m ≤ m + n. -/2/ qed. - -theorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n. -// qed. - -theorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m. -#a (elim a) normalize /3/ qed. - -theorem le_plus_to_le_r: ∀a,n,m. n + a ≤ m +a → n ≤ m. -/2/ qed. - -(* plus & lt *) - -theorem monotonic_lt_plus_r: -∀n:nat.monotonic nat lt (λm.n+m). -/2/ qed. - -(* -variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def -monotonic_lt_plus_r. *) - -theorem monotonic_lt_plus_l: -∀n:nat.monotonic nat lt (λm.m+n). -/2/ qed. - -(* -variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def -monotonic_lt_plus_l. *) - -theorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q. -#n #m #p #q #ltnm #ltpq -@(transitive_lt ? (n+q))/2/ qed. +/2/ qed. -theorem 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. +(* lt *) -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 -] +theorem transitive_lt: transitive nat lt. +#a #b #c #ltab #ltbc (elim ltbc) /2/ 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. *) +theorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p. +#n #m #p #H #H1 (elim H1) /2/ qed. -(* -theorem 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. *) +theorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. +#n #m #p #H (elim H) /3/ qed. -theorem monotonic_lt_times_l: - ∀c:nat. O < c → monotonic nat lt (λt.(t*c)). +theorem lt_S_to_lt: ∀n,m. S n < m → n < m. +/2/ qed. + +theorem ltn_to_ltO: ∀n,m:nat. n < m → 0 < m. +/2/ qed. + +theorem lt_O_S : ∀n:nat. O < S n. +/2/ qed. + +theorem monotonic_lt_plus_r: +∀n:nat.monotonic nat lt (λm.n+m). +/2/ qed. + +theorem monotonic_lt_plus_l: +∀n:nat.monotonic nat lt (λm.m+n). +/2/ qed. + +theorem lt_plus: ∀n,m,p,q:nat. n < m → p < q → n + p < m + q. +#n #m #p #q #ltnm #ltpq +@(transitive_lt ? (n+q))/2/ qed. + +theorem lt_plus_to_lt_l :∀n,p,q:nat. p+n < q+n → p 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 plus_minus_commutative: ∀x,y,z. z ≤ y → x + (y - z) = x + y - z. +/2 by plus_minus/ 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. +(* More atomic conclusion ***************************************************) -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. -*) +(* le *) -(* monotonicity and galois *) +theorem le_n_fn: ∀f:nat → nat. + increasing f → ∀n:nat. n ≤ f n. +#f #incr #n (elim n) /2/ +qed-. theorem monotonic_le_minus_l: ∀p,q,n:nat. q ≤ p → q-n ≤ p-n. @@ -921,6 +576,10 @@ theorem le_minus_to_plus: ∀n,m,p. n-m ≤ p → n≤ p+m. [|@le_plus_minus_m_m | @monotonic_le_plus_l // ] qed. +theorem le_minus_to_plus_r: ∀a,b,c. c ≤ b → a ≤ b - c → a + c ≤ b. +#a #b #c #Hlecb #H >(plus_minus_m_m … Hlecb) /2/ +qed. + theorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. #n #m #p #lep /2/ qed. @@ -930,6 +589,55 @@ theorem monotonic_le_minus_r: @(transitive_le … (le_plus_minus_m_m ? q)) /2/ qed. +theorem increasing_to_le: ∀f:nat → nat. + increasing f → ∀m:nat.∃i.m ≤ f i. +#f #incr #m (elim m) /2/#n * #a #lenfa +@(ex_intro ?? (S a)) /2/ +qed. + +(* thus is le_plus +lemma le_plus_compatible: ∀x1,x2,y1,y2. x1 ≤ y1 → x2 ≤ y2 → x1 + x2 ≤ y1 + y2. +#x1 #y1 #x2 #y2 #H1 #H2 /2/ @le_plus // /2/ /3 by le_minus_to_plus, monotonic_le_plus_r, transitive_le/ qed. +*) + +(* lt *) + +theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n(eq_minus_O …) /2/ >(eq_minus_O …) // + [> eq_minus_O [2:/2/] >eq_minus_O // @monotonic_le_times_r /2/ - |@sym_eq (applyS plus_to_minus) <(distributive_times_plus …) - (* STRANO *) - @(eq_f …b) (applyS plus_minus_m_m) /2/ + |@sym_eq (applyS plus_to_minus) eq_minus_O /2/ (* >eq_minus_O // *) + ] +qed. + +theorem minus_minus: ∀n,m,p:nat. p ≤ m → m ≤ n → + p+(n-m) = n-(m-p). +#n #m #p #lepm #lemn +@sym_eq @plus_to_minus (plus_minus_m_m b c) in ⊢ (? ? ?%); // qed. +lemma minus_minus_m_m: ∀m,n. n ≤ m → m - (m - n) = n. +/2 width=1/ qed. + +(* Stilll more atomic conclusion ********************************************) + +(* le *) + +lemma le_fwd_plus_plus_ge: ∀m1,m2. m2 ≤ m1 → ∀n1,n2. m1 + n1 ≤ m2 + n2 → n1 ≤ n2. +#m1 #m2 #H #n1 #n2 >commutative_plus +#H elim (le_inv_plus_l … H) -H >commutative_plus