]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/divisible_group.ma
Bertrand's conjecture (weak), some work in progress
[helm.git] / helm / software / matita / dama / divisible_group.ma
index 6830ef4e1751fbf8ebffb24dcb2a1a0a896f3e7a..3a79b11bbd969a884307f4a6a86d08e63cbcfff8 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/divisible_group/".
+
 
 include "nat/orders.ma".
 include "group.ma".
@@ -50,8 +50,8 @@ qed.
 lemma div1: ∀G:dgroup.∀x:G.x/O ≈ x.
 intro G; cases G; unfold divide; intros; simplify;
 cases (f x O); simplify; simplify in H; intro; apply H;
-apply (ap_rewl ???? (plus_comm ???));
-apply (ap_rewl ???w (zero_neutral ??)); assumption;
+apply (Ap≪ ? (plus_comm ???));
+apply (Ap≪ w (zero_neutral ??)); assumption;
 qed.
 
 lemma apmul_ap: ∀G:dgroup.∀x,y:G.∀n.S n * x # S n * y → x # y.