X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Ftimes.ma;h=57dc60c6f2a071b5e25400161da998fa88f12763;hb=063523ae5f8da7e6458232f4afb6744ec86dc8bd;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..57dc60c6f 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,7 +19,6 @@ 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). theorem times_n_O: \forall n:nat. O = n*O. @@ -68,6 +65,27 @@ 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.