]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/div_and_mod.ma
Towards chebyshev.
[helm.git] / matita / library / nat / div_and_mod.ma
index 658b07b686eb5e6c23f02f3247c64f02f8f25381..0323b18fb17849949b5ce4a49e3b33476f14fcec 100644 (file)
@@ -293,6 +293,63 @@ constructor 1.
 assumption.reflexivity.
 qed.
 
+theorem mod_SO: \forall n:nat. mod n (S O) = O.
+intro.
+apply sym_eq.
+apply le_n_O_to_eq.
+apply le_S_S_to_le.
+apply lt_mod_m_m.
+apply le_n.
+qed.
+
+theorem div_SO: \forall n:nat. div n (S O) = n.
+intro.
+rewrite > (div_mod ? (S O)) in \vdash (? ? ? %)
+  [rewrite > mod_SO.
+   rewrite < plus_n_O.
+   apply times_n_SO
+  |apply le_n
+  ]
+qed.
+
+theorem le_div: \forall n,m. O < n \to m/n \le m.
+intros.
+rewrite > (div_mod m n) in \vdash (? ? %)
+  [apply (trans_le ? (m/n*n))
+    [rewrite > times_n_SO in \vdash (? % ?).
+     apply le_times
+      [apply le_n|assumption]
+    |apply le_plus_n_r
+    ]
+  |assumption
+  ]
+qed.
+
+theorem or_div_mod: \forall n,q. O < q \to
+((S (n \mod q)=q) \land S n = (S (div n q)) * q \lor
+((S (n \mod q)<q) \land S n= (div n q) * q + S (n\mod q))).
+intros.
+elim (le_to_or_lt_eq ? ? (lt_mod_m_m n q H))
+  [right.split
+    [assumption
+    |rewrite < plus_n_Sm.
+     apply eq_f.
+     apply div_mod.
+     assumption
+    ]
+  |left.split
+    [assumption
+    |simplify.
+     rewrite > sym_plus.
+     rewrite < H1 in ⊢ (? ? ? (? ? %)).
+     rewrite < plus_n_Sm.
+     apply eq_f.
+     apply div_mod.
+     assumption
+    ]
+  ]
+qed.
+
 (* injectivity *)
 theorem injective_times_r: \forall n:nat.injective nat nat (\lambda m:nat.(S n)*m).
 change with (\forall n,p,q:nat.(S n)*p = (S n)*q \to p=q).