X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=9c420f9ec492e532cebae13dcd72fe76b5087c84;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=3152b0dbe3503aeee43bbd0f872b847f87979239;hpb=a79bf6edc13daaea8135ca71fdc92e02e229f030;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 3152b0dbe..9c420f9ec 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -12,12 +12,11 @@ (* *) (**************************************************************************) -(* include "higher_order_defs/functions.ma". *) include "hints_declaration.ma". include "basics/functions.ma". -include "basics/eq.ma". +include "basics/eq.ma". -ninductive nat : Type[0] ≝ +ninductive nat : Type ≝ | O : nat | S : nat → nat. @@ -36,7 +35,7 @@ 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. @@ -49,18 +48,19 @@ ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. //. nqed. *) ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m. -/2/; nqed. +/3/; nqed. ndefinition not_zero: nat → Prop ≝ λn: nat. match n with [ O ⇒ False | (S p) ⇒ True ]. - + ntheorem not_eq_O_S : ∀n:nat. O ≠ S n. -#n; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //. +#n; napply nmk; #eqOS; nchange with (not_zero O); +nrewrite > eqOS; //. nqed. -ntheorem not_eq_n_Sn : ∀n:nat. n ≠ S n. -#n; nelim n; /2/; nqed. +ntheorem not_eq_n_Sn: ∀n:nat. n ≠ S n. +#n; nelim n;/2/; nqed. ntheorem nat_case: ∀n:nat.∀P:nat → Prop. @@ -76,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). +ntheorem decidable_eq_nat : ∀n,m:nat.decidable (n=m). napply nat_elim2; #n; ##[ ncases n; /2/; ##| /3/; - ##| #m; #Hind; ncases Hind; /3/; + ##| #m; #Hind; ncases Hind;/3/; ##] nqed. @@ -112,9 +112,10 @@ 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. *) +(* deleterio? +ntheorem plus_n_1 : ∀n:nat. S n = n+1. +//; nqed. +*) ntheorem symmetric_plus: symmetric ? plus. #n; nelim n; nnormalize; //; nqed. @@ -123,7 +124,7 @@ ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. ntheorem assoc_plus1: ∀a,b,c. c + (b + a) = b + c + a. -//; nqed. +//; nqed. ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). #n; nelim n; nnormalize; /3/; nqed. @@ -159,7 +160,7 @@ ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). #n; nelim n; nnormalize; //; nqed. ntheorem symmetric_times : symmetric nat times. -#n; nelim n; nnormalize; //; nqed. +#n; nelim n; nnormalize; //; nqed. (* variant sym_times : \forall n,m:nat. n*m = m*n \def symmetric_times. *) @@ -167,9 +168,9 @@ 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 distributive_times_plus_r : + ∀a,b,c:nat. (b+c)*a = b*a + c*a. +//; nqed. ntheorem associative_times: associative nat times. #n; nelim n; nnormalize; //; nqed. @@ -220,13 +221,16 @@ 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. +(* nlemma eq_lt: ∀n,m. (n < m) = (S n ≤ m). +//; nqed. *) + +ndefinition ge: nat → nat → Prop ≝ +λn,m:nat.m ≤ 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 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. *) + +(* +ntheorem 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. *) + +ntheorem monotonic_lt_times_l: + ∀c:nat. O < c → monotonic nat lt (λt.(t*c)). +#c; #posc; #n; #m; #ltnm; +nelim ltnm; nnormalize; + ##[/2/; + ##|#a; #_; #lt1; napply (transitive_le ??? lt1);//; + ##] +nqed. + +ntheorem monotonic_lt_times_r: + ∀c:nat. O < c → monotonic nat lt (λt.(c*t)). +/2/; 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