]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/nlibrary/arithmetics/nat.ma
Unfolded exact letins during proof reconstruction.
[helm.git] / helm / software / matita / nlibrary / arithmetics / nat.ma
index 189f2f1524598296296b945f3a4841685f05acdb..b6683b267801bac94d220ad5ca573ac644eaf1a5 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 (* include "higher_order_defs/functions.ma". *)
+include "hints_declaration.ma".
 include "basics/functions.ma".
 include "basics/eq.ma". 
 
@@ -22,21 +23,26 @@ 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
- [ O ⇒  O
- | (S p) ⇒ p].
+ λn. match n with [ O ⇒  O | (S p) ⇒ p].
 
 ntheorem pred_Sn : ∀n. n = pred (S n).
-//. nqed.
+//; 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.
@@ -47,8 +53,7 @@ ntheorem not_eq_S: ∀n,m:nat. n ≠ m → S n ≠ S m.
 
 ndefinition not_zero: nat → Prop ≝
  λn: nat. match n with
-  [ O ⇒ False
-  | (S p) ⇒ True ].
+  [ O ⇒ False | (S p) ⇒ True ].
 
 ntheorem not_eq_O_S : ∀n:nat. O ≠ S n.
 #n; #eqOS; nchange with (not_zero O); nrewrite > eqOS; //.
@@ -72,10 +77,10 @@ ntheorem nat_elim2 :
 #n0; #Rn0m; #m; ncases m;/2/; nqed.
 
 ntheorem decidable_eq_nat : \forall n,m:nat.decidable (n=m).
-napply nat_elim2;#n;
+napply nat_elim2; #n;
  ##[ ncases n; /2/;
  ##| /3/;
- ##| #m; #Hind; ncases Hind; /3/; (* ??? /2/;#neqnm; /3/; *)
+ ##| #m; #Hind; ncases Hind; /3/;
  ##]
 nqed. 
 
@@ -88,32 +93,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_Sn_m: ∀n,m:nat. S (n + m) = S n + m.
+//; nqed.
+*)
 
-ntheorem plus_n_Sm : ∀n,m:nat. S (n+m) = n+(S m).
+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. c + (b + a) = b + c + a.
+//; 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 +149,498 @@ 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.
+#n; //; nqed.
+
+ntheorem times_SSO_n : ∀n:nat. n + n = (S(S O)) * n.
+nnormalize; //; nqed.
+
+nlemma times_SSO: \forall n.(S(S O))*(S n) = S(S((S(S O))*n)).
 //; nqed.
 
+ntheorem or_eq_eq_S: \forall n.\exists m. 
+n = (S(S O))*m \lor n = S ((S(S O))*m).
+#n; nelim n;
+  ##[@; /2/;
+  ##|#a; #H; nelim H; #b;#or;nelim or;#aeq;
+    ##[@ b; @ 2; //;
+    ##|@ (S b); @ 1; /2/;
+    ##]
+nqed.
+*)
+
+(******************** ordering relations ************************)
+
+ninductive le (n:nat) : nat → Prop ≝
+  | le_n : le n n
+  | le_S : ∀ m:nat. le n m → le n (S m).
+
+interpretation "natural 'less or equal to'" 'leq x y = (le x y).
+
+interpretation "natural 'neither less nor equal to'" 'nleq x y = (Not (le x y)).
+
+ndefinition lt: nat → nat → Prop ≝
+λn,m:nat. S n ≤ m.
+
+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.
+
+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<n.
+
+interpretation "natural 'greater than'" 'gt x y = (gt x y).
+
+interpretation "natural 'not greater than'" 'ngtr x y = (Not (gt x y)).
+
+ntheorem transitive_le : transitive nat le.
+#a; #b; #c; #leab; #lebc;nelim lebc;/2/;
+nqed.
+
+(*
+ntheorem trans_le: \forall n,m,p:nat. n \leq m \to m \leq p \to n \leq p
+\def transitive_le. *)
+
+ntheorem transitive_lt: transitive nat lt.
+#a; #b; #c; #ltab; #ltbc;nelim ltbc;/2/;nqed.
+
 (*
+theorem trans_lt: \forall n,m,p:nat. lt n m \to lt m p \to lt n p
+\def transitive_lt. *)
 
