| author | wenzelm | 
| Thu, 15 Nov 2001 23:21:57 +0100 | |
| changeset 12216 | dda8c04a8fb4 | 
| parent 11443 | 77ed7e2b56c8 | 
| child 12459 | 6978ab7cac64 | 
| permissions | -rw-r--r-- | 
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/GroupTheory/Group  | 
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Florian Kammueller, with new proofs by L C Paulson  | 
| 11394 | 4  | 
Copyright 1998-2001 University of Cambridge  | 
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 11394 | 6  | 
Group theory using locales  | 
7  | 
*)  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 11394 | 10  | 
fun afs thms = (asm_full_simp_tac (simpset() addsimps thms));  | 
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
12  | 
|
| 11394 | 13  | 
(* Proof of group theory theorems necessary for Sylow's theorem *)  | 
14  | 
Open_locale "group";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
15  | 
|
| 11394 | 16  | 
val e_def = thm "e_def";  | 
17  | 
val binop_def = thm "binop_def";  | 
|
18  | 
val inv_def = thm "inv_def";  | 
|
19  | 
val Group_G = thm "Group_G";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 11394 | 22  | 
val simp_G = simplify (simpset() addsimps [Group_def]) (Group_G);  | 
23  | 
Addsimps [simp_G, Group_G];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
24  | 
|
| 11394 | 25  | 
Goal "e \\<in> carrier G";  | 
26  | 
by (afs [e_def] 1);  | 
|
27  | 
qed "e_closed";  | 
|
28  | 
val unit_closed = export e_closed;  | 
|
29  | 
Addsimps [e_closed];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
30  | 
|
| 11394 | 31  | 
Goal "op ## \\<in> carrier G \\<rightarrow> carrier G \\<rightarrow> carrier G";  | 
32  | 
by (simp_tac (simpset() addsimps [binop_def]) 1);  | 
|
33  | 
qed "binop_funcset";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
34  | 
|
| 11394 | 35  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G |] ==> x ## y \\<in> carrier G";  | 
36  | 
by (afs [binop_funcset RS (funcset_mem RS funcset_mem)] 1);  | 
|
37  | 
qed "binop_closed";  | 
|
38  | 
val bin_op_closed = export binop_closed;  | 
|
39  | 
Addsimps [binop_closed];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
40  | 
|
| 11394 | 41  | 
Goal "INV \\<in> carrier G \\<rightarrow> carrier G";  | 
42  | 
by (simp_tac (simpset() addsimps [inv_def]) 1);  | 
|
43  | 
qed "inv_funcset";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
44  | 
|
| 11394 | 45  | 
Goal "x \\<in> carrier G ==> i(x) \\<in> carrier G";  | 
46  | 
by (afs [inv_funcset RS funcset_mem] 1);  | 
|
47  | 
qed "inv_closed";  | 
|
48  | 
val inverse_closed = export inv_closed;  | 
|
49  | 
Addsimps [inv_closed];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
50  | 
|
| 11394 | 51  | 
Goal "x \\<in> carrier G ==> e ## x = x";  | 
52  | 
by (afs [e_def, binop_def] 1);  | 
|
53  | 
qed "e_ax1";  | 
|
54  | 
Addsimps [e_ax1];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
55  | 
|
| 11394 | 56  | 
Goal "x \\<in> carrier G ==> i(x) ## x = e";  | 
57  | 
by (afs [binop_def, inv_def, e_def] 1);  | 
|
58  | 
qed "inv_ax2";  | 
|
59  | 
Addsimps [inv_ax2];  | 
|
60  | 
||
61  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G; z \\<in> carrier G |] \  | 
|
62  | 
\ ==> (x ## y) ## z = x ## (y ## z)";  | 
|
63  | 
by (afs [binop_def] 1);  | 
|
64  | 
qed "binop_assoc";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
65  | 
|
| 11394 | 66  | 
Goal "[|f \\<in> A \\<rightarrow> A \\<rightarrow> A; i1 \\<in> A \\<rightarrow> A; e1 \\<in> A;\  | 
67  | 
\ \\<forall>x \\<in> A. (f (i1 x) x = e1); \\<forall>x \\<in> A. (f e1 x = x);\  | 
|
68  | 
\ \\<forall>x \\<in> A. \\<forall>y \\<in> A. \\<forall>z \\<in> A. (f (f x y) z = f (x) (f y z)) |] \  | 
|
69  | 
\ ==> (| carrier = A, bin_op = f, inverse = i1, unit = e1 |) \\<in> Group";  | 
|
70  | 
by (afs [Group_def] 1);  | 
|
71  | 
qed "groupI";  | 
|
72  | 
val GroupI = export groupI;  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
73  | 
|
| 11394 | 74  | 
(*****)  | 
75  | 
(* Now the real derivations *)  | 
|
76  | 
||
77  | 
Goal "[| x\\<in>carrier G ; y\\<in>carrier G; z\\<in>carrier G; x ## y = x ## z |] \  | 
|
78  | 
\ ==> y = z";  | 
|
79  | 
by (dres_inst_tac [("f","%z. i x ## z")] arg_cong 1); 
 | 
