Documentation

Mathlib.RepresentationTheory.Rep

Rep k G is the category of k-linear representations of G. #

If V : Rep k G, there is a coercion that allows you to treat V as a type, and this type comes equipped with a Module k V instance. Also V.ρ gives the homomorphism G →* (V →ₗ[k] V).

Conversely, given a homomorphism ρ : G →* (V →ₗ[k] V), you can construct the bundled representation as Rep.of ρ.

We construct the categorical equivalence Rep k G ≌ ModuleCat (MonoidAlgebra k G). We verify that Rep k G is a k-linear abelian symmetric monoidal category with all (co)limits.

@[reducible, inline]
noncomputable abbrev Rep (k G : Type u) [Ring k] [Monoid G] :
Type (u + 1)

The category of k-linear representations of a monoid G.

Equations
noncomputable instance instLinearRep (k G : Type u) [CommRing k] [Monoid G] :
Equations
noncomputable instance Rep.instCoeSortType {k G : Type u} [CommRing k] [Monoid G] :
CoeSort (Rep k G) (Type u)
Equations
noncomputable instance Rep.instModuleCarrierVModuleCat {k G : Type u} [CommRing k] [Monoid G] (V : Rep k G) :
Module k V.V
Equations
noncomputable def Rep.ρ {k G : Type u} [CommRing k] [Monoid G] (V : Rep k G) :
Representation k G V.V

Specialize the existing Action.ρ, changing the type to Representation k G V.

Equations
@[reducible, inline]
noncomputable abbrev Rep.of {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) :
Rep k G

Lift an unbundled representation to Rep.

Equations
theorem Rep.coe_of {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) :
(of ρ).V = V
@[simp]
theorem Rep.of_ρ {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] (ρ : G →* V →ₗ[k] V) :
(of ρ).ρ = ρ
@[simp]
theorem Rep.ρ_hom {k G : Type u} [CommRing k] [Monoid G] {X : Rep k G} (g : G) :
@[simp]
theorem Rep.ofHom_ρ {k G : Type u} [CommRing k] [Monoid G] {X : Rep k G} (g : G) :
@[deprecated Representation.inv_self_apply (since := "2025-05-09")]
theorem Rep.ρ_inv_self_apply {k : Type u} [CommRing k] {G : Type u} [Group G] (A : Rep k G) (g : G) (x : A.V) :
(A.ρ g⁻¹) ((A.ρ g) x) = x
@[deprecated Representation.self_inv_apply (since := "2025-05-09")]
theorem Rep.ρ_self_inv_apply {k : Type u} [CommRing k] {G : Type u} [Group G] {A : Rep k G} (g : G) (x : A.V) :
(A.ρ g) ((A.ρ g⁻¹) x) = x
theorem Rep.hom_comm_apply {k G : Type u} [CommRing k] [Monoid G] {A B : Rep k G} (f : A B) (g : G) (x : A.V) :
@[reducible, inline]
noncomputable abbrev Rep.trivial (k G : Type u) [CommRing k] [Monoid G] (V : Type u) [AddCommGroup V] [Module k V] :
Rep k G

The trivial k-linear G-representation on a k-module V.

Equations
theorem Rep.trivial_def {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] (g : G) :
noncomputable def Rep.trivialFunctor (k G : Type u) [CommRing k] [Monoid G] :

The functor equipping a module with the trivial representation.

Equations
@[simp]
theorem Rep.trivialFunctor_map_hom (k G : Type u) [CommRing k] [Monoid G] {X✝ Y✝ : ModuleCat k} (f : X✝ Y✝) :
((trivialFunctor k G).map f).hom = f
@[simp]
theorem Rep.trivialFunctor_obj_V (k G : Type u) [CommRing k] [Monoid G] (V : ModuleCat k) :
((trivialFunctor k G).obj V).V = V
@[reducible, inline]
noncomputable abbrev Rep.IsTrivial {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) :

A predicate for representations that fix every element.

Equations
instance Rep.instIsTrivialTrivial {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] :
instance Rep.instIsTrivialOfOfIsTrivial {k G : Type u} [CommRing k] [Monoid G] {V : Type u} [AddCommGroup V] [Module k V] (ρ : Representation k G V) [ρ.IsTrivial] :
@[reducible, inline]
noncomputable abbrev Rep.ofQuotient {k : Type u} [CommRing k] {G : Type u} [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] [Representation.IsTrivial (MonoidHom.comp A.ρ S.subtype)] :
Rep k (G S)

