X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=b6683b267801bac94d220ad5ca573ac644eaf1a5;hb=42e022596d00f85eebb19b8f7c9029273ad15db1;hp=a1e6fbb230c6f7cd93e444aa96a4fc681bfbacd1;hpb=9f0b17f553522373b51fc4761d5a3e1e9dbbad73;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index a1e6fbb23..b6683b267 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -36,12 +36,10 @@ ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on *) 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. //; nqed. @@ -55,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; //. @@ -80,10 +77,10 @@ ntheorem nat_elim2 : #n0; #Rn0m; #m; ncases m;/2/; nqed. ntheorem decidable_eq_nat : \forall n,m:nat.decidable (n=m). -napply nat_elim2;#n; +napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; (* ??? /2/;#neqnm; /3/; *) + ##| #m; #Hind; ncases Hind; /3/; ##] nqed. @@ -125,7 +122,7 @@ ntheorem symmetric_plus: symmetric ? plus. ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. -ntheorem assoc_plus1: ∀a,b,c. b + (a + c) = a + (b + c). +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). @@ -206,5 +203,444 @@ n = (S(S O))*m \lor n = S ((S(S O))*m). nqed. *) -(************************** compare ****************************) +(******************** 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 (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 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. + +(* 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. + +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: +\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. + +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_times: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2 +\to n1*m1 \le n2*m2. +intros. +apply (trans_le ? (n2*m1)). +apply le_times_l.assumption. +apply le_times_r.assumption. +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. +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 + ] + ] +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 + ] +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