]> 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 0d682a268a9a69af3af26d8f36a99656d38a3952..0e2668c2d71c58923390538a8753fff9868158df 100644 (file)
 
 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.
 definition left_inverse ≝ λC:apartness.λop.λe:C.λinv:C→C. ∀x:C. op (inv x) x ≈ e.
 definition right_inverse ≝ λC:apartness.λop.λe:C.λ inv: C→ C. ∀x:C. op x (inv x) ≈ e. 
-definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y.
 (* ALLOW DEFINITION WITH SOME METAS *)
 
 definition distributive_left ≝
@@ -84,6 +83,14 @@ 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;
 apply (plus_strong_ext ???? A2);
 qed.
+
+lemma plus_cancl_ap: ∀G:abelian_group.∀x,y,z:G.z+x # z + y → x # y.
+intros; apply plus_strong_ext; assumption;
+qed.
+
+lemma plus_cancr_ap: ∀G:abelian_group.∀x,y,z:G.x+z # y+z → x # y.
+intros; apply plus_strong_extr; assumption;
+qed.
    
 lemma feq_plusr: ∀G:abelian_group.∀x,y,z:G. y ≈ z →  y+x ≈ z+x.
 intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_extr ? x));
@@ -131,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;
@@ -143,37 +156,38 @@ qed.
 theorem eq_opp_plus_plus_opp_opp: 
   ∀G:abelian_group.∀x,y:G. -(x+y) ≈ -x + -y.
 intros (G x y); apply (plus_cancr ??? (x+y));
-apply (eq_trans ?? 0 ? (opp_inverse ??));
-apply (eq_trans ?? (-x + -y + x + y)); [2: apply (eq_sym ??? (plus_assoc ????))]
-apply (eq_trans ?? (-y + -x + x + y)); [2: repeat apply feq_plusr; apply plus_comm]
-apply (eq_trans ?? (-y + (-x + x) + y)); [2: apply feq_plusr; apply plus_assoc;]
-apply (eq_trans ?? (-y + 0 + y)); 
+apply (Eq≈ 0 (opp_inverse ??));
+apply (Eq≈ (-x + -y + x + y)); [2: apply (eq_sym ??? (plus_assoc ????))]
+apply (Eq≈ (-y + -x + x + y)); [2: repeat apply feq_plusr; apply plus_comm]
+apply (Eq≈ (-y + (-x + x) + y)); [2: apply feq_plusr; apply plus_assoc;]
+apply (Eq≈ (-y + 0 + y)); 
   [2: apply feq_plusr; apply feq_plusl; apply eq_sym; apply opp_inverse]
-apply (eq_trans ?? (-y + y)); 
+apply (Eq≈ (-y + y)); 
   [2: apply feq_plusr; apply eq_sym; 
-      apply (eq_trans ?? (0+-y)); [apply plus_comm|apply zero_neutral]]
+      apply (Eq≈ (0+-y)); [apply plus_comm|apply zero_neutral]]
 apply eq_sym; apply opp_inverse.
 qed.
 
 theorem eq_opp_opp_x_x: ∀G:abelian_group.∀x:G.--x ≈ x.
 intros (G x); apply (plus_cancl ??? (-x));
-apply (eq_trans ?? (--x + -x)); [apply plus_comm]
-apply (eq_trans ?? 0); [apply opp_inverse]
+apply (Eq≈ (--x + -x) (plus_comm ???));
+apply (Eq≈ 0 (opp_inverse ??));
 apply eq_sym; apply opp_inverse;
 qed.
 
 theorem eq_zero_opp_zero: ∀G:abelian_group.0 ≈ -0. [assumption]
 intro G; apply (plus_cancr ??? 0);
-apply (eq_trans ?? 0); [apply zero_neutral;]
+apply (Eq≈ 0); [apply zero_neutral;]
 apply eq_sym; apply opp_inverse;
 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 ?z));
-apply (eq_trans ?? (-y + z) ? H2);
-apply (eq_trans ?? (-y + y) ? H1);
-apply (eq_trans ?? 0 ? (opp_inverse ??));
+(* apply (eq_trans ??? 0 ? (opp_inverse ??)); *)
+apply (Eq≈ 0 ? (opp_inverse ??));
+apply (Eq≈ (-y + z) H2);
+apply (Eq≈ (-y + y) H1);
+apply (Eq≈ 0 (opp_inverse ??));
 apply eq_reflexive;
 qed.
 
@@ -205,25 +219,3 @@ compose feq_plusl with feq_opp(H); apply H; assumption;
 qed.
 
 coercion cic:/matita/group/eq_opp_plusl.con nocomposites.
-
-lemma plus_cancr_ap: ∀G:abelian_group.∀x,y,z:G. x+z # y+z → x # y.
-intros (G x y z H); lapply (fap_plusr ? (-z) ?? H) as H1; clear H;
-lapply (ap_rewl ? (x + (z + -z)) ?? (plus_assoc ? x z (-z)) H1) as H2; clear H1;
-lapply (ap_rewl ? (x + (-z + z)) ?? (plus_comm ?z (-z)) H2) as H1; clear H2;
-lapply (ap_rewl ? (x + 0) ?? (opp_inverse ?z) H1) as H2; clear H1;
-lapply (ap_rewl ? (0+x) ?? (plus_comm ?x 0) H2) as H1; clear H2;
-lapply (ap_rewl ? x ?? (zero_neutral ?x) H1) as H2; clear H1;
-lapply (ap_rewr ? (y + (z + -z)) ?? (plus_assoc ? y z (-z)) H2) as H3;
-lapply (ap_rewr ? (y + (-z + z)) ?? (plus_comm ?z (-z)) H3) as H4;
-lapply (ap_rewr ? (y + 0) ?? (opp_inverse ?z) H4) as H5;
-lapply (ap_rewr ? (0+y) ?? (plus_comm ?y 0) H5) as H6;
-lapply (ap_rewr ? y ?? (zero_neutral ?y) H6);
-assumption;
-qed.
-
-lemma pluc_cancl_ap: ∀G:abelian_group.∀x,y,z:G. z+x # z+y → x # y.
-intros (G x y z H); apply (plus_cancr_ap ??? z);
-apply (ap_rewl ???? (plus_comm ???)); 
-apply (ap_rewr ???? (plus_comm ???));
-assumption;
-qed.