]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/le_arith.ma
Towards chebyshev.
[helm.git] / matita / library / nat / le_arith.ma
index 90f3e182b2f51441b34c06054816a6c1cdcc9a82..f391c85a9686e77900e11bee558cd07987e3fbf8 100644 (file)
@@ -20,9 +20,10 @@ include "nat/orders.ma".
 (* plus *)
 theorem monotonic_le_plus_r: 
 \forall n:nat.monotonic nat le (\lambda m.n + m).
-simplify.intros.elim n.
-simplify.assumption.
-simplify.apply le_S_S.assumption.
+simplify.intros.elim n
+  [simplify.assumption.
+  |simplify.apply le_S_S.assumption
+  ]
 qed.
 
 theorem le_plus_r: \forall p,n,m:nat. n \le m \to p + n \le p + m
@@ -41,9 +42,21 @@ theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p
 theorem le_plus: \forall n1,n2,m1,m2:nat. n1 \le n2  \to m1 \le m2 
 \to n1 + m1 \le n2 + m2.
 intros.
+(**
+auto.
+*)
+apply (transitive_le (plus n1 m1) (plus n1 m2) (plus n2 m2) ? ?);
+  [apply (monotonic_le_plus_r n1 m1 m2 ?).
+   apply (H1).
+  |apply (monotonic_le_plus_l m2 n1 n2 ?).
+   apply (H).
+  ]
+(* end auto($Revision$) proof: TIME=0.61 SIZE=100 DEPTH=100 *)
+(*
 apply (trans_le ? (n2 + m1)).
 apply le_plus_l.assumption.
 apply le_plus_r.assumption.
+*)
 qed.
 
 theorem le_plus_n :\forall n,m:nat. m \le n + m.
@@ -131,4 +144,27 @@ apply nat_elim2;intros
      assumption
     ]
   ]
-qed.
\ No newline at end of file
+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.