-ntheorem times_SSO_n : ∀n:nat. n + n = 2 * n.
+ntheorem le_S_S: ∀n,m:nat. n ≤ m → S n ≤ S m.
+#n; #m; #lenm; nelim lenm; /2/; nqed.
+
+ntheorem le_O_n : ∀n:nat. O ≤ n.
+#n; nelim n; /2/; nqed.
+
+ntheorem le_n_Sn : ∀n:nat. n ≤ S n.
+/2/; nqed.
+
+ntheorem le_pred_n : ∀n:nat. pred n ≤ n.
+#n; nelim n; //; nqed.
+
+ntheorem monotonic_pred: monotonic ? le pred.
+#n; #m; #lenm; nelim lenm; /2/; nqed.
+
+ntheorem le_S_S_to_le: ∀n,m:nat. S n ≤ S m → n ≤ m.
+/2/; nqed.
+
+ntheorem lt_S_S_to_lt: ∀n,m. S n < S m \to n < m.
+/2/; nqed. 
+
+ntheorem lt_to_lt_S_S: ∀n,m. n < m → S n < S m.
+/2/; nqed.
+
+ntheorem lt_to_not_zero : ∀n,m:nat. n < m → not_zero m.
+#n; #m; #Hlt; nelim Hlt;//; nqed.
+
+(* lt vs. le *)
+ntheorem not_le_Sn_O: ∀ n:nat. S n ≰ O.
+#n; #Hlen0; napply (lt_to_not_zero ?? Hlen0); nqed.
+
+ntheorem not_le_to_not_le_S_S: ∀ n,m:nat. n ≰ m → S n ≰ S m.
+/3/; nqed.
+
+ntheorem not_le_S_S_to_not_le: ∀ n,m:nat. S n ≰ S m → n ≰ m.
+/3/; nqed.
+
+ntheorem decidable_le: ∀n,m. decidable (n≤m).
+napply nat_elim2; #n; /2/;
+#m; #dec; ncases dec;/3/; nqed.
+
+ntheorem decidable_lt: ∀n,m. decidable (n < m).
+#n; #m; napply decidable_le ; nqed.
+
+ntheorem not_le_Sn_n: ∀n:nat. S n ≰ n.
+#n; nelim n; /2/; nqed.
+
+ntheorem lt_S_to_le: ∀n,m:nat. n < S m → n ≤ m.
+/2/; nqed.
+
+ntheorem not_le_to_lt: ∀n,m. n ≰ m → m < n.
+napply nat_elim2; #n;
+ ##[#abs; napply False_ind;/2/;
+ ##|/2/;
+ ##|#m;#Hind;#HnotleSS; napply lt_to_lt_S_S;/3/;
+ ##]
+nqed.
+
+ntheorem lt_to_not_le: ∀n,m. n < m → m ≰ n.
+#n; #m; #Hltnm; nelim Hltnm;/3/; nqed.
+
+ntheorem not_lt_to_le: ∀n,m:nat. n ≮ m → m ≤ n.
+#n; #m; #Hnlt; napply lt_S_to_le;
+(* something strange here: /2/ fails *)
+napply not_le_to_lt; napply Hnlt; nqed. 
+
+ntheorem le_to_not_lt: ∀n,m:nat. n ≤ m → m ≮ n.
+/2/; nqed.
+
+(* lt and le trans *)
+
+ntheorem lt_to_le_to_lt: ∀n,m,p:nat. n < m → m ≤ p → n < p.
+#n; #m; #p; #H; #H1; nelim H1; /2/; nqed.
+
+ntheorem le_to_lt_to_lt: ∀n,m,p:nat. n ≤ m → m < p → n < p.
+#n; #m; #p; #H; nelim H; /3/; nqed.
+
+ntheorem lt_S_to_lt: ∀n,m. S n < m → n < m.
+/2/; nqed.
+
+ntheorem ltn_to_ltO: ∀n,m:nat. n < m → O < m.
+/2/; nqed.
+
+(*
+theorem lt_SO_n_to_lt_O_pred_n: \forall n:nat.
+(S O) \lt n \to O \lt (pred n).
 intros.
-simplify.
-rewrite < plus_n_O.
-reflexivity.
+apply (ltn_to_ltO (pred (S O)) (pred n) ?).
+ apply (lt_pred (S O) n);
+ [ apply (lt_O_S O) 
+ | assumption
+ ]
+qed. *)
+
+ntheorem lt_O_n_elim: ∀n:nat. O < n → 
+  ∀P:nat → Prop.(∀m:nat.P (S m)) → P n.
+#n; nelim n; //; #abs; napply False_ind; /2/; nqed.
+
+(*
+theorem lt_pred: \forall n,m. 
+  O < n \to n < m \to pred n < pred m. 
+apply nat_elim2
+  [intros.apply False_ind.apply (not_le_Sn_O ? H)
+  |intros.apply False_ind.apply (not_le_Sn_O ? H1)
+  |intros.simplify.unfold.apply le_S_S_to_le.assumption
+  ]
 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.
+theorem S_pred: \forall n:nat.lt O n \to eq nat n (S (pred n)).
+intro.elim n.apply False_ind.exact (not_le_Sn_O O H).
+apply eq_f.apply pred_Sn.
 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
-    ]
+theorem le_pred_to_le:
+ ∀n,m. O < m → pred n ≤ pred m → n ≤ m.
+intros 2;
+elim n;
+[ apply le_O_n
+| simplify in H2;
+  rewrite > (S_pred m);
+  [ apply le_S_S;
+    assumption
+  | assumption
   ]
+].
 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.
