]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/library/nat/div_and_mod.ma
Much ado about nothing:
[helm.git] / matita / library / nat / div_and_mod.ma
index 65ba0e9e5127a77ed8c69bb13f50300a058e3213..537d3bcf0d11a4207f313f0103d069c1d9d87041 100644 (file)
@@ -215,11 +215,12 @@ qed.
 (* some properties of div and mod *)
 theorem div_times: \forall n,m:nat. ((S n)*m) / (S n) = m.
 intros.
-apply (div_mod_spec_to_eq ((S n)*m) (S n) ? ? ? O).
-goal 15. (* ?11 is closed with the following tactics *)
-apply div_mod_spec_div_mod.
-unfold lt.apply le_S_S.apply le_O_n.
-apply div_mod_spec_times.
+apply (div_mod_spec_to_eq ((S n)*m) (S n) ? ? ? O);
+[2: apply div_mod_spec_div_mod.
+    unfold lt.apply le_S_S.apply le_O_n.
+|   skip
+|   apply div_mod_spec_times
+]
 qed.
 
 theorem div_n_n: \forall n:nat. O < n \to n / n = S O.