]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/dama/group.ma
excedence -> excess
[helm.git] / helm / software / matita / dama / group.ma
index a04cd3bcc17542189eb1d3a4ad7b84e57bfcc6f9..0e2668c2d71c58923390538a8753fff9868158df 100644 (file)
@@ -14,7 +14,7 @@
 
 set "baseuri" "cic:/matita/group/".
 
-include "excedence.ma".
+include "excess.ma".
 
 definition left_neutral ≝ λC:apartness.λop.λe:C. ∀x:C. op e x ≈ x.
 definition right_neutral ≝ λC:apartness.λop. λe:C. ∀x:C. op x e ≈ x.
@@ -138,6 +138,12 @@ apply (ap_rewl ??? (y + (x + -x)));
         apply (ap_rewr ??? z (zero_neutral ??));
         assumption]]
 qed.
+
+lemma applus: ∀E:abelian_group.∀x,a,y,b:E.x + a # y + b → x # y ∨ a # b.
+intros; cases (ap_cotransitive ??? (y+a) a1); [left|right]
+[apply (plus_cancr_ap ??? a)|apply (plus_cancl_ap ??? y)]
+assumption;
+qed.
     
 lemma plus_cancl: ∀G:abelian_group.∀y,z,x:G. x+y ≈ x+z → y ≈ z. 
 intros 6 (G y z x E Ayz); apply E; apply fap_plusl; assumption;