1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/groups/".
17 include "excedence.ma".
19 definition left_neutral ≝ λC:apartness.λop.λe:C. ∀x:C. op e x ≈ x.
20 definition right_neutral ≝ λC:apartness.λop. λe:C. ∀x:C. op x e ≈ x.
21 definition left_inverse ≝ λC:apartness.λop.λe:C.λinv:C→C. ∀x:C. op (inv x) x ≈ e.
22 definition right_inverse ≝ λC:apartness.λop.λe:C.λ inv: C→ C. ∀x:C. op x (inv x) ≈ e.
23 definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y.
24 (* ALLOW DEFINITION WITH SOME METAS *)
26 definition distributive_left ≝
27 λA:apartness.λf:A→A→A.λg:A→A→A.
28 ∀x,y,z. f x (g y z) ≈ g (f x y) (f x z).
30 definition distributive_right ≝
31 λA:apartness.λf:A→A→A.λg:A→A→A.
32 ∀x,y,z. f (g x y) z ≈ g (f x z) (f y z).
34 record abelian_group : Type ≝
36 plus: carr → carr → carr;
39 plus_assoc: associative ? plus (eq carr);
40 plus_comm: commutative ? plus (eq carr);
41 zero_neutral: left_neutral ? plus zero;
42 opp_inverse: left_inverse ? plus zero opp;
43 plus_strong_ext: ∀z.strong_ext ? (plus z)
46 notation "0" with precedence 89 for @{ 'zero }.
48 interpretation "Abelian group zero" 'zero =
49 (cic:/matita/groups/zero.con _).
51 interpretation "Abelian group plus" 'plus a b =
52 (cic:/matita/groups/plus.con _ a b).
54 interpretation "Abelian group opp" 'uminus a =
55 (cic:/matita/groups/opp.con _ a).
58 λG:abelian_group.λa,b:G. a + -b.
60 interpretation "Abelian group minus" 'minus a b =
61 (cic:/matita/groups/minus.con _ a b).
63 definition ext ≝ λA:apartness.λop:A→A. ∀x,y. x ≈ y → op x ≈ op y.
65 lemma strong_ext_to_ext: ∀A:apartness.∀op:A→A. strong_ext ? op → ext ? op.
66 intros 6 (A op SEop x y Exy); intro Axy; apply Exy; apply SEop; assumption;
69 lemma feq_plusl: ∀G:abelian_group.∀x,y,z:G. y ≈ z → x+y ≈ x+z.
70 intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_ext ? x));
74 coercion cic:/matita/groups/feq_plusl.con.
76 lemma plus_strong_extr: ∀G:abelian_group.∀z:G.strong_ext ? (λx.x + z).
77 intros 5 (G z x y A); simplify in A;
78 lapply (plus_comm ? z x) as E1; lapply (plus_comm ? z y) as E2;
79 lapply (ap_rewl ???? E1 A) as A1; lapply (ap_rewr ???? E2 A1) as A2;
80 apply (plus_strong_ext ???? A2);
83 lemma feq_plusr: ∀G:abelian_group.∀x,y,z:G. y ≈ z → y+x ≈ z+x.
84 intros (G x y z Eyz); apply (strong_ext_to_ext ?? (plus_strong_extr ? x));
88 coercion cic:/matita/groups/feq_plusr.con.
90 lemma fap_plusl: ∀G:abelian_group.∀x,y,z:G. y # z → x+y # x+z.
91 intros (G x y z Ayz); apply (plus_strong_ext ? (-x));
92 apply (ap_rewl ??? ((-x + x) + y));
94 |2: apply (ap_rewr ??? ((-x +x) +z));
96 |2: apply (ap_rewl ??? (0 + y));
97 [1: apply (feq_plusr ???? (opp_inverse ??));
98 |2: apply (ap_rewl ???? (zero_neutral ? y)); apply (ap_rewr ??? (0 + z));
99 [1: apply (feq_plusr ???? (opp_inverse ??));
100 |2: apply (ap_rewr ???? (zero_neutral ??)); assumption;]]]]
105 lemma fap_plusr: ∀G:abelian_group.∀x,y,z:G. y # z → y+x # z+x.
106 intros (G x y z Ayz); apply (plus_strong_extr ? (-x));
107 apply (ap_rewl ??? (y + (x + -x)));
108 [1: apply (eq_symmetric ??? (plus_assoc ????));
109 |2: apply (ap_rewr ??? (z + (x + -x)));
110 [1: apply (eq_symmetric ??? (plus_assoc ????));
111 |2: apply (ap_rewl ??? (y + (-x+x)) (feq_plusl ???? (plus_comm ???)));
112 apply (ap_rewl ??? (y + 0) (feq_plusl ???? (opp_inverse ??)));
113 apply (ap_rewl ??? (0 + y) (plus_comm ???));
114 apply (ap_rewl ??? y (zero_neutral ??));
115 apply (ap_rewr ??? (z + (-x+x)) (feq_plusl ???? (plus_comm ???)));
116 apply (ap_rewr ??? (z + 0) (feq_plusl ???? (opp_inverse ??)));
117 apply (ap_rewr ??? (0 + z) (plus_comm ???));
118 apply (ap_rewr ??? z (zero_neutral ??));
122 lemma plus_cancl: ∀G:abelian_group.∀y,z,x:G. x+y ≈ x+z → y ≈ z.
123 intros 6 (G y z x E Ayz); apply E; apply fap_plusl; assumption;
126 lemma plus_cancr: ∀G:abelian_group.∀y,z,x:G. y+x ≈ z+x → y ≈ z.
127 intros 6 (G y z x E Ayz); apply E; apply fap_plusr; assumption;
130 theorem eq_opp_plus_plus_opp_opp:
131 ∀G:abelian_group.∀x,y:G. -(x+y) ≈ -x + -y.
132 intros (G x y); apply (plus_cancr ??? (x+y));
133 apply (eq_transitive ?? 0); [apply (opp_inverse ??)]
134 apply (eq_transitive ?? (-x + -y + x + y)); [2: apply (eq_symmetric ??? (plus_assoc ????))]
135 apply (eq_transitive ?? (-y + -x + x + y)); [2: repeat apply feq_plusr; apply plus_comm]
136 apply (eq_transitive ?? (-y + (-x + x) + y)); [2: apply feq_plusr; apply plus_assoc;]
137 apply (eq_transitive ?? (-y + 0 + y));
138 [2: apply feq_plusr; apply feq_plusl; apply eq_symmetric; apply opp_inverse]
139 apply (eq_transitive ?? (-y + y));
140 [2: apply feq_plusr; apply eq_symmetric;
141 apply (eq_transitive ?? (0+-y)); [apply plus_comm|apply zero_neutral]]
142 apply eq_symmetric; apply opp_inverse.
145 theorem eq_opp_opp_x_x: ∀G:abelian_group.∀x:G.--x ≈ x.
146 intros (G x); apply (plus_cancl ??? (-x));
147 apply (eq_transitive ?? (--x + -x)); [apply plus_comm]
148 apply (eq_transitive ?? 0); [apply opp_inverse]
149 apply eq_symmetric; apply opp_inverse;
152 theorem eq_zero_opp_zero: ∀G:abelian_group.0 ≈ -0. [assumption]
153 intro G; apply (plus_cancr ??? 0);
154 apply (eq_transitive ?? 0); [apply zero_neutral;]
155 apply eq_symmetric; apply opp_inverse;