]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/group.ma
many sed to use notation for rewriting
[helm.git] / helm / software / matita / dama / group.ma
index 9da386ef7bfefbb1797e86bcec43cddfb0c067c7..104dcf274e3943727d090dc8ff8fddd5f7bca59e 100644 (file)
@@ -80,7 +80,7 @@ coercion cic:/matita/group/feq_plusl.con nocomposites.
 lemma plus_strong_extr: ∀G:abelian_group.∀z:G.strong_ext ? (λx.x + z).
 intros 5 (G z x y A); simplify in A;
 lapply (plus_comm ? z x) as E1; lapply (plus_comm ? z y) as E2;
-lapply (ap_rewl ???? E1 A) as A1; lapply (ap_rewr ???? E2 A1) as A2;
+lapply (Ap≪ ? E1 A) as A1; lapply (Ap≫ ? E2 A1) as A2;
 apply (plus_strong_ext ???? A2);
 qed.
 
@@ -111,31 +111,31 @@ coercion cic:/matita/group/feq_plusl_sym_.con nocomposites.
       
 lemma fap_plusl: ∀G:abelian_group.∀x,y,z:G. y # z →  x+y # x+z. 
 intros (G x y z Ayz); apply (plus_strong_ext ? (-x));
-apply (ap_rewl ??? ((-x + x) + y));
+apply (Ap≪ ((-x + x) + y));
 [1: apply plus_assoc; 
-|2: apply (ap_rewr ??? ((-x +x) +z));
+|2: apply (Ap≫ ((-x +x) +z));
     [1: apply plus_assoc; 
-    |2: apply (ap_rewl ??? (0 + y));
+    |2: apply (Ap≪ (0 + y));
         [1: apply (feq_plusr ???? (opp_inverse ??)); 
-        |2: apply (ap_rewl ???? (zero_neutral ? y)); 
-            apply (ap_rewr ??? (0 + z) (opp_inverse ??)); 
-            apply (ap_rewr ???? (zero_neutral ??)); assumption;]]]
+        |2: apply (Ap≪ ? (zero_neutral ? y)); 
+            apply (Ap≫ (0 + z) (opp_inverse ??)); 
+            apply (Ap≫ ? (zero_neutral ??)); assumption;]]]
 qed.
 
 lemma fap_plusr: ∀G:abelian_group.∀x,y,z:G. y # z →  y+x # z+x. 
 intros (G x y z Ayz); apply (plus_strong_extr ? (-x));
-apply (ap_rewl ??? (y + (x + -x)));
+apply (Ap≪ (y + (x + -x)));
 [1: apply (eq_sym ??? (plus_assoc ????)); 
-|2: apply (ap_rewr ??? (z + (x + -x)));
+|2: apply (Ap≫ (z + (x + -x)));
     [1: apply (eq_sym ??? (plus_assoc ????)); 
-    |2: apply (ap_rewl ??? (y + (-x+x)) (plus_comm ? x (-x)));
-        apply (ap_rewl ??? (y + 0) (opp_inverse ??));
-        apply (ap_rewl ??? (0 + y) (plus_comm ???));
-        apply (ap_rewl ??? y (zero_neutral ??));
-        apply (ap_rewr ??? (z + (-x+x)) (plus_comm ? x (-x)));
-        apply (ap_rewr ??? (z + 0) (opp_inverse ??));
-        apply (ap_rewr ??? (0 + z) (plus_comm ???));
-        apply (ap_rewr ??? z (zero_neutral ??));
+    |2: apply (Ap≪ (y + (-x+x)) (plus_comm ? x (-x)));
+        apply (Ap≪ (y + 0) (opp_inverse ??));
+        apply (Ap≪ (0 + y) (plus_comm ???));
+        apply (Ap≪ y (zero_neutral ??));
+        apply (Ap≫ (z + (-x+x)) (plus_comm ? x (-x)));
+        apply (Ap≫ (z + 0) (opp_inverse ??));
+        apply (Ap≫ (0 + z) (plus_comm ???));
+        apply (Ap≫ z (zero_neutral ??));
         assumption]]
 qed.
 
@@ -183,7 +183,6 @@ qed.
 
 lemma feq_oppr: ∀G:abelian_group.∀x,y,z:G. y ≈ z → x ≈ -y → x ≈ -z.
 intros (G x y z H1 H2); apply (plus_cancr ??? z);
-(* apply (eq_trans ??? 0 ? (opp_inverse ??)); *)
 apply (Eq≈ 0 ? (opp_inverse ??));
 apply (Eq≈ (-y + z) H2);
 apply (Eq≈ (-y + y) H1);