]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/nat/times.ma
Committing all the recent development of Andrea after the merge between his
[helm.git] / helm / matita / library / nat / times.ma
index d2ce51bfa83dd9605bf5160e227edf3c532aeff0..2707c2ba1cdd8ff08a0404fa40de9f714390542c 100644 (file)
@@ -1,5 +1,5 @@
 (**************************************************************************)
-(*       ___                                                               *)
+(*       __                                                               *)
 (*      ||M||                                                             *)
 (*      ||A||       A project by Andrea Asperti                           *)
 (*      ||T||                                                             *)
 
 set "baseuri" "cic:/matita/nat/times".
 
-include "logic/equality.ma".
-include "nat/nat.ma".
 include "nat/plus.ma".
 
 let rec times n m \def 
  match n with 
  [ O \Rightarrow O
- | (S p) \Rightarrow (m+(times p m)) ].
+ | (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).
@@ -33,7 +31,7 @@ simplify.assumption.
 qed.
 
 theorem times_n_Sm : 
-\forall n,m:nat.n+n*m = n*(S m).
+\forall n,m:nat. n+(n*m) = n*(S m).
 intros.elim n.
 simplify.reflexivity.
 simplify.apply eq_f.rewrite < H.
@@ -45,16 +43,24 @@ reflexivity.
 apply assoc_plus.
 qed.
 
-(* same problem with symmetric: see plus 
-theorem symmetric_times : symmetric nat times. *)
+theorem times_n_SO : \forall n:nat. eq nat n (times n (S O)).
+intros.
+rewrite < times_n_Sm.
+rewrite < times_n_O.
+rewrite < plus_n_O.
+reflexivity.
+qed.
 
-theorem sym_times : 
-\forall n,m:nat.n*m = m*n.
-intros.elim n.
+theorem symmetric_times : symmetric nat times. 
+simplify.
+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.
 simplify.
 intros.elim x.
@@ -64,6 +70,18 @@ apply eq_f.rewrite < assoc_plus. rewrite < sym_plus ? z.
 rewrite > assoc_plus.reflexivity.
 qed.
 
-variant times_plus_distr: \forall n,m,p:nat. n*(m+p)=n*m+n*p
+variant times_plus_distr: \forall n,m,p:nat. n*(m+p) = n*m + n*p
 \def distributive_times_plus.
 
+theorem associative_times: associative nat times.
+simplify.intros.
+elim x.simplify.apply refl_eq.
+simplify.rewrite < sym_times.
+rewrite > times_plus_distr.
+rewrite < sym_times.
+rewrite < sym_times (times n y) z.
+rewrite < H.apply refl_eq.
+qed.
+
+variant assoc_times: \forall n,m,p:nat. (n*m)*p = n*(m*p) \def
+associative_times.