|
80  | 
by (asm_full_simp_tac (simpset() addsimps [binop_assoc RS sym]) 1);  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
81  | 
qed "left_cancellation";  | 
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
82  | 
|
| 11394 | 83  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G; z \\<in> carrier G |] \  | 
84  | 
\ ==> (x ## y = x ## z) = (y = z)";  | 
|
85  | 
by (blast_tac (claset() addIs [left_cancellation]) 1);  | 
|
86  | 
qed "left_cancellation_iff";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
88  | 
|
| 11394 | 89  | 
(* Now the other directions of basic lemmas; they need a left cancellation*)  | 
90  | 
||
91  | 
Goal "x \\<in> carrier G ==> x ## e = x";  | 
|
92  | 
by (res_inst_tac [("x","i x")] left_cancellation 1);
 | 
|
93  | 
by (auto_tac (claset(), simpset() addsimps [binop_assoc RS sym]));  | 
|
94  | 
qed "e_ax2";  | 
|
95  | 
Addsimps [e_ax2];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
96  | 
|
| 11394 | 97  | 
Goal "[| x \\<in> carrier G; x ## x = x |] ==> x = e";  | 
98  | 
by (res_inst_tac [("x","x")] left_cancellation 1);
 | 
|
99  | 
by Auto_tac;  | 
|
100  | 
qed "idempotent_e";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
101  | 
|
| 11394 | 102  | 
Goal "x \\<in> carrier G ==> x ## i(x) = e";  | 
103  | 
by (rtac idempotent_e 1);  | 
|
104  | 
by (auto_tac (claset(), simpset() addsimps [binop_assoc RS sym]));  | 
|
105  | 
by (asm_simp_tac (simpset() addsimps [inst "x" "x" binop_assoc]) 1);  | 
|
106  | 
qed "inv_ax1";  | 
|
107  | 
Addsimps [inv_ax1];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
108  | 
|
| 11394 | 109  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G; x ## y = e |] ==> y = i(x)";  | 
110  | 
by (res_inst_tac [("x","x")] left_cancellation 1);
 | 
|
111  | 
by Auto_tac;  | 
|
112  | 
qed "inv_unique";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
113  | 
|
| 11394 | 114  | 
Goal "x \\<in> carrier G ==> i(i(x)) = x";  | 
115  | 
by (res_inst_tac [("x","i x")] left_cancellation 1);
 | 
|
116  | 
by Auto_tac;  | 
|
117  | 
qed "inv_inv";  | 
|
118  | 
Addsimps [inv_inv];  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
119  | 
|
| 11394 | 120  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G |] ==> i(x ## y) = i(y) ## i(x)";  | 
121  | 
by (rtac (inv_unique RS sym) 1);  | 
|
122  | 
by (auto_tac (claset(), simpset() addsimps [binop_assoc RS sym]));  | 
|
123  | 
by (asm_simp_tac (simpset() addsimps [inst "x" "x" binop_assoc]) 1);  | 
|
124  | 
qed "inv_prod";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
125  | 
|
| 11394 | 126  | 
Goal "[| y ## x = z ## x; \  | 
127  | 
\ x \\<in> carrier G; y \\<in> carrier G; z \\<in> carrier G|] ==> y = z";  | 
|
128  | 
by (dres_inst_tac [("f","%z. z ## i x")] arg_cong 1); 
 | 
