The Transfer Homomorphism #
In this file we construct the transfer homomorphism.
Main definitions #
diff ϕ S T
: The difference of two left transversalsS
andT
under the homomorphismϕ
.transfer ϕ
: The transfer homomorphism induced byϕ
.transferCenterPow
: The transfer homomorphismG →* center G
.
Main results #
transferCenterPow_apply
: The transfer homomorphismG →* center G
is given byg ↦ g ^ (center G).index
.ker_transferSylow_isComplement'
: Burnside's transfer (or normalp
-complement) theorem: IfhP : N(P) ≤ C(P)
, then(transfer P hP).ker
is a normalp
-complement.
The difference of two left transversals
Equations
- Subgroup.leftTransversals.diff ϕ S T = ∏ q : G ⧸ H, ϕ ⟨(↑((Subgroup.MemLeftTransversals.toEquiv ⋯) q))⁻¹ * ↑((Subgroup.MemLeftTransversals.toEquiv ⋯) q), ⋯⟩
Instances For
The difference of two left transversals
Equations
- AddSubgroup.leftTransversals.diff ϕ S T = ∑ q : G ⧸ H, ϕ ⟨-↑((AddSubgroup.MemLeftTransversals.toEquiv ⋯) q) + ↑((AddSubgroup.MemLeftTransversals.toEquiv ⋯) q), ⋯⟩
Instances For
The transfer transversal as a function. Given a ⟨g⟩
-orbit q₀, g • q₀, ..., g ^ (m - 1) • q₀
in G ⧸ H
, an element g ^ k • q₀
is mapped to g ^ k • g₀
for a fixed choice of
representative g₀
of q₀
.
Equations
- H.transferFunction g q = g ^ ((H.quotientEquivSigmaZMod g) q).snd.cast * Quotient.out (Quotient.out ((H.quotientEquivSigmaZMod g) q).fst)
Instances For
The transfer transversal as a set. Contains elements of the form g ^ k • g₀
for fixed choices
of representatives g₀
of fixed choices of representatives q₀
of ⟨g⟩
-orbits in G ⧸ H
.
Instances For
The transfer transversal. Contains elements of the form g ^ k • g₀
for fixed choices
of representatives g₀
of fixed choices of representatives q₀
of ⟨g⟩
-orbits in G ⧸ H
.
Equations
- H.transferTransversal g = ⟨H.transferSet g, ⋯⟩
Instances For
Given ϕ : H →* A
from H : Subgroup G
to a commutative group A
,
the transfer homomorphism is transfer ϕ : G →* A
.
Equations
- ϕ.transfer = { toFun := fun (g : G) => Subgroup.leftTransversals.diff ϕ default (g • default), map_one' := ⋯, map_mul' := ⋯ }
Instances For
Given ϕ : H →+ A
from H : AddSubgroup G
to an additive commutative group A
,
the transfer homomorphism is transfer ϕ : G →+ A
.
Equations
- ϕ.transfer = { toFun := fun (g : G) => AddSubgroup.leftTransversals.diff ϕ default (g +ᵥ default), map_zero' := ⋯, map_add' := ⋯ }
Instances For
Explicit computation of the transfer homomorphism.
The transfer homomorphism G →* center G
.
Equations
- MonoidHom.transferCenterPow G = { toFun := fun (g : G) => ⟨g ^ (Subgroup.center G).index, ⋯⟩, map_one' := ⋯, map_mul' := ⋯ }
Instances For
The homomorphism G →* P
in Burnside's transfer theorem.
Equations
- MonoidHom.transferSylow P hP = (MonoidHom.id ↥↑P).transfer
Instances For
Auxiliary lemma in order to state transferSylow_eq_pow
.
Burnside's normal p-complement theorem: If N(P) ≤ C(P)
, then P
has a normal
complement.