+(* le to lt or eq *)
+ntheorem le_to_or_lt_eq: ∀n,m:nat. n ≤ m → n < m ∨ n = m.
+#n; #m; #lenm; nelim lenm; /3/; nqed.
+
+(* not eq *)
+ntheorem lt_to_not_eq : ∀n,m:nat. n < m → n ≠ m.
+/2/; nqed.
+
+(*not lt 
+ntheorem eq_to_not_lt: ∀a,b:nat. a = b → a ≮ b.
+intros.
+unfold Not.
+intros.
+rewrite > H in H1.
+apply (lt_to_not_eq b b)
+[ assumption
+| reflexivity
+]
+qed. 
+
+theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n.
+intros;
+unfold Not;
+intro;
+unfold lt in H;
+unfold lt in H1;
+generalize in match (le_S_S ? ? H);
+intro;
+generalize in match (transitive_le ? ? ? H2 H1);
+intro;
+apply (not_le_Sn_n ? H3).
+qed. *)
+
+ntheorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n<m.
+#n; #m; #Hneq; #Hle; ncases (le_to_or_lt_eq ?? Hle); //;
+#Heq; nelim (Hneq Heq); nqed.
+
+(* le elimination *)
+ntheorem le_n_O_to_eq : ∀n:nat. n ≤ O → O=n.
+#n; ncases n; //; #a ; #abs; nelim (not_le_Sn_O ? abs); nqed.
+
+ntheorem le_n_O_elim: ∀n:nat. n ≤ O → ∀P: nat →Prop. P O → P n.
+#n; ncases n; //; #a; #abs; nelim (not_le_Sn_O ? abs); nqed. 
+
+ntheorem le_n_Sm_elim : ∀n,m:nat.n ≤ S m → 
+∀P:Prop. (S n ≤ S m → P) → (n=S m → P) → P.
+#n; #m; #Hle; #P; nelim Hle; /3/; nqed.
+
+(* le and eq *)
+
+ntheorem le_to_le_to_eq: ∀n,m. n ≤ m → m ≤ n → n = m.
+napply nat_elim2; /3/; nqed.
+
+ntheorem lt_O_S : \forall n:nat. O < S n.
+/2/; nqed.
 
