Documentation

Mathlib.GroupTheory.GroupAction.Primitive

Primitive actions #

Definitions #

Relation with stabilizers #

Relation with normal subgroups #

class AddAction.IsPreprimitive (G : Type u_1) (X : Type u_2) [VAdd G X] extends AddAction.IsPretransitive G X :

An additive action is preprimitive if it is pretransitive and the only blocks are the trivial ones

Instances
    class MulAction.IsPreprimitive (G : Type u_1) (X : Type u_2) [SMul G X] extends MulAction.IsPretransitive G X :

    An action is preprimitive if it is pretransitive and the only blocks are the trivial ones

    • exists_smul_eq (x y : X) : ∃ (g : G), g x = y
    • isTrivialBlock_of_isBlock {B : Set X} : IsBlock G BIsTrivialBlock B

      An action is preprimitive if it is pretransitive and the only blocks are the trivial ones

    Instances

      An additive action of an additive group is quasipreprimitive if any normal subgroup that has no fixed point acts pretransitively

      Instances
        class MulAction.IsQuasiPreprimitive (G : Type u_1) (X : Type u_2) [Group G] [MulAction G X] extends MulAction.IsPretransitive G X :

        An action of a group is quasipreprimitive if any normal subgroup that has no fixed point acts pretransitively

        Instances
          theorem MulAction.IsBlock.subsingleton_or_eq_univ {G : Type u_1} {X : Type u_2} [SMul G X] [IsPreprimitive G X] {B : Set X} (hB : IsBlock G B) :
          theorem AddAction.IsBlock.subsingleton_or_eq_univ {G : Type u_1} {X : Type u_2} [VAdd G X] [IsPreprimitive G X] {B : Set X} (hB : IsBlock G B) :
          theorem MulAction.IsPreprimitive.of_isTrivialBlock_base {G : Type u_1} {X : Type u_2} [Group G] [MulAction G X] [IsPretransitive G X] (a : X) (H : ∀ {B : Set X}, a BIsBlock G BIsTrivialBlock B) :

          If the action is pretransitive, then the trivial blocks condition implies preprimitivity (based condition)

          theorem AddAction.IsPreprimitive.of_isTrivialBlock_base {G : Type u_1} {X : Type u_2} [AddGroup G] [AddAction G X] [IsPretransitive G X] (a : X) (H : ∀ {B : Set X}, a BIsBlock G BIsTrivialBlock B) :

          If the action is pretransitive, then the trivial blocks condition implies preprimitivity (based condition)

          theorem MulAction.IsPreprimitive.of_isTrivialBlock_of_not_mem_fixedPoints {G : Type u_1} {X : Type u_2} [Group G] [MulAction G X] {a : X} (ha : afixedPoints G X) (H : ∀ ⦃B : Set X⦄, a BIsBlock G BIsTrivialBlock B) :

          If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic) (based condition)

          theorem AddAction.IsPreprimitive.of_isTrivialBlock_of_not_mem_fixedPoints {G : Type u_1} {X : Type u_2} [AddGroup G] [AddAction G X] {a : X} (ha : afixedPoints G X) (H : ∀ ⦃B : Set X⦄, a BIsBlock G BIsTrivialBlock B) :

          If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic) (based condition)

          theorem MulAction.mk' {G : Type u_1} {X : Type u_2} [Group G] [MulAction G X] (Hnt : fixedPoints G X ) (H : ∀ {B : Set X}, IsBlock G BIsTrivialBlock B) :

          If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic)

          theorem AddAction.mk' {G : Type u_1} {X : Type u_2} [AddGroup G] [AddAction G X] (Hnt : fixedPoints G X ) (H : ∀ {B : Set X}, IsBlock G BIsTrivialBlock B) :

          If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic)

          theorem MulAction.IsPreprimitive.of_surjective {M : Type u_3} [Group M] {α : Type u_4} [MulAction M α] {N : Type u_5} {β : Type u_6} [Group N] [MulAction N β] {φ : M →* N} {f : α →ₑ[φ] β} [IsPreprimitive M α] (hf : Function.Surjective f) :
          theorem AddAction.IsPreprimitive.of_surjective {M : Type u_3} [AddGroup M] {α : Type u_4} [AddAction M α] {N : Type u_5} {β : Type u_6} [AddGroup N] [AddAction N β] {φ : M →+ N} {f : α →ₑ[φ] β} [IsPreprimitive M α] (hf : Function.Surjective f) :
          theorem MulAction.isPreprimitive_congr {M : Type u_3} [Group M] {α : Type u_4} [MulAction M α] {N : Type u_5} {β : Type u_6} [Group N] [MulAction N β] {φ : M →* N} {f : α →ₑ[φ] β} (hφ : Function.Surjective φ) (hf : Function.Bijective f) :
          theorem AddAction.isPreprimitive_congr {M : Type u_3} [AddGroup M] {α : Type u_4} [AddAction M α] {N : Type u_5} {β : Type u_6} [AddGroup N] [AddAction N β] {φ : M →+ N} {f : α →ₑ[φ] β} (hφ : Function.Surjective φ) (hf : Function.Bijective f) :
          @[simp]

          A pretransitive action on a nontrivial type is preprimitive iff the set of blocks containing a given element is a simple order

          @[simp]

          A pretransitive action on a nontrivial type is preprimitive iff the set of blocks containing a given element is a simple order

          A pretransitive action is preprimitive iff the stabilizer of any point is a maximal subgroup (Wielandt, th. 7.5)

          A pretransitive action is preprimitive iff the stabilizer of any point is a maximal subgroup (Wielandt, th. 7.5)

          In a preprimitive action, stabilizers are maximal subgroups

          In a preprimitive action, stabilizers are maximal subgroups.

          @[instance 100]

          In a preprimitive action, any normal subgroup that acts nontrivially is pretransitive (Wielandt, th. 7.1).

          @[instance 100]

          In a preprimitive additive action, any normal subgroup that acts nontrivially is pretransitive (Wielandt, th. 7.1).