Given a normal subgroup S ≤ G, a G-representation ρ which is trivial on S factors through G ⧸ S.

Equations
@[reducible, inline]
noncomputable abbrev Rep.resOfQuotientIso {k : Type u} [CommRing k] {G : Type u} [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] [Representation.IsTrivial (MonoidHom.comp A.ρ S.subtype)] :

A G-representation A on which a normal subgroup S ≤ G acts trivially induces a G ⧸ S-representation on A, and composing this with the quotient map G → G ⧸ S gives the original representation by definition. Useful for typechecking.

Equations
@[reducible, inline]
noncomputable abbrev Rep.subrepresentation {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
Rep k G

Given a k-linear G-representation (V, ρ), this is the representation defined by restricting ρ to a G-invariant k-submodule of V.

Equations
noncomputable def Rep.subtype {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
A.subrepresentation W le_comap A

The natural inclusion of a subrepresentation into the ambient representation.

Equations
@[simp]
theorem Rep.subtype_hom {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
@[reducible, inline]
noncomputable abbrev Rep.quotient {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
Rep k G

Given a k-linear G-representation (V, ρ) and a G-invariant k-submodule W ≤ V, this is the representation induced on V ⧸ W by ρ.

Equations
noncomputable def Rep.mkQ {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
A A.quotient W le_comap

The natural projection from a representation to its quotient by a subrepresentation.

Equations
@[simp]
theorem Rep.mkQ_hom {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (W : Submodule k A.V) (le_comap : ∀ (g : G), W Submodule.comap (A.ρ g) W) :
(A.mkQ W le_comap).hom = ModuleCat.ofHom W.mkQ
@[simp]
@[simp]
theorem Rep.res_obj_ρ {k G : Type u} [CommRing k] [Monoid G] {H : Type u} [Monoid H] (f : G →* H) (A : Rep k H) (g : G) :
(ρ ((Action.res (ModuleCat k) f).obj A)) g = A.ρ (f g)
noncomputable def Rep.linearization (k G : Type u) [CommRing k] [Monoid G] :

The monoidal functor sending a type H with a G-action to the induced k-linear G-representation on k[H].

Equations
@[simp]
theorem Rep.linearization_obj_ρ {k G : Type u} [CommRing k] [Monoid G] (X : Action (Type u) G) (g : G) (x : X.V →₀ k) :
(((linearization k G).obj X).ρ g) x = (Finsupp.lmapDomain k k (X.ρ g)) x
theorem Rep.linearization_of {k G : Type u} [CommRing k] [Monoid G] (X : Action (Type u) G) (g : G) (x : X.V) :
(((linearization k G).obj X).ρ g) (Finsupp.single x 1) = Finsupp.single (X.ρ g x) 1
theorem Rep.linearization_single {k G : Type u} [CommRing k] [Monoid G] (X : Action (Type u) G) (g : G) (x : X.V) (r : k) :
(((linearization k G).obj X).ρ g) (Finsupp.single x r) = Finsupp.single (X.ρ g x) r
@[simp]
theorem Rep.linearization_map_hom {k G : Type u} [CommRing k] [Monoid G] {X Y : Action (Type u) G} (f : X Y) :
theorem Rep.linearization_map_hom_single {k G : Type u} [CommRing k] [Monoid G] {X Y : Action (Type u) G} (f : X Y) (x : X.V) (r : k) :
noncomputable def Rep.linearizationTrivialIso (k G : Type u) [CommRing k] [Monoid G] (X : Type u) :
(linearization k G).obj { V := X, ρ := 1 } trivial k G (X →₀ k)

The linearization of a type X on which G acts trivially is the trivial G-representation on k[X].

Equations
@[simp]
theorem Rep.linearizationTrivialIso_inv_hom_hom_apply (k G : Type u) [CommRing k] [Monoid G] (X : Type u) (a : ((linearization k G).obj { V := X, ρ := 1 }).V) :
@[simp]
theorem Rep.linearizationTrivialIso_hom_hom_hom_apply (k G : Type u) [CommRing k] [Monoid G] (X : Type u) (a : ((linearization k G).obj { V := X, ρ := 1 }).V) :
@[reducible, inline]
noncomputable abbrev Rep.ofMulAction (k G : Type u) [CommRing k] [Monoid G] (H : Type u) [MulAction G H] :
Rep k G

Given a G-action on H, this is k[H] bundled with the natural representation G →* End(k[H]) as a term of type Rep k G.

Equations
@[reducible, inline]
noncomputable abbrev Rep.leftRegular (k G : Type u) [CommRing k] [Monoid G] :
Rep k G

The k-linear G-representation on k[G], induced by left multiplication.

Equations
noncomputable def Rep.diagonal (k G : Type u) [CommRing k] [Monoid G] (n : ) :
Rep k G

The k-linear G-representation on k[Gⁿ], induced by left multiplication.

Equations
noncomputable def Rep.linearizationOfMulActionIso (k G : Type u) [CommRing k] [Monoid G] (H : Type u) [MulAction G H] :

The linearization of a type H with a G-action is definitionally isomorphic to the k-linear G-representation on k[H] induced by the G-action on H.

Equations
noncomputable def Rep.ofDistribMulAction (k G A : Type u) [CommRing k] [Monoid G] [AddCommGroup A] [Module k A] [DistribMulAction G A] [SMulCommClass G k A] :
Rep k G

Turns a k-module A with a compatible DistribMulAction of a monoid G into a k-linear G-representation on A.

Equations
@[simp]
theorem Rep.ofDistribMulAction_ρ_apply_apply (k G A : Type u) [CommRing k] [Monoid G] [AddCommGroup A] [Module k A] [DistribMulAction G A] [SMulCommClass G k A] (g : G) (a : A) :
((ofDistribMulAction k G A).ρ g) a = g a
noncomputable def Rep.ofAlgebraAut (R S : Type) [CommRing R] [CommRing S] [Algebra R S] :

Given an R-algebra S, the -linear representation associated to the natural action of S ≃ₐ[R] S on S.

Equations
noncomputable def Rep.ofMulDistribMulAction (M G : Type) [Monoid M] [CommGroup G] [MulDistribMulAction M G] :

Turns a CommGroup G with a MulDistribMulAction of a monoid M into a -linear M-representation on Additive G.

Equations
noncomputable def Rep.ofAlgebraAutOnUnits (R S : Type) [CommRing R] [CommRing S] [Algebra R S] :

Given an R-algebra S, the -linear representation associated to the natural action of S ≃ₐ[R] S on .

Equations
noncomputable def Rep.leftRegularHom {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (x : A.V) :

Given an element x : A, there is a natural morphism of representations k[G] ⟶ A sending g ↦ A.ρ(g)(x).

Equations
@[simp]
theorem Rep.leftRegularHom_hom {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (x : A.V) :
(A.leftRegularHom x).hom = ModuleCat.ofHom ((Finsupp.lift (↑A.V) k G) fun (g : G) => (A.ρ g) x)
theorem Rep.leftRegularHom_hom_single {k G : Type u} [CommRing k] [Monoid G] {A : Rep k G} (g : G) (x : A.V) (r : k) :
noncomputable def Rep.leftRegularHomEquiv {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) :
(leftRegular k G A) ≃ₗ[k] A.V

Given a k-linear G-representation A, there is a k-linear isomorphism between representation morphisms Hom(k[G], A) and A.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem Rep.leftRegularHomEquiv_symm_apply {k G : Type u} [CommRing k] [Monoid G] (A : Rep k G) (x : A.V) :
@[reducible, inline]
noncomputable abbrev Rep.finsupp {k G : Type u} [CommRing k] [Monoid G] (α : Type u) (A : Rep k G) :
Rep k G

The representation on α →₀ A defined pointwise by a representation on A.

Equations
@[reducible, inline]
noncomputable abbrev Rep.free (k G : Type u) [CommRing k] [Monoid G] (α : Type u) :
Rep k G

The representation on α →₀ k[G] defined pointwise by the left regular representation on k[G].

Equations
noncomputable def Rep.freeLift {k G : Type u} [CommRing k] [Monoid G] {α : Type u} (A : Rep k G) [DecidableEq α] (f : αA.V) :
free k G α A

Given f : α → A, the natural representation morphism (α →₀ k[G]) ⟶ A sending single a (single g r) ↦ r • A.ρ g (f a).

Equations
@[simp]
theorem Rep.freeLift_hom {k G : Type u} [CommRing k] [Monoid G] {α : Type u} (A : Rep k G) [DecidableEq α] (f : αA.V) :
(A.freeLift f).hom = ModuleCat.ofHom ((Finsupp.linearCombination k fun (x : α × G) => (A.ρ x.2) (f x.1)) ∘ₗ (Finsupp.finsuppProdLEquiv k).symm)
theorem Rep.freeLift_hom_single_single {k G : Type u} [CommRing k] [Monoid G] {α : Type u} {A : Rep k G} [DecidableEq α] (f : αA.V) (i : α) (g : G) (r : k) :
noncomputable def Rep.freeLiftLEquiv {k G : Type u} [CommRing k] [Monoid G] (α : Type u) (A : Rep k G) [DecidableEq α] :
(free k G α A) ≃ₗ[k] αA.V

The natural linear equivalence between functions α → A and representation morphisms (α →₀ k[G]) ⟶ A.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem Rep.freeLiftLEquiv_apply {k G : Type u} [CommRing k] [Monoid G] (α : Type u) (A : Rep k G) [DecidableEq α] (f : free k G α A) (i : α) :
@[simp]
theorem Rep.freeLiftLEquiv_symm_apply {k G : Type u} [CommRing k] [Monoid G] (α : Type u) (A : Rep k G) [DecidableEq α] (f : αA.V) :
theorem Rep.free_ext {k G : Type u} [CommRing k] [Monoid G] {α : Type u} {A : Rep k G} [DecidableEq α] (f g : free k G α A) (h : ∀ (i : α), (CategoryTheory.ConcreteCategory.hom f.hom) (Finsupp.single i (Finsupp.single 1 1)) = (CategoryTheory.ConcreteCategory.hom g.hom) (Finsupp.single i (Finsupp.single 1 1))) :
f = g

Given representations A, B and a type α, this is the natural representation isomorphism (α →₀ A) ⊗ B ≅ (A ⊗ B) →₀ α sending single x a ⊗ₜ b ↦ single x (a ⊗ₜ b).

Equations
@[simp]
theorem Rep.finsuppTensorLeft_hom_hom {k G : Type u} [CommRing k] [Monoid G] (A B : Rep k G) (α : Type u) [DecidableEq α] :
@[simp]
theorem Rep.finsuppTensorLeft_inv_hom {k G : Type u} [CommRing k] [Monoid G] (A B : Rep k G) (α : Type u) [DecidableEq α] :

Given representations A, B and a type α, this is the natural representation isomorphism A ⊗ (α →₀ B) ≅ (A ⊗ B) →₀ α sending a ⊗ₜ single x b ↦ single x (a ⊗ₜ b).

Equations
@[simp]
theorem Rep.finsuppTensorRight_inv_hom {k G : Type u} [CommRing k] [Monoid G] (A B : Rep k G) (α : Type u) [DecidableEq α] :
@[simp]
theorem Rep.finsuppTensorRight_hom_hom {k G : Type u} [CommRing k] [Monoid G] (A B : Rep k G) (α : Type u) [DecidableEq α] :

The natural isomorphism sending single g r₁ ⊗ single a r₂ ↦ single a (single g r₁r₂).

Equations
  • One or more equations did not get rendered due to their size.
noncomputable def Rep.ihom {k G : Type u} [CommRing k] [Group G] (A : Rep k G) :

Given a k-linear G-representation (A, ρ₁), this is the 'internal Hom' functor sending (B, ρ₂) to the representation Homₖ(A, B) that maps g : G and f : A →ₗ[k] B to (ρ₂ g) ∘ₗ f ∘ₗ (ρ₁ g⁻¹).

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem Rep.ihom_obj_V_carrier {k G : Type u} [CommRing k] [Group G] (A B : Rep k G) :
(A.ihom.obj B).V = (A.V →ₗ[k] B.V)
@[simp]
theorem Rep.ihom_obj_ρ {k G : Type u} [CommRing k] [Group G] (A B : Rep k G) :
@[simp]
theorem Rep.ihom_map_hom {k G : Type u} [CommRing k] [Group G] (A : Rep k G) {X Y : Rep k G} (f : X Y) :
@[simp]
theorem Rep.ihom_obj_V_isModule {k G : Type u} [CommRing k] [Group G] (A B : Rep k G) :
@[simp]
theorem Rep.ihom_obj_ρ_apply {k G : Type u} [CommRing k] [Group G] {A B : Rep k G} (g : G) (x : A.V →ₗ[k] B.V) :
((A.ihom.obj B).ρ g) x = B.ρ g ∘ₗ x ∘ₗ A.ρ g⁻¹
noncomputable def Rep.homEquiv {k G : Type u} [CommRing k] [Group G] (A B C : Rep k G) :

Given a k-linear G-representation A, this is the Hom-set bijection in the adjunction A ⊗ - ⊣ ihom(A, -). It sends f : A ⊗ B ⟶ C to a Rep k G morphism defined by currying the k-linear map underlying f, giving a map A →ₗ[k] B →ₗ[k] C, then flipping the arguments.

Equations
  • One or more equations did not get rendered due to their size.

Porting note: if we generate this with @[simps] the linter complains some types in the LHS simplify.

theorem Rep.homEquiv_symm_apply_hom {k G : Type u} [CommRing k] [Group G] {A B C : Rep k G} (f : B A.ihom.obj C) :

Porting note: if we generate this with @[simps] the linter complains some types in the LHS simplify.

noncomputable instance Rep.instMonoidalClosed {k G : Type u} [CommRing k] [Group G] :
Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem Rep.ihom_obj_ρ_def {k G : Type u} [CommRing k] [Group G] (A B : Rep k G) :
@[simp]
theorem Rep.homEquiv_def {k G : Type u} [CommRing k] [Group G] (A B C : Rep k G) :
@[simp]
theorem Rep.ihom_ev_app_hom {k G : Type u} [CommRing k] [Group G] (A B : Rep k G) :

There is a k-linear isomorphism between the sets of representation morphismsHom(A ⊗ B, C) and Hom(B, Homₖ(A, C)).

Equations
  • One or more equations did not get rendered due to their size.

There is a k-linear isomorphism between the sets of representation morphismsHom(A ⊗ B, C) and Hom(A, Homₖ(B, C)).

Equations
noncomputable def Representation.repOfTprodIso {k G : Type u} [CommRing k] [Monoid G] {V W : Type u} [AddCommGroup V] [AddCommGroup W] [Module k V] [Module k W] (ρ : Representation k G V) (τ : Representation k G W) :

Tautological isomorphism to help Lean in typechecking.

Equations

The categorical equivalence Rep k G ≌ Module.{u} (MonoidAlgebra k G). #

theorem Rep.to_Module_monoidAlgebra_map_aux {k : Type u_1} {G : Type u_2} [CommRing k] [Monoid G] (V : Type u_3) (W : Type u_4) [AddCommGroup V] [AddCommGroup W] [Module k V] [Module k W] (ρ : G →* V →ₗ[k] V) (σ : G →* W →ₗ[k] W) (f : V →ₗ[k] W) (w : ∀ (g : G), f ∘ₗ ρ g = σ g ∘ₗ f) (r : MonoidAlgebra k G) (x : V) :
f ((((MonoidAlgebra.lift k G (V →ₗ[k] V)) ρ) r) x) = (((MonoidAlgebra.lift k G (W →ₗ[k] W)) σ) r) (f x)

Auxiliary lemma for toModuleMonoidAlgebra.

noncomputable def Rep.toModuleMonoidAlgebraMap {k G : Type u} [CommRing k] [Monoid G] {V W : Rep k G} (f : V W) :

Auxiliary definition for toModuleMonoidAlgebra.

Equations

Functorially convert a representation of G into a module over MonoidAlgebra k G.

Equations
  • One or more equations did not get rendered due to their size.

Functorially convert a module over MonoidAlgebra k G into a representation of G.

Equations
  • One or more equations did not get rendered due to their size.

Auxiliary definition for equivalenceModuleMonoidAlgebra.

Equations
theorem Rep.unit_iso_comm {k G : Type u} [CommRing k] [Monoid G] (V : Rep k G) (g : G) (x : V.V) :
noncomputable def Rep.unitIso {k G : Type u} [CommRing k] [Monoid G] (V : Rep k G) :

Auxiliary definition for equivalenceModuleMonoidAlgebra.

Equations
  • One or more equations did not get rendered due to their size.
noncomputable def Rep.equivalenceModuleMonoidAlgebra {k G : Type u} [CommRing k] [Monoid G] :

The categorical equivalence Rep k G ≌ ModuleCat (MonoidAlgebra k G).

Equations
  • One or more equations did not get rendered due to their size.