|
129  | 
by (asm_full_simp_tac (simpset() addsimps [binop_assoc]) 1);  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
130  | 
qed "right_cancellation";  | 
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
131  | 
|
| 11394 | 132  | 
Goal "[| x \\<in> carrier G; y \\<in> carrier G; z \\<in> carrier G |] \  | 
133  | 
\ ==> (y ## x = z ## x) = (y = z)";  | 
|
134  | 
by (blast_tac (claset() addIs [right_cancellation]) 1);  | 
|
135  | 
qed "right_cancellation_iff";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
137  | 
|
| 11394 | 138  | 
(* subgroup *)  | 
139  | 
Goal "[| H <= carrier G; H \\<noteq> {}; \\<forall>a \\<in> H. i(a)\\<in>H; \\<forall>a\\<in>H. \\<forall>b\\<in>H. a ## b\\<in>H|] \
 | 
|
140  | 
\ ==> e \\<in> H";  | 
|
141  | 
by (Force_tac 1);  | 
|
142  | 
qed "e_in_H";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
143  | 
|
| 11394 | 144  | 
(* subgroupI: a characterization of subgroups *)  | 
145  | 
Goal "[| H <= carrier G; H \\<noteq> {}; \\<forall>a \\<in> H. i(a)\\<in>H;\
 | 
|
146  | 
\ \\<forall>a\\<in>H. \\<forall>b\\<in>H. a ## b\\<in>H |] ==> H <<= G";  | 
|
147  | 
by (asm_full_simp_tac (simpset() addsimps [subgroup_def]) 1);  | 
|
148  | 
(* Fold the locale definitions: the top level definition of subgroup gives  | 
|
149  | 
the verbose form, which does not immediately match rules like inv_ax1 *)  | 
|
150  | 
by (rtac groupI 1);  | 
|
151  | 
by (ALLGOALS (asm_full_simp_tac  | 
|
152  | 
(simpset() addsimps [subsetD, restrictI, e_in_H, binop_assoc] @  | 
|
153  | 
(map symmetric [binop_def, inv_def, e_def]))));  | 
|
154  | 
qed "subgroupI";  | 
|
155  | 
val SubgroupI = export subgroupI;  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
156  | 
|
| 11394 | 157  | 
Goal "H <<= G ==> \  | 
158  | 
\ (|carrier = H, bin_op = lam x:H. lam y:H. x ## y, \  | 
|
159  | 
\ inverse = lam x:H. i(x), unit = e|)\\<in>Group";  | 
|
160  | 
by (afs [subgroup_def, binop_def, inv_def, e_def] 1);  | 
|
161  | 
qed "subgroupE2";  | 
|
162  | 
val SubgroupE2 = export subgroupE2;  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
163  | 
|
| 11394 | 164  | 
Goal "H <<= G ==> H <= carrier G";  | 
165  | 
by (afs [subgroup_def, binop_def, inv_def, e_def] 1);  | 
|
166  | 
qed "subgroup_imp_subset";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
167  | 
|
| 11394 | 168  | 
Goal "[| H <<= G; x \\<in> H; y \\<in> H |] ==> x ## y \\<in> H";  | 
169  | 
by (dtac subgroupE2 1);  | 
|
170  | 
by (dtac bin_op_closed 1);  | 
|
171  | 
by (Asm_full_simp_tac 1);  | 
|
172  | 
by (thin_tac "x\\<in>H" 1);  | 
|
173  | 
by Auto_tac;  | 
|
174  | 
qed "subgroup_f_closed";  | 
|
175  | 
||
176  | 
Goal "[| H <<= G; x \\<in> H |] ==> i x \\<in> H";  | 
|
177  | 
by (dtac (subgroupE2 RS inverse_closed) 1);  | 
|
178  | 
by Auto_tac;  | 
|
179  | 
qed "subgroup_inv_closed";  | 
|
180  | 
val Subgroup_inverse_closed = export subgroup_inv_closed;  | 
|
181  | 
||
182  | 
Goal "H <<= G ==> e\\<in>H";  | 
|
183  | 
by (dtac (subgroupE2 RS unit_closed) 1);  | 
|
184  | 
by (Asm_full_simp_tac 1);  | 
|
185  | 
qed "subgroup_e_closed";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
186  | 
|
| 11394 | 187  | 
Goal "[| finite(carrier G); H <<= G |] ==> 0 < card(H)";  | 
188  | 
by (subgoal_tac "finite H" 1);  | 
|
189  | 
by (blast_tac (claset() addIs [finite_subset] addDs [subgroup_imp_subset]) 2);  | 
|
190  | 
by (rtac ccontr 1);  | 
|
191  | 
by (asm_full_simp_tac (simpset() addsimps [card_0_eq]) 1);  | 
|
192  | 
by (blast_tac (claset() addDs [subgroup_e_closed]) 1);  | 
|
193  | 
qed "SG_card1";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
194  | 
|
| 11394 | 195  | 
(* Abelian Groups *)  | 
196  | 
Goal "[|G' \\<in> AbelianGroup; x: carrier G'; y: carrier G'|] \  | 
|
197  | 
\ ==> (G'.<f>) x y = (G'.<f>) y x";  | 
|
198  | 
by (auto_tac (claset(), simpset() addsimps [AbelianGroup_def]));  | 
|
199  | 
qed "Group_commute";  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
200  | 
|
| 11394 | 201  | 
Goal "AbelianGroup <= Group";  | 
202  | 
by (auto_tac (claset(), simpset() addsimps [AbelianGroup_def]));  | 
|
203  | 
qed "Abel_subset_Group";  | 
|
204  | 
||
205  | 
val Abel_imp_Group = Abel_subset_Group RS subsetD;  | 
|
206  | 
||
| 11443 | 207  | 
Delsimps [simp_G, Group_G];  | 
| 11394 | 208  | 
Close_locale "group";  | 
| 11443 | 209  | 
|
210  | 
Goal "G \\<in> Group ==> (| carrier = G .<cr>, bin_op = G .<f>, inverse = G .<inv>,\  | 
|
211  | 
\ unit = G .<e> |) \\<in> Group";  | 
|
212  | 
by (blast_tac  | 
|
213  | 
(claset() addIs [GroupI, export binop_funcset, export inv_funcset, export e_closed, export inv_ax2, export e_ax1, export binop_assoc]) 1);  | 
|
214  | 
qed "Group_Group";  | 
|
215  | 
||
216  | 
Goal "G \\<in> AbelianGroup \  | 
|
217  | 
\ ==> (| carrier = G .<cr>, bin_op = G .<f>, inverse = G .<inv>,\  | 
|
218  | 
\ unit = G .<e> |) \\<in> AbelianGroup";  | 
|
219  | 
by (asm_full_simp_tac (simpset() addsimps [AbelianGroup_def]) 1);  | 
|
220  | 
by (rtac Group_Group 1);  | 
|
221  | 
by Auto_tac;  | 
|
222  | 
qed "Abel_Abel";  | 
|
223  |