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/algebra/groups/".
17 include "algebra/monoids.ma".
18 include "nat/le_arith.ma".
19 include "datatypes/bool.ma".
20 include "nat/compare.ma".
22 record PreGroup : Type ≝
23 { premonoid:> PreMonoid;
24 inv: premonoid -> premonoid
27 record isGroup (G:PreGroup) : Prop ≝
28 { is_monoid:> isMonoid G;
29 inv_is_left_inverse: is_left_inverse (mk_Monoid ? is_monoid) (inv G);
30 inv_is_right_inverse: is_right_inverse (mk_Monoid ? is_monoid) (inv G)
34 { pregroup:> PreGroup;
35 group_properties:> isGroup pregroup
38 notation "hvbox(x \sup (-1))" with precedence 89
41 interpretation "Group inverse" 'ginv x =
42 (cic:/matita/algebra/groups/inv.con _ x).
44 definition left_cancellable ≝
45 λT:Type. λop: T -> T -> T.
46 ∀x. injective ? ? (op x).
48 definition right_cancellable ≝
49 λT:Type. λop: T -> T -> T.
50 ∀x. injective ? ? (λz.op z x).
52 theorem eq_op_x_y_op_x_z_to_eq:
53 ∀G:Group. left_cancellable G (op G).
55 unfold left_cancellable;
58 rewrite < (e_is_left_unit ? G);
59 rewrite < (e_is_left_unit ? G z);
60 rewrite < (inv_is_left_inverse ? G x);
61 rewrite > (op_associative ? G);
62 rewrite > (op_associative ? G);
68 theorem eq_op_x_y_op_z_y_to_eq:
69 ∀G:Group. right_cancellable G (op G).
71 unfold right_cancellable;
73 simplify;fold simplify (op G);
75 rewrite < (e_is_right_unit ? G);
76 rewrite < (e_is_right_unit ? G z);
77 rewrite < (inv_is_right_inverse ? G x);
78 rewrite < (op_associative ? G);
79 rewrite < (op_associative ? G);
84 theorem eq_inv_inv_x_x: ∀G:Group. ∀x:G. x \sup -1 \sup -1 = x.
86 apply (eq_op_x_y_op_z_y_to_eq ? (x \sup -1));
87 rewrite > (inv_is_right_inverse ? G);
88 rewrite > (inv_is_left_inverse ? G);
92 theorem eq_opxy_e_to_eq_x_invy:
93 ∀G:Group. ∀x,y:G. x·y=1 → x=y \sup -1.
95 apply (eq_op_x_y_op_z_y_to_eq ? y);
96 rewrite > (inv_is_left_inverse ? G);
100 theorem eq_opxy_e_to_eq_invx_y:
101 ∀G:Group. ∀x,y:G. x·y=1 → x \sup -1=y.
103 apply (eq_op_x_y_op_x_z_to_eq ? x);
104 rewrite > (inv_is_right_inverse ? G);
109 theorem eq_opxy_z_to_eq_x_opzinvy:
110 ∀G:Group. ∀x,y,z:G. x·y=z → x = z·y \sup -1.
112 apply (eq_op_x_y_op_z_y_to_eq ? y);
113 rewrite > (op_associative ? G);
114 rewrite > (inv_is_left_inverse ? G);
115 rewrite > (e_is_right_unit ? G);
119 theorem eq_opxy_z_to_eq_y_opinvxz:
120 ∀G:Group. ∀x,y,z:G. x·y=z → y = x \sup -1·z.
122 apply (eq_op_x_y_op_x_z_to_eq ? x);
123 rewrite < (op_associative ? G);
124 rewrite > (inv_is_right_inverse ? G);
125 rewrite > (e_is_left_unit ? G);
129 theorem eq_inv_op_x_y_op_inv_y_inv_x:
130 ∀G:Group. ∀x,y:G. (x·y) \sup -1 = y \sup -1 · x \sup -1.
132 apply (eq_op_x_y_op_z_y_to_eq ? (x·y));
133 rewrite > (inv_is_left_inverse ? G);
134 rewrite < (op_associative ? G);
135 rewrite > (op_associative ? G (y \sup -1));
136 rewrite > (inv_is_left_inverse ? G);
137 rewrite > (e_is_right_unit ? G);
138 rewrite > (inv_is_left_inverse ? G);
144 record morphism (G,G':Group) : Type ≝
146 f_morph: ∀x,y:G.image(x·y) = image x · image y
149 notation "hvbox(f˜ x)" with precedence 79
150 for @{ 'morimage $f $x }.
152 interpretation "Morphism image" 'morimage f x =
153 (cic:/matita/algebra/groups/image.con _ _ f x).
155 theorem morphism_to_eq_f_1_1:
156 ∀G,G'.∀f:morphism G G'.f˜1 = 1.
158 apply (eq_op_x_y_op_z_y_to_eq ? (f˜1));
159 rewrite > (e_is_left_unit ? G');
161 rewrite > (e_is_left_unit ? G);
165 theorem eq_image_inv_inv_image:
166 ∀G,G'.∀f:morphism G G'.
167 ∀x.f˜(x \sup -1) = (f˜x) \sup -1.
169 apply (eq_op_x_y_op_z_y_to_eq ? (f˜x));
170 rewrite > (inv_is_left_inverse ? G');
172 rewrite > (inv_is_left_inverse ? G);
173 apply (morphism_to_eq_f_1_1 ? ? f).
176 record monomorphism (G,G':Group) : Type ≝
177 { morphism:> morphism G G';
178 injective: injective ? ? (image ? ? morphism)
183 record subgroup (G:Group) : Type ≝
185 embed:> monomorphism group G
188 notation "hvbox(x \sub H)" with precedence 79
189 for @{ 'subgroupimage $H $x }.
191 interpretation "Subgroup image" 'subgroupimage H x =
192 (cic:/matita/algebra/groups/image.con _ _
193 (cic:/matita/algebra/groups/morphism_of_subgroup.con _ H) x).
195 definition member_of_subgroup ≝
196 λG.λH:subgroup G.λx:G.∃y.x=y \sub H.
198 notation "hvbox(x break ∈ H)" with precedence 79
199 for @{ 'member_of $x $H }.
201 notation "hvbox(x break ∉ H)" with precedence 79
202 for @{ 'not_member_of $x $H }.
204 interpretation "Member of subgroup" 'member_of x H =
205 (cic:/matita/algebra/groups/member_of_subgroup.con _ H x).
207 interpretation "Not member of subgroup" 'not_member_of x H =
208 (cic:/matita/logic/connectives/Not.con
209 (cic:/matita/algebra/groups/member_of_subgroup.con _ H x)).
213 record left_coset (G:Group) : Type ≝
218 (* Here I would prefer 'magma_op, but this breaks something in the next definition *)
219 interpretation "Left_coset" 'times x C =
220 (cic:/matita/algebra/groups/left_coset.ind#xpointer(1/1/1) _ x C).
222 definition member_of_left_coset ≝
223 λG:Group.λC:left_coset G.λx:G.
224 ∃y.x=(element ? C)·y \sub (subgrp ? C).
226 interpretation "Member of left_coset" 'member_of x C =
227 (cic:/matita/algebra/groups/member_of_left_coset.con _ C x).
229 definition left_coset_eq ≝
230 λG.λC,C':left_coset G.
231 ∀x.((element ? C)·x \sub (subgrp ? C)) ∈ C'.
233 interpretation "Left cosets equality" 'eq C C' =
234 (cic:/matita/algebra/groups/left_coset_eq.con _ C C').
236 definition left_coset_disjoint ≝
237 λG.λC,C':left_coset G.
238 ∀x.¬(((element ? C)·x \sub (subgrp ? C)) ∈ C').
240 notation "hvbox(a break ∥ b)"
241 non associative with precedence 45
242 for @{ 'disjoint $a $b }.
244 interpretation "Left cosets disjoint" 'disjoint C C' =
245 (cic:/matita/algebra/groups/left_coset_disjoint.con _ C C').
247 (* The following should be a one-shot alias! *)
248 alias symbol "member_of" (instance 0) = "Member of subgroup".
249 theorem member_of_subgroup_op_inv_x_y_to_left_coset_eq:
250 ∀G.∀x,y.∀H:subgroup G. (x \sup -1 ·y) ∈ H → x*H = y*H.
254 unfold member_of_subgroup in H1;
258 [ apply (a\sup-1 · x1)
260 rewrite > eq_image_inv_inv_image;
262 rewrite > eq_inv_op_x_y_op_inv_y_inv_x;
263 rewrite > eq_inv_inv_x_x;
264 rewrite < (op_associative ? G);
265 rewrite < (op_associative ? G);
266 rewrite > (inv_is_right_inverse ? G);
267 rewrite > (e_is_left_unit ? G);
272 theorem Not_member_of_subgroup_to_left_coset_disjoint:
273 ∀G.∀x,y.∀H:subgroup G.(x \sup -1 ·y) ∉ H → x*H ∥ y*H.
279 unfold member_of_subgroup;
281 apply (ex_intro ? ? (x'·a \sup -1));
283 apply (eq_op_x_y_op_z_y_to_eq ? (a \sub H));
284 rewrite > (op_associative ? G);
286 rewrite > (op_associative ? G);
288 rewrite > (inv_is_left_inverse ? H);
289 rewrite < (op_associative ? G);
290 rewrite > (inv_is_left_inverse ? G);
291 rewrite > (e_is_left_unit ? G);
292 rewrite < (f_morph ? ? H);
293 rewrite > (e_is_right_unit ? H);
297 (*CSC: here the coercion Type_of_Group cannot be omitted. Why? *)
298 theorem in_x_mk_left_coset_x_H:
299 ∀G.∀x:Type_of_Group G.∀H:subgroup G.x ∈ (x*H).
302 apply (ex_intro ? ? 1);
303 rewrite > morphism_to_eq_f_1_1;
304 rewrite > (e_is_right_unit ? G);
308 (* Normal Subgroups *)
310 record normal_subgroup (G:Group) : Type ≝
311 { ns_subgroup:> subgroup G;
312 normal:> ∀x:G.∀y:ns_subgroup.(x·y \sub ns_subgroup·x \sup -1) ∈ ns_subgroup
315 (*CSC: I have not defined yet right cosets
317 ∀G.∀H:normal_subgroup G.∀x.x*H=H*x.
320 theorem member_of_left_coset_mk_left_coset_x_H_a_to_member_of_left_coset_mk_left_coset_y_H_b_to_member_of_left_coset_mk_left_coset_op_x_y_H_op_a_b:
321 ∀G.∀H:normal_subgroup G.∀x,y,a,b.
322 a ∈ (x*H) → b ∈ (y*H) → (a·b) ∈ ((x·y)*H).