-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. *)
+(* other abstract properties *)
+theorem antisymmetric_le : antisymmetric nat le.
+unfold antisymmetric.intros 2.
+apply (nat_elim2 (\lambda n,m.(n \leq m \to m \leq n \to n=m))).
+intros.apply le_n_O_to_eq.assumption.
+intros.apply False_ind.apply (not_le_Sn_O ? H).
+intros.apply eq_f.apply H.
+apply le_S_S_to_le.assumption.
+apply le_S_S_to_le.assumption.
 qed.
 
-variant distr_times_plus: \forall n,m,p:nat. n*(m+p) = n*m + n*p
-\def distributive_times_plus.
+theorem antisym_le: \forall n,m:nat. n \leq m \to m \leq n \to n=m
+\def antisymmetric_le.
 
-theorem associative_times: associative nat times.
-unfold associative.
-intros.
-elim x. simplify.apply refl_eq. 
-simplify.
-demodulate all.
+theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m.
+intros;
+unfold lt in H1;
+generalize in match (le_S_S_to_le ? ? H1);
+intro;
+apply antisym_le;
+assumption.
+qed.
+*)
+
+(* well founded induction principles *)
+
+ntheorem nat_elim1 : ∀n:nat.∀P:nat → Prop. 
+(∀m.(∀p. p < m → P p) → P m) → P n.
+#n; #P; #H; 
+ncut (∀q:nat. q ≤ n → P q);/2/;
+nelim n; 
+ ##[#q; #HleO; (* applica male *) 
+    napply (le_n_O_elim ? HleO);
+    napply H; #p; #ltpO;
+    napply False_ind; /2/; 
+ ##|#p; #Hind; #q; #HleS; 
+    napply H; #a; #lta; napply Hind;
+    napply le_S_S_to_le;/2/;
+ ##]
+nqed.
+
+(* some properties of functions *)
 (*
-rewrite < sym_times.
->>>>>>> .r9770
-rewrite > distr_times_plus.
-rewrite < sym_times.
-rewrite < (sym_times (times n y) z).
-rewrite < H.apply refl_eq.
+definition increasing \def \lambda f:nat \to nat. 
+\forall n:nat. f n < f (S n).
+
+theorem increasing_to_monotonic: \forall f:nat \to nat.
+increasing f \to monotonic nat lt f.
+unfold monotonic.unfold lt.unfold increasing.unfold lt.intros.elim H1.apply H.
+apply (trans_le ? (f n1)).
+assumption.apply (trans_le ? (S (f n1))).
+apply le_n_Sn.
+apply H.
+qed.
+
+theorem le_n_fn: \forall f:nat \to nat. (increasing f) 
+\to \forall n:nat. n \le (f n).
+intros.elim n.
+apply le_O_n.
+apply (trans_le ? (S (f n1))).
+apply le_S_S.apply H1.
+simplify in H. unfold increasing in H.unfold lt in H.apply H.
+qed.
+
+theorem increasing_to_le: \forall f:nat \to nat. (increasing f) 
+\to \forall m:nat. \exists i. m \le (f i).
+intros.elim m.
+apply (ex_intro ? ? O).apply le_O_n.
+elim H1.
+apply (ex_intro ? ? (S a)).
+apply (trans_le ? (S (f a))).
+apply le_S_S.assumption.
+simplify in H.unfold increasing in H.unfold lt in H.
+apply H.
+qed.
+
+theorem increasing_to_le2: \forall f:nat \to nat. (increasing f) 
+\to \forall m:nat. (f O) \le m \to 
+\exists i. (f i) \le m \land m <(f (S i)).
+intros.elim H1.
+apply (ex_intro ? ? O).
+split.apply le_n.apply H.
+elim H3.elim H4.
+cut ((S n1) < (f (S a)) \lor (S n1) = (f (S a))).
+elim Hcut.
+apply (ex_intro ? ? a).
+split.apply le_S. assumption.assumption.
+apply (ex_intro ? ? (S a)).
+split.rewrite < H7.apply le_n.
+rewrite > H7.
+apply H.
+apply le_to_or_lt_eq.apply H6.
+qed.
 *)
