Coinduced representations #
Given a commutative ring k
, a monoid homomorphism φ : G →* H
, and a k
-linear
G
-representation A
, this file introduces the coinduced representation $Coind_G^H(A)$ of A
as
an H
-representation.
By coind φ A
we mean the submodule of functions H → A
such that for all g : G
, h : H
,
f (φ g * h) = ρ g (f h)
. We define a representation of H
on this submodule by sending h : H
and f : coind φ A
to the function H → A
sending h₁
to f (h₁ * h)
.
Alternatively, we could define $Coind_G^H(A)$ as the morphisms $Hom(k[H], A)$ in the category
Rep k G
, which we equip with the H
-representation sending h : H
and f : k[H] ⟶ A
to the
representation morphism sending r · h₁
to r • f (h₁ * h)
. We include this definition as
coind' φ A
and prove the two representations are isomorphic.
We also prove that the restriction functor Rep k H ⥤ Rep k G
along φ
is left adjoint to the
coinduction functor and hence that the coinduction functor preserves limits.
Main definitions #
Representation.coind φ ρ
: the coinduction ofρ
alongφ
defined as thek
-submodule ofG
-equivariant functionsH → A
, withH
-action given by(h • f) h₁ := f (h₁ * h)
forf : H → A
,h, h₁ : H
.Representation.coind' φ A
: the coinduction ofA
alongφ
defined as the set ofG
-representation morphismsk[H] ⟶ A
, withH
-action given by(h • f) (r • h₁) := r • f(h₁ * h)
forf : k[H] ⟶ A
,h, h₁ : H
,r : k
.Rep.resCoindAdjunction k φ
: given a monoid homomorphismφ : G →* H
, this is the adjunction between the restriction functorRep k H ⥤ Rep k G
alongφ
and the coinduction functor alongφ
.
If ρ : Representation k G A
and φ : G →* H
then coindV φ ρ
is the sub-k
-module of
functions H → A
underlying the coinduction of ρ
along φ
, i.e., the functions f : H → A
such that f (φ g * h) = (ρ g) (f h)
for all g : G
and h : H
.
Equations
Instances For
If ρ : Representation k G A
and φ : G →* H
then coind φ ρ
is the representation
coinduced by ρ
along φ
, defined as the following action of H
on the submodule coindV φ ρ
of G
-equivariant functions from H
to A
: we let h : H
send the function f : H → A
to the function sending h₁
to f (h₁ * h)
.
See also Rep.coind
and Representation.coind'
for variants involving the category Rep k G
.
Equations
- Representation.coind φ ρ = { toFun := fun (h : H) => (LinearMap.funLeft k A fun (x : H) => x * h).restrict ⋯, map_one' := ⋯, map_mul' := ⋯ }
Instances For
If φ : G →* H
and A : Rep k G
then coind φ A
is the coinduction of A
along φ
,
defined by letting H
act on the G
-equivariant functions H → A
by (h • f) h₁ := f (h₁ * h)
.
Equations
- Rep.coind φ A = Rep.of (Representation.coind φ A.ρ)
Instances For
Given a monoid morphism φ : G →* H
and a morphism of G
-representations f : A ⟶ B
, there
is a natural H
-representation morphism coind φ A ⟶ coind φ B
, given by postcomposition by
f
.
Equations
- Rep.coindMap φ f = { hom := ModuleCat.ofHom (((ModuleCat.Hom.hom f.hom).compLeft H).restrict ⋯), comm := ⋯ }
Instances For
Given a monoid homomorphism φ : G →* H
, this is the functor sending a G
-representation A
to the coinduced H
-representation coind φ A
, with action on maps given by postcomposition.
Equations
- Rep.coindFunctor k φ = { obj := fun (A : Rep k G) => Rep.coind φ A, map := fun {X Y : Rep k G} (f : X ⟶ Y) => Rep.coindMap φ f, map_id := ⋯, map_comp := ⋯ }
Instances For
If φ : G →* H
and A : Rep k G
then coind' φ A
, the coinduction of A
along φ
,
is defined as an H
-action on Hom_{k[G]}(k[H], A)
. If f : k[H] → A
is G
-equivariant
then (h • f) (r • h₁) := r • f (h₁ * h)
, where r : k
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If φ : G →* H
and A : Rep k G
then coind' φ A
, the coinduction of A
along φ
,
is defined as an H
-action on Hom_{k[G]}(k[H], A)
. If f : k[H] → A
is G
-equivariant
then (h • f) (r • h₁) := r • f (h₁ * h)
, where r : k
.
Equations
- Rep.coind' φ A = Rep.of (Representation.coind' φ A)
Instances For
Given a monoid morphism φ : G →* H
and a morphism of G
-representations f : A ⟶ B
, there
is a natural H
-representation morphism coind' φ A ⟶ coind' φ B
, given by postcomposition
by f
.
Equations
- Rep.coindMap' φ f = { hom := ModuleCat.ofHom (CategoryTheory.Linear.rightComp k ((Action.res (ModuleCat k) φ).obj (Rep.leftRegular k H)) f), comm := ⋯ }
Instances For
Given a monoid homomorphism φ : G →* H
, this is the functor sending a G
-representation A
to the coinduced H
-representation coind' φ A
, with action on maps given by postcomposition.
Equations
- Rep.coindFunctor' k φ = { obj := fun (A : Rep k G) => Rep.coind' φ A, map := fun {X Y : Rep k G} (f : X ⟶ Y) => Rep.coindMap' φ f, map_id := ⋯, map_comp := ⋯ }
Instances For
If φ : G →* H
and A : Rep k G
then the k
-submodule of functions f : H → A
such that for all g : G
, h : H
, f (φ g * h) = A.ρ g (f h)
, is k
-linearly equivalent
to the G
-representation morphisms k[H] ⟶ A
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
coind φ A
and coind' φ A
are isomorphic representations, with the underlying
k
-linear equivalence given by coindVLEquiv
.
Equations
- Rep.coindIso φ A = Action.mkIso (Rep.coindVLEquiv φ A).toModuleIso ⋯
Instances For
Given a monoid homomorphism φ : G →* H
, the coinduction functors Rep k G ⥤ Rep k H
given by
coindFunctor k φ
and coindFunctor' k φ
are naturally isomorphic, with isomorphism on objects
given by coindIso φ
.
Equations
Instances For
Given a monoid homomorphism φ : G →* H
, an H
-representation B
, and a G
-representation
A
, there is a k
-linear equivalence between the G
-representation morphisms B ⟶ A
and the
H
-representation morphisms B ⟶ coind φ A
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Given a monoid homomorphism φ : G →* H
, the coinduction functor Rep k G ⥤ Rep k H
is right
adjoint to the restriction functor along φ
.
Equations
- One or more equations did not get rendered due to their size.