Documentation

Mathlib.RepresentationTheory.Coinvariants

Coinvariants of a group representation #

Given a commutative ring k and a monoid G, this file introduces the coinvariants of a k-linear G-representation (V, ρ).

We first define Representation.Coinvariants.ker, the submodule of V generated by elements of the form ρ g x - x for x : V, g : G. Then the coinvariants of (V, ρ) are the quotient of V by this submodule. We show that the functor sending a representation to its coinvariants is left adjoint to the functor equipping a module with the trivial representation.

Main definitions #

def Representation.Coinvariants.ker {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :

The submodule of a representation generated by elements of the form ρ g x - x.

Equations
Instances For
    def Representation.Coinvariants {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :
    Type u_3

    The coinvariants of a representation, V ⧸ ⟨{ρ g x - x | g ∈ G, x ∈ V}⟩.

    Equations
    Instances For
      theorem Representation.Coinvariants.sub_mem_ker {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} (g : G) (x : V) :
      (ρ g) x - x ker ρ
      theorem Representation.Coinvariants.mem_ker_of_eq {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} (g : G) (x a : V) (h : (ρ g) x - x = a) :
      a ker ρ
      def Representation.Coinvariants.mk {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :

      The quotient map from a representation to its coinvariants as a linear map.

      Equations
      Instances For
        theorem Representation.Coinvariants.mk_eq_iff {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) {x y : V} :
        (mk ρ) x = (mk ρ) y x - y ker ρ
        theorem Representation.Coinvariants.mk_eq_zero {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) {x : V} :
        (mk ρ) x = 0 x ker ρ
        theorem Representation.Coinvariants.mk_surjective {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :
        @[simp]
        theorem Representation.Coinvariants.mk_self_apply {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (g : G) (x : V) :
        (mk ρ) ((ρ g) x) = (mk ρ) x
        theorem Representation.Coinvariants.induction_on {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} {motive : ρ.CoinvariantsProp} (x : ρ.Coinvariants) (h : ∀ (v : V), motive ((mk ρ) v)) :
        motive x
        def Representation.Coinvariants.lift {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (f : V →ₗ[k] W) (h : ∀ (x : G), f ∘ₗ ρ x = f) :

        A G-invariant linear map induces a linear map out of the coinvariants of a G-representation.

        Equations
        Instances For
          @[simp]
          theorem Representation.Coinvariants.lift_comp_mk {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (f : V →ₗ[k] W) (h : ∀ (x : G), f ∘ₗ ρ x = f) :
          lift ρ f h ∘ₗ mk ρ = f
          @[simp]
          theorem Representation.Coinvariants.lift_mk {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (f : V →ₗ[k] W) (h : ∀ (x : G), f ∘ₗ ρ x = f) (x : V) :
          (lift ρ f h) ((mk ρ) x) = f x
          theorem Representation.Coinvariants.hom_ext {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] {ρ : Representation k G V} {f g : ρ.Coinvariants →ₗ[k] W} (H : f ∘ₗ mk ρ = g ∘ₗ mk ρ) :
          f = g
          theorem Representation.Coinvariants.hom_ext_iff {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] {ρ : Representation k G V} {f g : ρ.Coinvariants →ₗ[k] W} :
          f = g f ∘ₗ mk ρ = g ∘ₗ mk ρ
          noncomputable def Representation.Coinvariants.map {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (τ : Representation k G W) (f : V →ₗ[k] W) (hf : ∀ (g : G), f ∘ₗ ρ g = τ g ∘ₗ f) :

          Given G-representations on k-modules V, W, a linear map V →ₗ[k] W commuting with the representations induces a k-linear map between the coinvariants.

          Equations
          Instances For
            @[simp]
            theorem Representation.Coinvariants.map_comp_mk {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] {ρ : Representation k G V} {τ : Representation k G W} (f : V →ₗ[k] W) (hf : ∀ (g : G), f ∘ₗ ρ g = τ g ∘ₗ f) :
            map ρ τ f hf ∘ₗ mk ρ = mk τ ∘ₗ f
            @[simp]
            theorem Representation.Coinvariants.map_mk {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] {ρ : Representation k G V} {τ : Representation k G W} (f : V →ₗ[k] W) (hf : ∀ (g : G), f ∘ₗ ρ g = τ g ∘ₗ f) (x : V) :
            (map ρ τ f hf) ((mk ρ) x) = (mk τ) (f x)
            @[simp]
            theorem Representation.Coinvariants.map_id {k : Type u_1} {G : Type u_2} {V : Type u_3} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :
            @[simp]
            theorem Representation.Coinvariants.map_comp {k : Type u_1} {G : Type u_2} {V : Type u_3} {W : Type u_4} {X : Type u_5} [CommRing k] [Monoid G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] [AddCommGroup X] [Module k X] {ρ : Representation k G V} {τ : Representation k G W} (υ : Representation k G X) (φ : V →ₗ[k] W) (ψ : W →ₗ[k] X) (H : ∀ (g : G), φ ∘ₗ ρ g = τ g ∘ₗ φ) (h : ∀ (g : G), ψ ∘ₗ τ g = υ g ∘ₗ ψ) :
            map τ υ ψ h ∘ₗ map ρ τ φ H = map ρ υ (ψ ∘ₗ φ)
            theorem Representation.Coinvariants.le_comap_ker {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (S : Subgroup G) [S.Normal] (g : G) :
            @[reducible, inline]
            noncomputable abbrev Representation.toCoinvariantsKer {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (S : Subgroup G) [S.Normal] :

            Given a normal subgroup S ≤ G, a G-representation ρ restricts to a G-representation on the kernel of the quotient map to the coinvariants of ρ|_S.

            Equations
            Instances For
              @[reducible, inline]
              noncomputable abbrev Representation.toCoinvariants {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (S : Subgroup G) [S.Normal] :

              Given a normal subgroup S ≤ G, a G-representation ρ induces a G-representation on the coinvariants of ρ|_S.

              Equations
              Instances For
                @[reducible, inline]
                noncomputable abbrev Representation.quotientToCoinvariants {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (S : Subgroup G) [S.Normal] :

                Given a normal subgroup S ≤ G, a G-representation ρ induces a G ⧸ S-representation on the coinvariants of ρ|_S.

                Equations
                Instances For
                  noncomputable def Representation.coinvariantsToFinsupp {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (α : Type u_9) :

                  Given a G-representation (V, ρ) and a type α, this is the map (α →₀ V)_G →ₗ (α →₀ V_G) sending ⟦single a v⟧ ↦ single a ⟦v⟧.

                  Equations
                  Instances For
                    @[simp]
                    theorem Representation.coinvariantsToFinsupp_mk_single {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} {α : Type u_9} (x : α) (a : V) :
                    noncomputable def Representation.finsuppToCoinvariants {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (α : Type u_9) :

                    Given a G-representation (V, ρ) and a type α, this is the map (α →₀ V_G) →ₗ (α →₀ V)_G sending single a ⟦v⟧ ↦ ⟦single a v⟧.

                    Equations
                    Instances For
                      @[simp]
                      theorem Representation.finsuppToCoinvariants_single_mk {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} {α : Type u_9} (a : α) (x : V) :
                      noncomputable def Representation.coinvariantsFinsuppLEquiv {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (α : Type u_9) :

                      Given a G-representation (V, ρ) and a type α, this is the linear equivalence (α →₀ V)_G ≃ₗ (α →₀ V_G) sending ⟦single a v⟧ ↦ single a ⟦v⟧.

                      Equations
                      Instances For
                        @[simp]
                        theorem Representation.coinvariantsFinsuppLEquiv_symm_apply {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (α : Type u_9) (a : α →₀ ρ.Coinvariants) :
                        @[simp]
                        theorem Representation.coinvariantsFinsuppLEquiv_apply {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] {ρ : Representation k G V} {α : Type u_9} (x : (ρ.finsupp α).Coinvariants) :
                        @[simp]
                        theorem Representation.Coinvariants.mk_inv_tmul {k : Type u_6} {G : Type u_7} {V : Type u_8} {W : Type u_9} [CommRing k] [Group G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (τ : Representation k G W) (x : V) (y : W) (g : G) :
                        (mk (ρ.tprod τ)) ((ρ g⁻¹) x ⊗ₜ[k] y) = (mk (ρ.tprod τ)) (x ⊗ₜ[k] (τ g) y)
                        @[simp]
                        theorem Representation.Coinvariants.mk_tmul_inv {k : Type u_6} {G : Type u_7} {V : Type u_8} {W : Type u_9} [CommRing k] [Group G] [AddCommGroup V] [Module k V] [AddCommGroup W] [Module k W] (ρ : Representation k G V) (τ : Representation k G W) (x : V) (y : W) (g : G) :
                        (mk (ρ.tprod τ)) (x ⊗ₜ[k] (τ g⁻¹) y) = (mk (ρ.tprod τ)) ((ρ g) x ⊗ₜ[k] y)
                        noncomputable def Representation.ofCoinvariantsTprodLeftRegular {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :

                        Given a k-linear G-representation V, ρ, this is the map (V ⊗ k[G])_G →ₗ[k] V sending ⟦v ⊗ single g r⟧ ↦ r • ρ(g⁻¹)(v).

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          @[simp]
                          theorem Representation.ofCoinvariantsTprodLeftRegular_mk_tmul_single {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) (x : V) (g : G) (r : k) :
                          noncomputable def Representation.coinvariantsTprodLeftRegularLEquiv {k : Type u_6} {G : Type u_7} {V : Type u_8} [CommRing k] [Group G] [AddCommGroup V] [Module k V] (ρ : Representation k G V) :

                          Given a k-linear G-representation (V, ρ), this is the linear equivalence (V ⊗ k[G])_G ≃ₗ[k] V sending ⟦v ⊗ single g r⟧ ↦ r • ρ(g⁻¹)(v).

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            @[reducible, inline]
                            abbrev Rep.toCoinvariantsKer {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                            Rep k G

                            Given a normal subgroup S ≤ G, a G-representation A restricts to a G-representation on the kernel of the quotient map to the S-coinvariants A_S.

                            Equations
                            Instances For
                              @[reducible, inline]
                              abbrev Rep.toCoinvariants {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                              Rep k G

                              Given a normal subgroup S ≤ G, a G-representation A induces a G-representation on the S-coinvariants A_S.

                              Equations
                              Instances For
                                @[reducible, inline]
                                abbrev Rep.quotientToCoinvariants {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                                Rep k (G S)

                                Given a normal subgroup S ≤ G, a G-representation ρ induces a G ⧸ S-representation on the coinvariants of ρ|_S.

                                Equations
                                Instances For

                                  Given a normal subgroup S ≤ G, a G-representation A induces a short exact sequence of G-representations 0 ⟶ Ker(mk) ⟶ A ⟶ A_S ⟶ 0 where mk is the quotient map to the S-coinvariants A_S.

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    @[simp]
                                    theorem Rep.coinvariantsShortComplex_X₂ {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                                    noncomputable def Rep.coinvariantsFunctor (k G : Type u) [CommRing k] [Monoid G] :

                                    The functor sending a representation to its coinvariants.

                                    Equations
                                    • One or more equations did not get rendered due to their size.
                                    Instances For
                                      @[simp]
                                      theorem Rep.coinvariantsFunctor_map_hom (k G : Type u) [CommRing k] [Monoid G] {X✝ Y✝ : Rep k G} (f : X✝ Y✝) :
                                      @[simp]
                                      noncomputable def Rep.coinvariantsMk (k G : Type u) [CommRing k] [Monoid G] :

                                      The quotient map from a representation to its coinvariants induces a natural transformation from the forgetful functor Rep k G ⥤ ModuleCat k to the coinvariants functor.

                                      Equations
                                      Instances For
                                        @[reducible, inline]
                                        noncomputable abbrev Rep.desc {k G : Type u} [CommRing k] [Monoid G] {A B : Rep k G} [B.ρ.IsTrivial] (f : A B) :

                                        The linear map underlying a G-representation morphism A ⟶ B, where B has the trivial representation, factors through A_G.

                                        Equations
                                        Instances For

                                          The adjunction between the functor sending a representation to its coinvariants and the functor equipping a module with the trivial representation.

                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            @[simp]
                                            @[reducible, inline]
                                            noncomputable abbrev Rep.coinvariantsTensor (k G : Type u) [CommRing k] [Monoid G] :

                                            The functor sending A, B to (A ⊗[k] B)_G. This is naturally isomorphic to the functor sending A, B to A ⊗[k[G]] B, where we give A the k[G]ᵐᵒᵖ-module structure defined by g • a := A.ρ g⁻¹ a.

                                            Equations
                                            Instances For
                                              @[reducible, inline]
                                              noncomputable abbrev Rep.coinvariantsTensorMk {k G : Type u} [CommRing k] [Monoid G] (A B : Rep k G) :
                                              A.V →ₗ[k] B.V →ₗ[k] (((coinvariantsTensor k G).obj A).obj B)

                                              The bilinear map sending a : A, b : B to ⟦a ⊗ₜ b⟧ in (A ⊗[k] B)_G.

                                              Equations
                                              Instances For

                                                Given a k-linear G-representation (A, ρ) and a type α, this is the map (A ⊗ (α →₀ k[G]))_G →ₗ[k] (α →₀ A) sending ⟦a ⊗ single x (single g r)⟧ ↦ single x (r • ρ(g⁻¹)(a)).

                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  @[simp]
                                                  theorem Rep.coinvariantsTensorFreeToFinsupp_mk_tmul_single {k G : Type u} [CommRing k] [Group G] {A : Rep k G} {α : Type u} [DecidableEq α] (x : A.V) (i : α) (g : G) (r : k) :

                                                  Given a k-linear G-representation (A, ρ) and a type α, this is the map (α →₀ A) →ₗ[k] (A ⊗ (α →₀ k[G]))_G sending single x a ↦ ⟦a ⊗ₜ single x 1⟧.

                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    @[reducible, inline]
                                                    noncomputable abbrev Rep.coinvariantsTensorFreeLEquiv {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (α : Type u) [DecidableEq α] :

                                                    Given a k-linear G-representation (A, ρ) and a type α, this is the linear equivalence (A ⊗ (α →₀ k[G]))_G ≃ₗ[k] (α →₀ A) sending ⟦a ⊗ single x (single g r)⟧ ↦ single x (r • ρ(g⁻¹)(a)).

                                                    Equations
                                                    Instances For