]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/le_arith.ma
A few extensions for the moebius inversion theorem
[helm.git] / matita / library / nat / le_arith.ma
index ae386d313882f2b37903ae28c60c77dfa9fda072..90f3e182b2f51441b34c06054816a6c1cdcc9a82 100644 (file)
@@ -62,6 +62,16 @@ rewrite < sym_plus.
 apply le_plus_n.
 qed.
 
+theorem le_plus_to_le: 
+\forall a,n,m. a + n \le a + m \to n \le m.
+intro.
+elim a
+  [assumption
+  |apply H.
+   apply le_S_S_to_le.assumption
+  ]
+qed.
+
 (* times *)
 theorem monotonic_le_times_r: 
 \forall n:nat.monotonic nat le (\lambda m. n * m).
@@ -98,3 +108,27 @@ 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.
\ No newline at end of file