X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fnlibrary%2Farithmetics%2Fnat.ma;h=a1e6fbb230c6f7cd93e444aa96a4fc681bfbacd1;hb=661403facfb7ca53b58635a95904787ae393bde5;hp=189f2f1524598296296b945f3a4841685f05acdb;hpb=b913b3ee973dfc9f1e6d7c0fd0c720d03699689e;p=helm.git diff --git a/helm/software/matita/nlibrary/arithmetics/nat.ma b/helm/software/matita/nlibrary/arithmetics/nat.ma index 189f2f152..a1e6fbb23 100644 --- a/helm/software/matita/nlibrary/arithmetics/nat.ma +++ b/helm/software/matita/nlibrary/arithmetics/nat.ma @@ -13,6 +13,7 @@ (**************************************************************************) (* include "higher_order_defs/functions.ma". *) +include "hints_declaration.ma". include "basics/functions.ma". include "basics/eq.ma". @@ -22,9 +23,17 @@ ninductive nat : Type[0] ≝ interpretation "Natural numbers" 'N = nat. -default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind. +alias num (instance 0) = "nnatural number". -alias num (instance 0) = "natural number". +(* +nrecord pos : Type ≝ + {n:>nat; is_pos: n ≠ 0}. + +ncoercion nat_to_pos: ∀n:nat. n ≠0 →pos ≝ mk_pos on +*) + +(* default "natural numbers" cic:/matita/ng/arithmetics/nat/nat.ind. +*) ndefinition pred ≝ λn. match n with @@ -35,8 +44,7 @@ ntheorem pred_Sn : ∀n. n = pred (S n). //. nqed. ntheorem injective_S : injective nat nat S. -(* nwhd; #a; #b;#H;napplyS pred_Sn. *) -//. nqed. +//; nqed. (* ntheorem inj_S : \forall n,m:nat.(S n)=(S m) \to n=m. @@ -88,32 +96,46 @@ nlet rec plus n m ≝ interpretation "natural plus" 'plus x y = (plus x y). -ntheorem plus_n_O: ∀n:nat. n = n+O. -#n; nelim n; /2/; nqed. +ntheorem plus_O_n: ∀n:nat. n = 0+n. +//; nqed. -ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n+(S m). +(* +ntheorem plus_Sn_m: ∀n,m:nat. S (n + m) = S n + m. +//; nqed. +*) + +ntheorem plus_n_O: ∀n:nat. n = n+0. #n; nelim n; nnormalize; //; nqed. -ntheorem plus_n_SO : ∀n:nat. S n = n+(S O). -//; nqed. +ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n + S m. +#n; nelim n; nnormalize; //; nqed. -ntheorem sym_plus: ∀n,m:nat. n+m = m+n. +(* +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. *) + +ntheorem symmetric_plus: symmetric ? plus. +#n; nelim n; nnormalize; //; nqed. ntheorem associative_plus : associative nat plus. #n; nelim n; nnormalize; //; nqed. -(* ntheorem assoc_plus : \forall n,m,p:nat. (n+m)+p = n+(m+p) -\def associative_plus. *) +ntheorem assoc_plus1: ∀a,b,c. b + (a + c) = a + (b + c). +//; nqed. ntheorem injective_plus_r: ∀n:nat.injective nat nat (λm.n+m). #n; nelim n; nnormalize; /3/; nqed. (* ntheorem inj_plus_r: \forall p,n,m:nat. p+n = p+m \to n=m -\def injective_plus_r. *) +\def injective_plus_r. ntheorem injective_plus_l: ∀m:nat.injective nat nat (λn.n+m). -/2/; nqed. +/2/; nqed. *) (* ntheorem inj_plus_l: \forall p,n,m:nat. n+p = m+p \to n=m \def injective_plus_l. *) @@ -130,103 +152,59 @@ interpretation "natural times" 'times x y = (times x y). ntheorem times_Sn_m: ∀n,m:nat. m+n*m = S n*m. //; nqed. +ntheorem times_O_n: ∀n:nat. O = O*n. +//; nqed. + ntheorem times_n_O: ∀n:nat. O = n*O. #n; nelim n; //; nqed. ntheorem times_n_Sm : ∀n,m:nat. n+(n*m) = n*(S m). -#n; nelim n; nnormalize; /2/; nqed. +#n; nelim n; nnormalize; //; nqed. + +ntheorem symmetric_times : symmetric nat times. +#n; nelim n; nnormalize; //; nqed. + +(* variant sym_times : \forall n,m:nat. n*m = m*n \def +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 associative_times: associative nat times. +#n; nelim n; nnormalize; //; nqed. +nlemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). +//; nqed. + +(* ci servono questi risultati? ntheorem times_O_to_O: ∀n,m:nat.n*m=O → n=O ∨ m=O. -napply nat_elim2; (* /2/ slow! *) - ##[#n; #H; @1; //; - ##|#n; #H; @2; //; (* ?? *) - ##|#n; #m; #H; #H1; napply False_ind;napply not_eq_O_S; - (* why two steps? *) /2/; - ##] -nqed. +napply nat_elim2; /2/; +#n; #m; #H; nnormalize; #H1; napply False_ind;napply not_eq_O_S; +//; nqed. ntheorem times_n_SO : ∀n:nat. n = n * S O. -//; nqed. - -(* +#n; //; nqed. -ntheorem times_SSO_n : ∀n:nat. n + n = 2 * n. -intros. -simplify. -rewrite < plus_n_O. -reflexivity. -qed. +ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n. +nnormalize; //; nqed. -alias num (instance 0) = "natural number". -lemma times_SSO: \forall n.2*(S n) = S(S(2*n)). -intro. -simplify.rewrite < plus_n_Sm.reflexivity. -qed. +nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)). +//; nqed. -theorem or_eq_eq_S: \forall n.\exists m. +ntheorem or_eq_eq_S: \forall n.\exists m. n = (S(S O))*m \lor n = S ((S(S O))*m). -intro.elim n - [apply (ex_intro ? ? O). - left.reflexivity - |elim H. elim H1 - [apply (ex_intro ? ? a). - right.apply eq_f.assumption - |apply (ex_intro ? ? (S a)). - left. rewrite > H2. - apply sym_eq. - apply times_SSO - ] - ] -qed. - -theorem symmetric_times : symmetric nat times. -unfold symmetric. -intros.elim x - [simplify.apply times_n_O. - | simplify.rewrite > H.apply times_n_Sm.] -qed. - -variant sym_times : \forall n,m:nat. n*m = m*n \def -symmetric_times. - -theorem distributive_times_plus : distributive nat times plus. -unfold distributive. -intros.elim x. -simplify.reflexivity. -simplify. -demodulate all. -(* -rewrite > H. rewrite > assoc_plus.rewrite > assoc_plus. -apply eq_f.rewrite < assoc_plus. rewrite < (sym_plus ? z). -rewrite > assoc_plus.reflexivity. *) -qed. - -variant distr_times_plus: \forall n,m,p:nat. n*(m+p) = n*m + n*p -\def distributive_times_plus. - -theorem associative_times: associative nat times. -unfold associative. -intros. -elim x. simplify.apply refl_eq. -simplify. -demodulate all. -(* -rewrite < sym_times. ->>>>>>> .r9770 -rewrite > distr_times_plus. -rewrite < sym_times. -rewrite < (sym_times (times n y) z). -rewrite < H.apply refl_eq. +#n; nelim n; + ##[@; /2/; + ##|#a; #H; nelim H; #b;#or;nelim or;#aeq; + ##[@ b; @ 2; //; + ##|@ (S b); @ 1; /2/; + ##] +nqed. *) -qed. - -variant assoc_times: \forall n,m,p:nat. (n*m)*p = n*(m*p) \def -associative_times. -lemma times_times: ∀x,y,z. x*(y*z) = y*(x*z). -intros. -demodulate. reflexivity. -(* autobatch paramodulation. *) -qed. +(************************** compare ****************************) -*) \ No newline at end of file