Primitive actions #
Definitions #
MulAction.IsPreprimitive G X
A structure that says that the action of a typeG
on a typeX
(defined by an instanceSMul G X
) is preprimitive, namely, it is pretransitive and the only blocks are ⊤ and subsingletons. (The pretransitivity assumption is essentially trivial, because orbits are blocks, unless the action itself is trivial.)The notion which is introduced in classical books on group theory is restricted to group actions. In fact, it may be irrelevant if the action is degenerate, when “trivial blocks” might not be blocks. Moreover, the classical notion is primitive, which further assumes that
X
is not empty.MulAction.IsQuasiPreprimitive G X
A structure that says that the action of the groupG
on the typeX
is quasipreprimitive, namely, normal subgroups ofG
which act nontrivially act pretransitively.We prove some straightforward theorems that relate preprimitivity under equivariant maps, for images and preimages.
Relation with stabilizers #
MulAction.isSimpleOrderBlockMem_iff_isPreprimitive
relates primitivity and the fact that the inclusion order on blocks containing is simple.MulAction.isCoatom_stabilizer_iff_preprimitive
An action is preprimitive iff the stabilizers of points are maximal subgroups.MulAction.IsPreprimitive.isCoatom_stabilizer_of_isPreprimitive
Stabilizers of points under a preprimitive action are maximal subgroups.
Relation with normal subgroups #
MulAction.IsPreprimitive.isQuasipreprimitive
Preprimitive actions are quasipreprimitive.
An additive action is preprimitive if it is pretransitive and the only blocks are the trivial ones
- isTrivialBlock_of_isBlock {B : Set X} : IsBlock G B → IsTrivialBlock B
An action is preprimitive if it is pretransitive and the only blocks are the trivial ones
Instances
An action is preprimitive if it is pretransitive and the only blocks are the trivial ones
- isTrivialBlock_of_isBlock {B : Set X} : IsBlock G B → IsTrivialBlock 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
- isPretransitive_of_normal {N : AddSubgroup G} [N.Normal] : fixedPoints (↥N) X ≠ Set.univ → IsPretransitive (↥N) X
Instances
An action of a group is quasipreprimitive if any normal subgroup that has no fixed point acts pretransitively
- isPretransitive_of_normal {N : Subgroup G} [N.Normal] : fixedPoints (↥N) X ≠ Set.univ → IsPretransitive (↥N) X
Instances
If the action is pretransitive, then the trivial blocks condition implies preprimitivity (based condition)
If the action is pretransitive, then the trivial blocks condition implies preprimitivity (based condition)
If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic) (based condition)
If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic) (based condition)
If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic)
If the action is not trivial, then the trivial blocks condition implies preprimitivity (pretransitivity is automatic)
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 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.
In a preprimitive action, any normal subgroup that acts nontrivially is pretransitive (Wielandt, th. 7.1).
In a preprimitive additive action, any normal subgroup that acts nontrivially is pretransitive (Wielandt, th. 7.1).