X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Ftimes.ma;h=6bbdcec4d2083e0200a3bdb7b431d9d577f00997;hb=4dc47c9675ffd5fa50296ffaa9b5997501518c98;hp=24a2846d7a684eb337f932d502fcc6fb9dab206f;hpb=06a19bec47845ecffe3bf9d9a95d3d4dadf76861;p=helm.git diff --git a/helm/software/matita/library/nat/times.ma b/helm/software/matita/library/nat/times.ma index 24a2846d7..6bbdcec4d 100644 --- a/helm/software/matita/library/nat/times.ma +++ b/helm/software/matita/library/nat/times.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/nat/times". - include "nat/plus.ma". let rec times n m \def @@ -21,8 +19,12 @@ let rec times n m \def [ O \Rightarrow O | (S p) \Rightarrow m+(times p m) ]. -(*CSC: the URI must disappear: there is a bug now *) -interpretation "natural times" 'times x y = (cic:/matita/nat/times/times.con x y). +interpretation "natural times" 'times x y = (times x y). + +theorem times_Sn_m: +\forall n,m:nat. m+n*m = S n*m. +intros. reflexivity. +qed. theorem times_n_O: \forall n:nat. O = n*O. intros.elim n. @@ -68,10 +70,32 @@ rewrite < plus_n_O. reflexivity. qed. +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. + +theorem 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. +(* applyS times_n_Sm. *) simplify.rewrite > H.apply times_n_Sm. qed. @@ -91,8 +115,9 @@ 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. +unfold associative. +intros. +elim x. simplify.apply refl_eq. simplify.rewrite < sym_times. rewrite > distr_times_plus. rewrite < sym_times. @@ -102,3 +127,8 @@ 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.autobatch paramodulation. +qed. +