+
+(******************* monotonicity ******************************)
+ntheorem monotonic_le_plus_r: 
+∀n:nat.monotonic nat le (λm.n + m).
+#n; #a; #b; nelim n; nnormalize; //;
+#m; #H; #leab;napply le_S_S; /2/; nqed.
+
+ntheorem le_plus_r: ∀p,n,m:nat. n ≤ m → p + n ≤ p + m
+≝ monotonic_le_plus_r.
+
+ntheorem monotonic_le_plus_l: 
+∀m:nat.monotonic nat le (λn.n + m).
+/2/; nqed.
+
+ntheorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
+\def monotonic_le_plus_l. 
+
+ntheorem le_plus: ∀n1,n2,m1,m2:nat. n1 ≤ n2  \to m1 ≤ m2 
+→ n1 + m1 ≤ n2 + m2.
+#n1; #n2; #m1; #m2; #len; #lem; napply transitive_le;
+/2/; nqed.
+
+ntheorem le_plus_n :∀n,m:nat. m ≤ n + m.
+/2/; nqed. 
+
+ntheorem le_plus_n_r :∀n,m:nat. m ≤ m + n.
+/2/; nqed.
+
+ntheorem eq_plus_to_le: ∀n,m,p:nat.n=m+p → m ≤ n.
+//; nqed.
+
+ntheorem le_plus_to_le: ∀a,n,m. a + n ≤ a + m → n ≤ m.
+#a; nelim a; /3/; nqed.
+
+(* times 
+theorem monotonic_le_times_r: 
+\forall n:nat.monotonic nat le (\lambda m. n * m).
+simplify.intros.elim n.
+simplify.apply le_O_n.
+simplify.apply le_plus.
+assumption.
+assumption.
+qed.
+
+theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m
+\def monotonic_le_times_r.
+
+theorem monotonic_le_times_l: 
+\forall m:nat.monotonic nat le (\lambda n.n*m).
+simplify.intros.
+rewrite < sym_times.rewrite < (sym_times m).
+apply le_times_r.assumption.
 qed.
 
-variant assoc_times: \forall n,m,p:nat. (n*m)*p = n*(m*p) \def
-associative_times.
+theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p
+\def monotonic_le_times_l.
 
-lemma times_times: ∀x,y,z. x*(y*z) = y*(x*z).
-intros. 
-demodulate. reflexivity.
-(* autobatch paramodulation. *)
+theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2  \to m1 \le m2 
+\to n1*m1 \le n2*m2.
+intros.
+apply (trans_le ? (n2*m1)).
+apply le_times_l.assumption.
+apply le_times_r.assumption.
 qed.
 
-*)
\ No newline at end of file
+theorem le_times_n: \forall n,m:nat.(S O) \le n \to m \le n*m.
+intros.elim H.simplify.
+elim (plus_n_O ?).apply le_n.
+simplify.rewrite < sym_plus.apply le_plus_n.
+qed.
+
+theorem le_times_to_le: 
+\forall a,n,m. S O \le a \to a * n \le a * m \to n \le m.
+intro.
+apply nat_elim2;intros
+  [apply le_O_n
+  |apply False_ind.
+   rewrite < times_n_O in H1.
+   generalize in match H1.
+   apply (lt_O_n_elim ? H).
+   intros.
+   simplify in H2.
+   apply (le_to_not_lt ? ? H2).
+   apply lt_O_S
+  |apply le_S_S.
+   apply H
+    [assumption
+    |rewrite < times_n_Sm in H2.
+     rewrite < times_n_Sm in H2.
+     apply (le_plus_to_le a).
+     assumption
+    ]
+  ]
+qed.
+
+theorem le_S_times_SSO: \forall n,m.O < m \to
+n \le m \to S n \le (S(S O))*m.
+intros.
+simplify.
+rewrite > plus_n_O.
+simplify.rewrite > plus_n_Sm.
+apply le_plus
+  [assumption
+  |rewrite < plus_n_O.
+   assumption
+  ]
+qed.
+(*0 and times *)
+theorem O_lt_const_to_le_times_const:  \forall a,c:nat.
+O \lt c \to a \le a*c.
+intros.
+rewrite > (times_n_SO a) in \vdash (? % ?).
+apply le_times
+[ apply le_n
+| assumption
+]
+qed. *)
\ No newline at end of file