X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fdama%2Fgroups.ma;h=c00740ee957bd3b433de704cf029d876cc5b6341;hb=9791cd146bc0b8df953aee7bb8a3df60553b530c;hp=f676258172ad7f298aab4d25b20a2ed4eae0ebe7;hpb=eb9bc52a705fd347a5e1906ec32fd12f86507fe9;p=helm.git diff --git a/helm/software/matita/dama/groups.ma b/helm/software/matita/dama/groups.ma index f67625817..c00740ee9 100644 --- a/helm/software/matita/dama/groups.ma +++ b/helm/software/matita/dama/groups.ma @@ -36,10 +36,10 @@ record abelian_group : Type ≝ plus: carr → carr → carr; zero: carr; opp: carr → carr; - plus_assoc: associative ? plus (eq carr); - plus_comm: commutative ? plus (eq carr); - zero_neutral: left_neutral ? plus zero; - opp_inverse: left_inverse ? plus zero opp; + plus_assoc_: associative ? plus (eq carr); + plus_comm_: commutative ? plus (eq carr); + zero_neutral_: left_neutral ? plus zero; + opp_inverse_: left_inverse ? plus zero opp; plus_strong_ext: ∀z.strong_ext ? (plus z) }. @@ -60,6 +60,11 @@ definition minus ≝ interpretation "Abelian group minus" 'minus a b = (cic:/matita/groups/minus.con _ a b). +lemma plus_assoc: ∀G:abelian_group.∀x,y,z:G.x+(y+z)≈x+y+z ≝ plus_assoc_. +lemma plus_comm: ∀G:abelian_group.∀x,y:G.x+y≈y+x ≝ plus_comm_. +lemma zero_neutral: ∀G:abelian_group.∀x:G.0+x≈x ≝ zero_neutral_. +lemma opp_inverse: ∀G:abelian_group.∀x:G.-x+x≈0 ≝ opp_inverse_. + definition ext ≝ λA:apartness.λop:A→A. ∀x,y. x ≈ y → op x ≈ op y. lemma strong_ext_to_ext: ∀A:apartness.∀op:A→A. strong_ext ? op → ext ? op. @@ -71,7 +76,7 @@ intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_ext ? x)); assumption; qed. -coercion cic:/matita/groups/feq_plusl.con. +coercion cic:/matita/groups/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; @@ -85,8 +90,18 @@ intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_extr ? x)); assumption; qed. -coercion cic:/matita/groups/feq_plusr.con. - +coercion cic:/matita/groups/feq_plusr.con nocomposites. + +(* generation of coercions to make *_rew[lr] easier *) +lemma feq_plusr_sym_: ∀G:abelian_group.∀x,y,z:G.z ≈ y → y+x ≈ z+x. +compose feq_plusr with eq_sym (H); apply H; assumption; +qed. +coercion cic:/matita/groups/feq_plusr_sym_.con nocomposites. +lemma feq_plusl_sym_: ∀G:abelian_group.∀x,y,z:G.z ≈ y → x+y ≈ x+z. +compose feq_plusl with eq_sym (H); apply H; assumption; +qed. +coercion cic:/matita/groups/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)); @@ -95,25 +110,23 @@ apply (ap_rewl ??? ((-x + x) + y)); [1: apply plus_assoc; |2: apply (ap_rewl ??? (0 + y)); [1: apply (feq_plusr ???? (opp_inverse ??)); - |2: apply (ap_rewl ???? (zero_neutral ? y)); apply (ap_rewr ??? (0 + z)); - [1: apply (feq_plusr ???? (opp_inverse ??)); - |2: apply (ap_rewr ???? (zero_neutral ??)); assumption;]]]] + |2: apply (ap_rewl ???? (zero_neutral ? y)); + apply (ap_rewr ??? (0 + z) (opp_inverse ??)); + apply (ap_rewr ???? (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))); -[1: apply (eq_symmetric ??? (plus_assoc ????)); +[1: apply (eq_sym ??? (plus_assoc ????)); |2: apply (ap_rewr ??? (z + (x + -x))); - [1: apply (eq_symmetric ??? (plus_assoc ????)); - |2: apply (ap_rewl ??? (y + (-x+x)) (feq_plusl ???? (plus_comm ???))); - apply (ap_rewl ??? (y + 0) (feq_plusl ???? (opp_inverse ??))); + [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)) (feq_plusl ???? (plus_comm ???))); - apply (ap_rewr ??? (z + 0) (feq_plusl ???? (opp_inverse ??))); + 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 ??)); assumption]] @@ -130,27 +143,87 @@ 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_transitive ?? 0); [apply (opp_inverse ??)] -apply (eq_transitive ?? (-x + -y + x + y)); [2: apply (eq_symmetric ??? (plus_assoc ????))] -apply (eq_transitive ?? (-y + -x + x + y)); [2: repeat apply feq_plusr; apply plus_comm] -apply (eq_transitive ?? (-y + (-x + x) + y)); [2: apply feq_plusr; apply plus_assoc;] -apply (eq_transitive ?? (-y + 0 + y)); - [2: apply feq_plusr; apply feq_plusl; apply eq_symmetric; apply opp_inverse] -apply (eq_transitive ?? (-y + y)); - [2: apply feq_plusr; apply eq_symmetric; - apply (eq_transitive ?? (0+-y)); [apply plus_comm|apply zero_neutral]] -apply eq_symmetric; apply opp_inverse. +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)); + [2: apply feq_plusr; apply feq_plusl; apply eq_sym; apply opp_inverse] +apply (eq_trans ?? (-y + y)); + [2: apply feq_plusr; apply eq_sym; + apply (eq_trans ?? (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_transitive ?? (--x + -x)); [apply plus_comm] -apply (eq_transitive ?? 0); [apply opp_inverse] -apply eq_symmetric; apply opp_inverse; +apply (eq_trans ?? (--x + -x)); [apply plus_comm] +apply (eq_trans ?? 0); [apply 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_transitive ?? 0); [apply zero_neutral;] -apply eq_symmetric; apply opp_inverse; +apply (eq_trans ?? 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_reflexive; +qed. + +lemma feq_oppl: ∀G:abelian_group.∀x,y,z:G. y ≈ z → -y ≈ x → -z ≈ x. +intros (G x y z H1 H2); apply eq_sym; apply (feq_oppr ??y); +[2:apply eq_sym] assumption; +qed. + +lemma feq_opp: ∀G:abelian_group.∀x,y:G. x ≈ y → -x ≈ -y. +intros (G x y H); apply (feq_oppl ??y ? H); apply eq_reflexive; +qed. + +coercion cic:/matita/groups/feq_opp.con nocomposites. + +lemma eq_opp_sym: ∀G:abelian_group.∀x,y:G. y ≈ x → -x ≈ -y. +compose feq_opp with eq_sym (H); apply H; assumption; +qed. + +coercion cic:/matita/groups/eq_opp_sym.con nocomposites. + +lemma eq_opp_plusr: ∀G:abelian_group.∀x,y,z:G. x ≈ y → -(x + z) ≈ -(y + z). +compose feq_plusr with feq_opp(H); apply H; assumption; +qed. + +coercion cic:/matita/groups/eq_opp_plusr.con nocomposites. + +lemma eq_opp_plusl: ∀G:abelian_group.∀x,y,z:G. x ≈ y → -(z + x) ≈ -(z + y). +compose feq_plusl with feq_opp(H); apply H; assumption; +qed. + +coercion cic:/matita/groups/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.