begin
:: deftheorem Def1 defines being_left_operation GROUP_10:def 1 :
for S being non empty unital multMagma
for E being set
for A being Action of the carrier of S,E holds
( A is being_left_operation iff ( A ^ (1_ S) = id E & ( for s1, s2 being Element of S holds A ^ (s1 * s2) = (A ^ s1) * (A ^ s2) ) ) );
theorem
canceled;
:: deftheorem Def2 defines the_left_operation_of GROUP_10:def 2 :
for S being non empty Group-like associative multMagma
for b2 being LeftOperation of S, the carrier of S holds
( b2 = the_left_operation_of S iff for s being Element of S holds b2 . s = the_left_translation_of s );
:: deftheorem defines the_subsets_of_card GROUP_10:def 3 :
for E, n being set holds the_subsets_of_card (n,E) = { X where X is Subset of E : card X = n } ;
theorem Th2:
theorem Th3:
definition
let E be non
empty set ;
let n be
natural number ;
let S be non
empty Group-like multMagma ;
let s be
Element of
S;
let LO be
LeftOperation of
S,
E;
assume A1:
card n c= card E
;
func the_extension_of_left_translation_of (
n,
s,
LO)
-> Function of
(the_subsets_of_card (n,E)),
(the_subsets_of_card (n,E)) means :
Def4:
for
X being
Element of
the_subsets_of_card (
n,
E) holds
it . X = (LO ^ s) .: X;
existence
ex b1 being Function of (the_subsets_of_card (n,E)),(the_subsets_of_card (n,E)) st
for X being Element of the_subsets_of_card (n,E) holds b1 . X = (LO ^ s) .: X
uniqueness
for b1, b2 being Function of (the_subsets_of_card (n,E)),(the_subsets_of_card (n,E)) st ( for X being Element of the_subsets_of_card (n,E) holds b1 . X = (LO ^ s) .: X ) & ( for X being Element of the_subsets_of_card (n,E) holds b2 . X = (LO ^ s) .: X ) holds
b1 = b2
end;
:: deftheorem Def4 defines the_extension_of_left_translation_of GROUP_10:def 4 :
for E being non empty set
for n being natural number
for S being non empty Group-like multMagma
for s being Element of S
for LO being LeftOperation of S,E st card n c= card E holds
for b6 being Function of (the_subsets_of_card (n,E)),(the_subsets_of_card (n,E)) holds
( b6 = the_extension_of_left_translation_of (n,s,LO) iff for X being Element of the_subsets_of_card (n,E) holds b6 . X = (LO ^ s) .: X );
definition
let E be non
empty set ;
let n be
natural number ;
let S be non
empty Group-like multMagma ;
let LO be
LeftOperation of
S,
E;
assume A1:
card n c= card E
;
func the_extension_of_left_operation_of (
n,
LO)
-> LeftOperation of
S,
(the_subsets_of_card (n,E)) means :
Def5:
for
s being
Element of
S holds
it . s = the_extension_of_left_translation_of (
n,
s,
LO);
existence
ex b1 being LeftOperation of S,(the_subsets_of_card (n,E)) st
for s being Element of S holds b1 . s = the_extension_of_left_translation_of (n,s,LO)
uniqueness
for b1, b2 being LeftOperation of S,(the_subsets_of_card (n,E)) st ( for s being Element of S holds b1 . s = the_extension_of_left_translation_of (n,s,LO) ) & ( for s being Element of S holds b2 . s = the_extension_of_left_translation_of (n,s,LO) ) holds
b1 = b2
end;
:: deftheorem Def5 defines the_extension_of_left_operation_of GROUP_10:def 5 :
for E being non empty set
for n being natural number
for S being non empty Group-like multMagma
for LO being LeftOperation of S,E st card n c= card E holds
for b5 being LeftOperation of S,(the_subsets_of_card (n,E)) holds
( b5 = the_extension_of_left_operation_of (n,LO) iff for s being Element of S holds b5 . s = the_extension_of_left_translation_of (n,s,LO) );
definition
let S be non
empty multMagma ;
let s be
Element of
S;
let Z be non
empty set ;
func the_left_translation_of (
s,
Z)
-> Function of
[: the carrier of S,Z:],
[: the carrier of S,Z:] means :
Def6:
for
z1 being
Element of
[: the carrier of S,Z:] ex
z2 being
Element of
[: the carrier of S,Z:] ex
s1,
s2 being
Element of
S ex
z being
Element of
Z st
(
z2 = it . z1 &
s2 = s * s1 &
z1 = [s1,z] &
z2 = [s2,z] );
existence
ex b1 being Function of [: the carrier of S,Z:],[: the carrier of S,Z:] st
for z1 being Element of [: the carrier of S,Z:] ex z2 being Element of [: the carrier of S,Z:] ex s1, s2 being Element of S ex z being Element of Z st
( z2 = b1 . z1 & s2 = s * s1 & z1 = [s1,z] & z2 = [s2,z] )
uniqueness
for b1, b2 being Function of [: the carrier of S,Z:],[: the carrier of S,Z:] st ( for z1 being Element of [: the carrier of S,Z:] ex z2 being Element of [: the carrier of S,Z:] ex s1, s2 being Element of S ex z being Element of Z st
( z2 = b1 . z1 & s2 = s * s1 & z1 = [s1,z] & z2 = [s2,z] ) ) & ( for z1 being Element of [: the carrier of S,Z:] ex z2 being Element of [: the carrier of S,Z:] ex s1, s2 being Element of S ex z being Element of Z st
( z2 = b2 . z1 & s2 = s * s1 & z1 = [s1,z] & z2 = [s2,z] ) ) holds
b1 = b2
end;
:: deftheorem Def6 defines the_left_translation_of GROUP_10:def 6 :
for S being non empty multMagma
for s being Element of S
for Z being non empty set
for b4 being Function of [: the carrier of S,Z:],[: the carrier of S,Z:] holds
( b4 = the_left_translation_of (s,Z) iff for z1 being Element of [: the carrier of S,Z:] ex z2 being Element of [: the carrier of S,Z:] ex s1, s2 being Element of S ex z being Element of Z st
( z2 = b4 . z1 & s2 = s * s1 & z1 = [s1,z] & z2 = [s2,z] ) );
definition
let S be non
empty Group-like associative multMagma ;
let Z be non
empty set ;
func the_left_operation_of (
S,
Z)
-> LeftOperation of
S,
[: the carrier of S,Z:] means :
Def7:
for
s being
Element of
S holds
it . s = the_left_translation_of (
s,
Z);
existence
ex b1 being LeftOperation of S,[: the carrier of S,Z:] st
for s being Element of S holds b1 . s = the_left_translation_of (s,Z)
uniqueness
for b1, b2 being LeftOperation of S,[: the carrier of S,Z:] st ( for s being Element of S holds b1 . s = the_left_translation_of (s,Z) ) & ( for s being Element of S holds b2 . s = the_left_translation_of (s,Z) ) holds
b1 = b2
end;
:: deftheorem Def7 defines the_left_operation_of GROUP_10:def 7 :
for S being non empty Group-like associative multMagma
for Z being non empty set
for b3 being LeftOperation of S,[: the carrier of S,Z:] holds
( b3 = the_left_operation_of (S,Z) iff for s being Element of S holds b3 . s = the_left_translation_of (s,Z) );
definition
let G be
Group;
let H,
P be
Subgroup of
G;
let h be
Element of
H;
func the_left_translation_of (
h,
P)
-> Function of
(Left_Cosets P),
(Left_Cosets P) means :
Def8:
for
P1 being
Element of
Left_Cosets P ex
P2 being
Element of
Left_Cosets P ex
A1,
A2 being
Subset of
G ex
g being
Element of
G st
(
P2 = it . P1 &
A2 = g * A1 &
A1 = P1 &
A2 = P2 &
g = h );
existence
ex b1 being Function of (Left_Cosets P),(Left_Cosets P) st
for P1 being Element of Left_Cosets P ex P2 being Element of Left_Cosets P ex A1, A2 being Subset of G ex g being Element of G st
( P2 = b1 . P1 & A2 = g * A1 & A1 = P1 & A2 = P2 & g = h )
uniqueness
for b1, b2 being Function of (Left_Cosets P),(Left_Cosets P) st ( for P1 being Element of Left_Cosets P ex P2 being Element of Left_Cosets P ex A1, A2 being Subset of G ex g being Element of G st
( P2 = b1 . P1 & A2 = g * A1 & A1 = P1 & A2 = P2 & g = h ) ) & ( for P1 being Element of Left_Cosets P ex P2 being Element of Left_Cosets P ex A1, A2 being Subset of G ex g being Element of G st
( P2 = b2 . P1 & A2 = g * A1 & A1 = P1 & A2 = P2 & g = h ) ) holds
b1 = b2
end;
:: deftheorem Def8 defines the_left_translation_of GROUP_10:def 8 :
for G being Group
for H, P being Subgroup of G
for h being Element of H
for b5 being Function of (Left_Cosets P),(Left_Cosets P) holds
( b5 = the_left_translation_of (h,P) iff for P1 being Element of Left_Cosets P ex P2 being Element of Left_Cosets P ex A1, A2 being Subset of G ex g being Element of G st
( P2 = b5 . P1 & A2 = g * A1 & A1 = P1 & A2 = P2 & g = h ) );
definition
let G be
Group;
let H,
P be
Subgroup of
G;
func the_left_operation_of (
H,
P)
-> LeftOperation of
H,
(Left_Cosets P) means :
Def9:
for
h being
Element of
H holds
it . h = the_left_translation_of (
h,
P);
existence
ex b1 being LeftOperation of H,(Left_Cosets P) st
for h being Element of H holds b1 . h = the_left_translation_of (h,P)
uniqueness
for b1, b2 being LeftOperation of H,(Left_Cosets P) st ( for h being Element of H holds b1 . h = the_left_translation_of (h,P) ) & ( for h being Element of H holds b2 . h = the_left_translation_of (h,P) ) holds
b1 = b2
end;
:: deftheorem Def9 defines the_left_operation_of GROUP_10:def 9 :
for G being Group
for H, P being Subgroup of G
for b4 being LeftOperation of H,(Left_Cosets P) holds
( b4 = the_left_operation_of (H,P) iff for h being Element of H holds b4 . h = the_left_translation_of (h,P) );
begin
:: deftheorem Def10 defines the_strict_stabilizer_of GROUP_10:def 10 :
for G being Group
for E being non empty set
for T being LeftOperation of G,E
for A being Subset of E
for b5 being strict Subgroup of G holds
( b5 = the_strict_stabilizer_of (A,T) iff the carrier of b5 = { g where g is Element of G : (T ^ g) .: A = A } );
:: deftheorem defines the_strict_stabilizer_of GROUP_10:def 11 :
for G being Group
for E being non empty set
for T being LeftOperation of G,E
for x being Element of E holds the_strict_stabilizer_of (x,T) = the_strict_stabilizer_of ({x},T);
:: deftheorem Def12 defines is_fixed_under GROUP_10:def 12 :
for S being non empty unital multMagma
for E being set
for T being LeftOperation of S,E
for x being Element of E holds
( x is_fixed_under T iff for s being Element of S holds x = (T ^ s) . x );
:: deftheorem Def13 defines the_fixed_points_of GROUP_10:def 13 :
for S being non empty unital multMagma
for E being set
for T being LeftOperation of S,E holds
( ( not E is empty implies the_fixed_points_of T = { x where x is Element of E : x is_fixed_under T } ) & ( E is empty implies the_fixed_points_of T = {} E ) );
:: deftheorem Def14 defines are_conjugated_under GROUP_10:def 14 :
for S being non empty unital multMagma
for E being set
for T being LeftOperation of S,E
for x, y being Element of E holds
( x,y are_conjugated_under T iff ex s being Element of S st y = (T ^ s) . x );
theorem Th4:
theorem Th5:
theorem Th6:
:: deftheorem defines the_orbit_of GROUP_10:def 15 :
for S being non empty unital multMagma
for E being non empty set
for T being LeftOperation of S,E
for x being Element of E holds the_orbit_of (x,T) = { y where y is Element of E : x,y are_conjugated_under T } ;
theorem
canceled;
theorem Th8:
theorem Th9:
theorem Th10:
:: deftheorem defines the_orbits_of GROUP_10:def 16 :
for G being Group
for E being non empty set
for T being LeftOperation of G,E holds the_orbits_of T = { X where X is Subset of E : ex x being Element of E st X = the_orbit_of (x,T) } ;
begin
:: deftheorem Def17 defines -group GROUP_10:def 17 :
for p being natural number
for G being Group holds
( G is p -group iff ex r being natural number st card G = p |^ r );
Lm1:
for n being natural number st n > 0 holds
card (INT.Group n) = n
theorem Th11:
begin
:: deftheorem GROUP_10:def 18 :
canceled;
:: deftheorem Def19 defines is_Sylow_p-subgroup_of_prime GROUP_10:def 19 :
for p being natural number
for G being Group
for P being Subgroup of G holds
( P is_Sylow_p-subgroup_of_prime p iff ( P is p -group & not p divides index P ) );
Lm2:
for n being natural number
for p being natural prime number st n <> 0 holds
ex m, r being natural number st
( n = (p |^ r) * m & not p divides m )
Lm3:
for X, Y being non empty set holds card { [:X,{y}:] where y is Element of Y : verum } = card Y
Lm4:
for G being finite Group
for E, T being non empty set
for LO being LeftOperation of G,E
for n being natural number st LO = the_left_operation_of (G,T) & n = card G & E = [: the carrier of G,T:] & card n c= card E holds
the_fixed_points_of (the_extension_of_left_operation_of (n,LO)) = { [: the carrier of G,{t}:] where t is Element of T : verum }
Lm5:
for n, m, r being natural number
for p being natural prime number st n = (p |^ r) * m & not p divides m holds
(n choose (p |^ r)) mod p <> 0
Lm6:
for r, n, m1, m2 being natural number
for p being natural prime number st (p |^ r) * n = m1 * m2 & not p divides n & not p divides m2 holds
p |^ r divides m1
Lm7:
for G being Group
for A being non empty Subset of G
for x being Element of G holds card A = card (A * x)
theorem Th12:
Lm8:
for n, r being natural number
for p being natural prime number st n divides p |^ r & n > 1 holds
p divides n
theorem
theorem Th14:
:: deftheorem defines the_sylow_p-subgroups_of_prime GROUP_10:def 20 :
for G being Group
for p being natural number holds the_sylow_p-subgroups_of_prime (p,G) = { H where H is Element of Subgroups G : ex P being strict Subgroup of G st
( P = H & P is_Sylow_p-subgroup_of_prime p ) } ;
definition
let G be
finite Group;
let p be
natural prime number ;
let H be
Subgroup of
G;
let h be
Element of
H;
func the_left_translation_of (
h,
p)
-> Function of
(the_sylow_p-subgroups_of_prime (p,G)),
(the_sylow_p-subgroups_of_prime (p,G)) means :
Def21:
for
P1 being
Element of
the_sylow_p-subgroups_of_prime (
p,
G) ex
P2 being
Element of
the_sylow_p-subgroups_of_prime (
p,
G) ex
H1,
H2 being
strict Subgroup of
G ex
g being
Element of
G st
(
P2 = it . P1 &
P1 = H1 &
P2 = H2 &
h " = g &
H2 = H1 |^ g );
existence
ex b1 being Function of (the_sylow_p-subgroups_of_prime (p,G)),(the_sylow_p-subgroups_of_prime (p,G)) st
for P1 being Element of the_sylow_p-subgroups_of_prime (p,G) ex P2 being Element of the_sylow_p-subgroups_of_prime (p,G) ex H1, H2 being strict Subgroup of G ex g being Element of G st
( P2 = b1 . P1 & P1 = H1 & P2 = H2 & h " = g & H2 = H1 |^ g )
uniqueness
for b1, b2 being Function of (the_sylow_p-subgroups_of_prime (p,G)),(the_sylow_p-subgroups_of_prime (p,G)) st ( for P1 being Element of the_sylow_p-subgroups_of_prime (p,G) ex P2 being Element of the_sylow_p-subgroups_of_prime (p,G) ex H1, H2 being strict Subgroup of G ex g being Element of G st
( P2 = b1 . P1 & P1 = H1 & P2 = H2 & h " = g & H2 = H1 |^ g ) ) & ( for P1 being Element of the_sylow_p-subgroups_of_prime (p,G) ex P2 being Element of the_sylow_p-subgroups_of_prime (p,G) ex H1, H2 being strict Subgroup of G ex g being Element of G st
( P2 = b2 . P1 & P1 = H1 & P2 = H2 & h " = g & H2 = H1 |^ g ) ) holds
b1 = b2
end;
:: deftheorem Def21 defines the_left_translation_of GROUP_10:def 21 :
for G being finite Group
for p being natural prime number
for H being Subgroup of G
for h being Element of H
for b5 being Function of (the_sylow_p-subgroups_of_prime (p,G)),(the_sylow_p-subgroups_of_prime (p,G)) holds
( b5 = the_left_translation_of (h,p) iff for P1 being Element of the_sylow_p-subgroups_of_prime (p,G) ex P2 being Element of the_sylow_p-subgroups_of_prime (p,G) ex H1, H2 being strict Subgroup of G ex g being Element of G st
( P2 = b5 . P1 & P1 = H1 & P2 = H2 & h " = g & H2 = H1 |^ g ) );
definition
let G be
finite Group;
let p be
natural prime number ;
let H be
Subgroup of
G;
func the_left_operation_of (
H,
p)
-> LeftOperation of
H,
(the_sylow_p-subgroups_of_prime (p,G)) means :
Def22:
for
h being
Element of
H holds
it . h = the_left_translation_of (
h,
p);
existence
ex b1 being LeftOperation of H,(the_sylow_p-subgroups_of_prime (p,G)) st
for h being Element of H holds b1 . h = the_left_translation_of (h,p)
uniqueness
for b1, b2 being LeftOperation of H,(the_sylow_p-subgroups_of_prime (p,G)) st ( for h being Element of H holds b1 . h = the_left_translation_of (h,p) ) & ( for h being Element of H holds b2 . h = the_left_translation_of (h,p) ) holds
b1 = b2
end;
:: deftheorem Def22 defines the_left_operation_of GROUP_10:def 22 :
for G being finite Group
for p being natural prime number
for H being Subgroup of G
for b4 being LeftOperation of H,(the_sylow_p-subgroups_of_prime (p,G)) holds
( b4 = the_left_operation_of (H,p) iff for h being Element of H holds b4 . h = the_left_translation_of (h,p) );
Lm9:
for G, H being finite Group
for p being natural prime number
for I being Subgroup of G
for P being Subgroup of H st I = P & I is_Sylow_p-subgroup_of_prime p & H is Subgroup of G holds
P is_Sylow_p-subgroup_of_prime p
theorem
begin
theorem
theorem
theorem
theorem