theorem
mul_action.mem_fixed_points_iff_card_orbit_eq_one
{G : Type u}
{α : Type v}
[group G]
[mul_action G α]
{a : α}
[fintype ↥(mul_action.orbit G a)] :
a ∈ mul_action.fixed_points G α ↔ fintype.card ↥(mul_action.orbit G a) = 1
theorem
mul_action.card_modeq_card_fixed_points
{G : Type u}
{α : Type v}
[group G]
[mul_action G α]
[fintype α]
[fintype G]
[fintype ↥(mul_action.fixed_points G α)]
(p : ℕ)
{n : ℕ}
[hp : fact (nat.prime p)] :
fintype.card G = p ^ n → fintype.card α ≡ fintype.card ↥(mul_action.fixed_points G α) [MOD p]
theorem
quotient_group.card_preimage_mk
{G : Type u}
[group G]
[fintype G]
(s : subgroup G)
(t : set (quotient_group.quotient s)) :
theorem
sylow.mem_vectors_prod_eq_one_iff
{G : Type u}
[group G]
{n : ℕ}
(v : vector G (n + 1)) :
v ∈ sylow.vectors_prod_eq_one G (n + 1) ↔ v ∈ set.range (sylow.mk_vector_prod_eq_one n)
def
sylow.rotate_vectors_prod_eq_one
(G : Type u_1)
[group G]
(n : ℕ) :
multiplicative (zmod n) → ↥(sylow.vectors_prod_eq_one G n) → ↥(sylow.vectors_prod_eq_one G n)
The rotation action of zmod n
(viewed as multiplicative group) on
vectors_prod_eq_one G n
, where G
is a multiplicative group.
@[instance]
def
sylow.rotate_vectors_prod_eq_one.mul_action
{G : Type u}
[group G]
(n : ℕ)
[fact (0 < n)] :
mul_action (multiplicative (zmod n)) ↥(sylow.vectors_prod_eq_one G n)
Equations
- sylow.rotate_vectors_prod_eq_one.mul_action n = {to_has_scalar := {smul := sylow.rotate_vectors_prod_eq_one G n}, one_smul := _, mul_smul := _}
theorem
sylow.one_mem_vectors_prod_eq_one
{G : Type u}
[group G]
(n : ℕ) :
vector.repeat 1 n ∈ sylow.vectors_prod_eq_one G n
theorem
sylow.one_mem_fixed_points_rotate
{G : Type u}
[group G]
(n : ℕ)
[fact (0 < n)] :
⟨vector.repeat 1 n, _⟩ ∈ mul_action.fixed_points (multiplicative (zmod n)) ↥(sylow.vectors_prod_eq_one G n)
theorem
sylow.mem_fixed_points_mul_left_cosets_iff_mem_normalizer
{G : Type u}
[group G]
{H : subgroup G}
[fintype ↥↑H]
{x : G} :
↑x ∈ mul_action.fixed_points ↥H (quotient_group.quotient H) ↔ x ∈ H.normalizer
def
sylow.fixed_points_mul_left_cosets_equiv_quotient
{G : Type u}
[group G]
(H : subgroup G)
[fintype ↥↑H] :
Equations
- sylow.fixed_points_mul_left_cosets_equiv_quotient H = equiv.subtype_quotient_equiv_quotient_subtype ↑(H.normalizer) (mul_action.fixed_points ↥H (quotient (id {r := λ (x y : G), x⁻¹ * y ∈ H, iseqv := _}))) _ _