More lemmas about group actions #
This file contains lemmas about group actions that require more imports than
Mathlib.Algebra.Group.Action.Defs offers.
Given an action of a group α on β, each g : α defines a permutation of β.
Equations
- MulAction.toPerm a = { toFun := fun (x : β) => a • x, invFun := fun (x : β) => a⁻¹ • x, left_inv := ⋯, right_inv := ⋯ }
Instances For
Given an action of an additive group α on β, each g : α defines a permutation of β.
Equations
- AddAction.toPerm a = { toFun := fun (x : β) => a +ᵥ x, invFun := fun (x : β) => -a +ᵥ x, left_inv := ⋯, right_inv := ⋯ }
Instances For
theorem
MulAction.toPerm_injective
{α : Type u_1}
{β : Type u_2}
[Group α]
[MulAction α β]
[FaithfulSMul α β]
 :
MulAction.toPerm is injective on faithful actions.
theorem
AddAction.toPerm_injective
{α : Type u_1}
{β : Type u_2}
[AddGroup α]
[AddAction α β]
[FaithfulVAdd α β]
 :
AddAction.toPerm is injective on faithful actions.
Given an action of a group α on a set β, each g : α defines a permutation of β.
Equations
- MulAction.toPermHom α β = { toFun := MulAction.toPerm, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Given an action of an additive group α on a set β, each g : α defines a permutation of
β.
Equations
Instances For
@[simp]
theorem
AddAction.toPermHom_apply_symm_apply
(β : Type u_2)
(α : Type u_3)
[AddGroup α]
[AddAction α β]
(a : α)
(x : β)
 :
Equiv.Perm.applyMulAction is faithful.
theorem
MulAction.bijective
{α : Type u_1}
{β : Type u_2}
[Group α]
[MulAction α β]
(g : α)
 :
Function.Bijective fun (x : β) => g • x
theorem
AddAction.bijective
{α : Type u_1}
{β : Type u_2}
[AddGroup α]
[AddAction α β]
(g : α)
 :
Function.Bijective fun (x : β) => g +ᵥ x
theorem
MulAction.injective
{α : Type u_1}
{β : Type u_2}
[Group α]
[MulAction α β]
(g : α)
 :
Function.Injective fun (x : β) => g • x
theorem
AddAction.injective
{α : Type u_1}
{β : Type u_2}
[AddGroup α]
[AddAction α β]
(g : α)
 :
Function.Injective fun (x : β) => g +ᵥ x
theorem
MulAction.surjective
{α : Type u_1}
{β : Type u_2}
[Group α]
[MulAction α β]
(g : α)
 :
Function.Surjective fun (x : β) => g • x
theorem
AddAction.surjective
{α : Type u_1}
{β : Type u_2}
[AddGroup α]
[AddAction α β]
(g : α)
 :
Function.Surjective fun (x : β) => g +ᵥ x
def
arrowAction
{G : Type u_3}
{A : Type u_4}
{B : Type u_5}
[DivisionMonoid G]
[MulAction G A]
 :
MulAction G (A → B)
If G acts on A, then it acts also on A → B, by (g • F) a = F (g⁻¹ • a).
Equations
- arrowAction = MulAction.mk ⋯ ⋯
Instances For
def
arrowAddAction
{G : Type u_3}
{A : Type u_4}
{B : Type u_5}
[SubtractionMonoid G]
[AddAction G A]
 :
AddAction G (A → B)
If G acts on A, then it acts also on A → B, by (g +ᵥ F) a = F (g⁻¹ +ᵥ a)
Equations
Instances For
@[simp]
theorem
isUnit_smul_iff
{α : Type u_1}
{β : Type u_2}
[Group α]
[Monoid β]
[MulAction α β]
[SMulCommClass α β β]
[IsScalarTower α β β]
(g : α)
(m : β)
 :