| author | paulson | 
| Mon, 23 Jul 2001 17:37:29 +0200 | |
| changeset 11443 | 77ed7e2b56c8 | 
| parent 11394 | e88c2c89f98e | 
| child 11468 | 02cd5d5bc497 | 
| permissions | -rw-r--r-- | 
| 11394 | 1  | 
(* Title: HOL/GroupTheory/Sylow  | 
| 
11370
 
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  | 
Sylow's theorem using locales (combinatorial argument in Exponent.thy)  | 
| 11443 | 7  | 
|
8  | 
See Florian Kamm\"uller and L. C. Paulson,  | 
|
9  | 
A Formal Proof of Sylow's theorem:  | 
|
10  | 
An Experiment in Abstract Algebra with Isabelle HOL  | 
|
11  | 
J. Automated Reasoning 23 (1999), 235-264  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
12  | 
*)  | 
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
13  | 
|
| 11394 | 14  | 
Sylow = Coset +  | 
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
15  | 
|
| 11394 | 16  | 
locale sylow = coset +  | 
17  | 
fixes  | 
|
18  | 
p :: "nat"  | 
|
19  | 
a :: "nat"  | 
|
20  | 
m :: "nat"  | 
|
21  | 
calM :: "'a set set"  | 
|
22  | 
    RelM      ::  "('a set * 'a set)set"
 | 
|
23  | 
assumes  | 
|
24  | 
prime_p "p\\<in>prime"  | 
|
25  | 
order_G "order(G) = (p^a) * m"  | 
|
26  | 
finite_G "finite (carrier G)"  | 
|
27  | 
defines  | 
|
28  | 
    calM_def "calM == {s. s <= carrier(G) & card(s) = p^a}"
 | 
|
29  | 
    RelM_def "RelM == {(N1,N2). (N1,N2) \\<in> calM \\<times> calM &
 | 
|
30  | 
(\\<exists>g \\<in> carrier(G). N1 = (N2 #> g) )}"  | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
31  | 
|
| 11394 | 32  | 
locale sylow_central = sylow +  | 
33  | 
fixes  | 
|
34  | 
H :: "'a set"  | 
|
35  | 
M :: "'a set set"  | 
|
36  | 
M1 :: "'a set"  | 
|
37  | 
assumes  | 
|
38  | 
M_in_quot "M \\<in> calM // RelM"  | 
|
39  | 
not_dvd_M "~(p ^ (exponent p m + 1) dvd card(M))"  | 
|
40  | 
M1_in_M "M1 \\<in> M"  | 
|
41  | 
defines  | 
|
42  | 
   H_def "H == {g. g\\<in>carrier G & M1 #> g = M1}"
 | 
|
| 
11370
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
680946254afe
new GroupTheory example, e.g. the Sylow theorem (preliminary version)
 
paulson 
parents:  
diff
changeset
 | 
44  | 
end  |