From 433a742bd425848f273c7faf192c96a51ab9d8c5 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Mon, 12 Dec 2011 13:37:38 +0000 Subject: [PATCH] nat library reorganized .... --- matita/matita/lib/arithmetics/nat.ma | 971 ++++++------------ .../matita/lib/arithmetics/nat_commented.ma | 380 +++++++ 2 files changed, 673 insertions(+), 678 deletions(-) create mode 100644 matita/matita/lib/arithmetics/nat_commented.ma diff --git a/matita/matita/lib/arithmetics/nat.ma b/matita/matita/lib/arithmetics/nat.ma index a20311ced..56482a042 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. -interpretation "natural 'greater than'" 'gt x y = (gt x y). -interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)). +theorem not_eq_n_Sn: ∀n:nat. n ≠ S n. +#n (elim n) /2/ qed. -theorem transitive_le : transitive nat le. -#a #b #c #leab #lebc (elim lebc) /2/ -qed. +(* Atomic conclusion *******************************************************) -(* -theorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p -\def transitive_le. *) +(* not_zero *) -theorem transitive_lt: transitive nat lt. -#a #b #c #ltab #ltbc (elim ltbc) /2/qed. +theorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m. +#n #m #Hlt (elim Hlt) // qed. -(* -theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p -\def transitive_lt. *) +(* le *) theorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m. #n #m #lenm (elim lenm) /2/ qed. -theorem le_O_n : ∀n:nat. O ≤ n. +theorem le_O_n : ∀n:nat. 0 ≤ n. #n (elim n) /2/ qed. theorem le_n_Sn : ∀n:nat. n ≤ S n. /2/ qed. +theorem transitive_le : transitive nat le. +#a #b #c #leab #lebc (elim lebc) /2/ +qed. + theorem le_pred_n : ∀n:nat. pred n ≤ n. #n (elim n) // qed. @@ -241,279 +224,15 @@ theorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m. (* demo *) /2/ qed-. -(* this are instances of the le versions -theorem lt_S_S_to_lt: ∀n,m. S n < S m → n < m. -/2/ qed. - -theorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m. -/2/ qed. *) - -theorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m. -#n #m #Hlt (elim Hlt) // qed. - -(* lt vs. le *) -theorem not_le_Sn_O: ∀ n:nat. S n ≰ O. -#n @nmk #Hlen0 @(lt_to_not_zero ?? Hlen0) qed. - -theorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m. -/3/ qed. - -theorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m. -/3/ qed. - -theorem decidable_le: ∀n,m. decidable (n≤m). -@nat_elim2 #n /2/ #m * /3/ qed. - -theorem decidable_lt: ∀n,m. decidable (n < m). -#n #m @decidable_le qed. - -theorem not_le_Sn_n: ∀n:nat. S n ≰ n. -#n (elim n) /2/ qed. - -(* this is le_S_S_to_le -theorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m. -/2/ qed. -*) - -lemma le_gen: ∀P:nat → Prop.∀n.(∀i. i ≤ n → P i) → P n. -/2/ qed. - -theorem not_le_to_lt: ∀n,m. n ≰ m → m < n. -@nat_elim2 #n - [#abs @False_ind /2/ - |/2/ - |#m #Hind #HnotleSS @le_S_S /3/ - ] -qed. - -theorem lt_to_not_le: ∀n,m. n < m → m ≰ n. -#n #m #Hltnm (elim Hltnm) /3/ qed. - -theorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n. -/4/ qed. - -theorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n. -#n #m #H @lt_to_not_le /2/ (* /3/ *) qed. - -(* lt and le trans *) - -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 le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. -#n #m #p #H (elim H) /3/ qed. - -theorem lt_S_to_lt: ∀n,m. S n < m → n < m. -/2/ qed. - -theorem ltn_to_ltO: ∀n,m:nat. n < m → O < m. -/2/ qed. - -(* -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. *) - -theorem lt_O_n_elim: ∀n:nat. O < n → - ∀P:nat → Prop.(∀m:nat.P (S m)) → P n. -#n (elim n) // #abs @False_ind /2/ -qed. - -theorem S_pred: ∀n. 0 < n → S(pred n) = n. -#n #posn (cases posn) // -qed. - -(* -theorem lt_pred: \forall n,m. - O < n \to n < m \to pred n < pred m. -apply nat_elim2 - [intros.apply False_ind.apply (not_le_Sn_O ? H) - |intros.apply False_ind.apply (not_le_Sn_O ? H1) - |intros.simplify.unfold.apply le_S_S_to_le.assumption - ] -qed. - -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 *) -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 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 le_plus_minus_m_m: ∀n,m:nat. n ≤ (n-m)+m. +#n (elim n) // #a #Hind #m (cases m) // normalize #n/2/ +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 -] +theorem le_plus_to_minus_r: ∀a,b,c. a + b ≤ c → a ≤ c -b. +#a #b #c #H @(le_plus_to_le_r … b) /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. *) +(* lt *) -(* -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 transitive_lt: transitive nat lt. +#a #b #c #ltab #ltbc (elim ltbc) /2/ +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 le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p. +#n #m #p #H (elim H) /3/ qed. + +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 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. -*) +(* 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. @@ -920,10 +549,35 @@ qed. theorem le_plus_to_minus: ∀n,m,p. n ≤ p+m → n-m ≤ p. #n #m #p #lep /2/ qed. -theorem le_plus_to_minus_r: ∀a,b,c. a + b ≤ c → a ≤ c -b. -#a #b #c #H @(le_plus_to_le_r … b) /2/ +theorem monotonic_le_minus_r: +∀p,q,n:nat. q ≤ p → n-p ≤ n-q. +#p #q #n #lepq @le_plus_to_minus +@(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. + +(* lt *) + +theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n(commutative_plus p) ->associative_plus (S_pred m) + [ apply le_S_S + assumption + | assumption + ] +]. +qed. + +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. + +(* other abstract properties *) +theorem antisymmetric_le : antisymmetric nat le. +unfold antisymmetric.intros 2. +apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))). +intros.apply le_n_O_to_eq.assumption. +intros.apply False_ind.apply (not_le_Sn_O ? H). +intros.apply eq_f.apply H. +apply le_S_S_to_le.assumption. +apply le_S_S_to_le.assumption. +qed. + +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. + +theorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m +≝ monotonic_le_plus_r. + +theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p +\def monotonic_le_plus_l. + +variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def +monotonic_lt_plus_r. + +variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def +monotonic_lt_plus_l. + +theorem le_to_lt_to_lt_plus: ∀a,b,c,d:nat. +a ≤ c → b < d → a + b < c+d. +(* bello /2/ un po' lento *) +#a #b #c #d #leac #lebd +normalize napplyS le_plus // qed. + +theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m +\def monotonic_le_times_r. + +theorem monotonic_le_times_l: +∀m:nat.monotonic nat le (λn.n*m). +/2/ qed. + +theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p +\def monotonic_le_times_l. + +theorem le_S_times_2: ∀n,m.O < m → n ≤ m → S n ≤ 2*m. +#n #m #posm #lenm (* interessante *) +applyS (le_plus n m) // qed. + +theorem lt_O_times_S_S: ∀n,m:nat.O < (S n)*(S m). +intros.simplify.unfold lt.apply le_S_S.apply le_O_n. +qed. + +theorem lt_times_eq_O: \forall a,b:nat. +O < a → a * b = O → b = O. +intros. +apply (nat_case1 b) +[ intros. + reflexivity +| intros. + rewrite > 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. + +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 nat_compare_times_l : \forall n,p,q:nat. +nat_compare p q = nat_compare ((S n) * p) ((S n) * q). +intros.apply nat_compare_elim.intro. +apply nat_compare_elim. +intro.reflexivity. +intro.absurd (p=q). +apply (inj_times_r n).assumption. +apply lt_to_not_eq. assumption. +intro.absurd (q 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. + +theorem plus_minus: ∀n,m,p. p ≤ m → (n+m)-p = n +(m-p). +#n #m #p #lepm @plus_to_minus >(commutative_plus p) +>associative_plus