Monoids of endomorphisms, groups of automorphisms #
This file defines
- the endomorphism monoid structure on
Function.End α := α → α
- the endomorphism monoid structure on
Monoid.End M := M →* M
andAddMonoid.End M := M →+ M
- the automorphism group structure on
Equiv.Perm α := α ≃ α
- the automorphism group structure on
MulAut M := M ≃* M
andAddAut M := M ≃+ M
.
Implementation notes #
The definition of multiplication in the endomorphism monoids and automorphism groups agrees with
function composition, and multiplication in CategoryTheory.End
, but not with
CategoryTheory.comp
.
Tags #
end monoid, aut group
Type endomorphisms #
The monoid of endomorphisms.
Note that this is generalized by CategoryTheory.End
to categories other than Type u
.
Equations
- Function.End α = (α → α)
Instances For
Equations
- instMonoidEnd = Monoid.mk ⋯ ⋯ (fun (n : ℕ) (f : Function.End α) => f^[n]) ⋯ ⋯
Equations
- instInhabitedEnd = { default := 1 }
Monoid endomorphisms #
Equations
- Equiv.Perm.instOne = { one := Equiv.refl α }
Equations
- Equiv.Perm.instMul = { mul := fun (f g : Equiv.Perm α) => Equiv.trans g f }
Equations
- Equiv.Perm.instInv = { inv := Equiv.symm }
Equations
- Equiv.Perm.instPowNat = { pow := fun (f : Equiv.Perm α) (n : ℕ) => { toFun := (⇑f)^[n], invFun := (⇑(Equiv.symm f))^[n], left_inv := ⋯, right_inv := ⋯ } }
Equations
The permutation of a type is equivalent to the units group of the endomorphisms monoid of this type.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Lift a monoid homomorphism f : G →* Function.End α
to a monoid homomorphism
f : G →* Equiv.Perm α
.
Equations
Instances For
Lemmas about mixing Perm
with Equiv
. Because we have multiple ways to express
Equiv.refl
, Equiv.symm
, and Equiv.trans
, we want simp lemmas for every combination.
The assumption made here is that if you're using the group structure, you want to preserve it after
simp.
Lemmas about Equiv.Perm.sumCongr
re-expressed via the group structure.
Equiv.Perm.sumCongr
as a MonoidHom
, with its two arguments bundled into a single Prod
.
This is particularly useful for its MonoidHom.range
projection, which is the subgroup of
permutations which do not exchange elements between α
and β
.
Equations
- Equiv.Perm.sumCongrHom α β = { toFun := fun (a : Equiv.Perm α × Equiv.Perm β) => a.1.sumCongr a.2, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Lemmas about Equiv.Perm.sigmaCongrRight
re-expressed via the group structure.
Equiv.Perm.sigmaCongrRight
as a MonoidHom
.
This is particularly useful for its MonoidHom.range
projection, which is the subgroup of
permutations which do not exchange elements between fibers.
Equations
- Equiv.Perm.sigmaCongrRightHom β = { toFun := Equiv.Perm.sigmaCongrRight, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Equiv.Perm.subtypeCongr
as a MonoidHom
.
Equations
- Equiv.Perm.subtypeCongrHom p = { toFun := fun (pair : Equiv.Perm { a : α // p a } × Equiv.Perm { a : α // ¬p a }) => pair.1.subtypeCongr pair.2, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Lemmas about Equiv.Perm.extendDomain
re-expressed via the group structure.
extendDomain
as a group homomorphism
Equations
- Equiv.Perm.extendDomainHom f = { toFun := fun (e : Equiv.Perm α) => e.extendDomain f, map_one' := ⋯, map_mul' := ⋯ }
Instances For
If the permutation f
fixes the subtype {x // p x}
, then this returns the permutation
on {x // p x}
induced by f
.
Equations
Instances For
The inclusion map of permutations on a subtype of α
into permutations of α
,
fixing the other points.
Equations
- Equiv.Perm.ofSubtype = { toFun := fun (f : Equiv.Perm (Subtype p)) => f.extendDomain (Equiv.refl (Subtype p)), map_one' := ⋯, map_mul' := ⋯ }
Instances For
Permutations on a subtype are equivalent to permutations on the original type that fix pointwise the rest.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Left-multiplying a permutation with swap i j
twice gives the original permutation.
This specialization of swap_mul_self
is useful when using cosets of permutations.
Right-multiplying a permutation with swap i j
twice gives the original permutation.
This specialization of swap_mul_self
is useful when using cosets of permutations.
A stronger version of mul_right_injective
A stronger version of mul_left_injective
The group operation on multiplicative automorphisms is defined by g h => MulEquiv.trans h g
.
This means that multiplication agrees with composition, (g*h)(x) = g (h x)
.
Equations
Equations
- MulAut.instInhabited M = { default := 1 }
Monoid hom from the group of multiplicative automorphisms to the group of permutations.
Equations
- MulAut.toPerm M = { toFun := MulEquiv.toEquiv, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Group conjugation, MulAut.conj g h = g * h * g⁻¹
, as a monoid homomorphism
mapping multiplication in G
into multiplication in the automorphism group MulAut G
.
See also the type ConjAct G
for any group G
, which has a MulAction (ConjAct G) G
instance
where conj G
acts on G
by conjugation.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The group operation on additive automorphisms is defined by g h => AddEquiv.trans h g
.
This means that multiplication agrees with composition, (g*h)(x) = g (h x)
.
Equations
- AddAut.group A = Group.mk ⋯
Equations
- AddAut.instInhabited A = { default := 1 }
Monoid hom from the group of multiplicative automorphisms to the group of permutations.
Equations
- AddAut.toPerm A = { toFun := AddEquiv.toEquiv, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Additive group conjugation, AddAut.conj g h = g + h - g
, as an additive monoid
homomorphism mapping addition in G
into multiplication in the automorphism group AddAut G
(written additively in order to define the map).
Equations
- One or more equations did not get rendered due to their size.
Instances For
Multiplicative G
and G
have isomorphic automorphism groups.
Equations
- MulAutMultiplicative G = { toEquiv := AddEquiv.toMultiplicative.symm, map_mul' := ⋯ }
Instances For
Additive G
and G
have isomorphic automorphism groups.
Equations
- AddAutAdditive G = { toEquiv := MulEquiv.toAdditive.symm, map_mul' := ⋯ }