Documentation

Mathlib.RepresentationTheory.Homological.GroupHomology.Functoriality

Functoriality of group homology #

Given a commutative ring k, a group homomorphism f : G →* H, a k-linear G-representation A, a k-linear H-representation B, and a representation morphism A ⟶ Res(f)(B), we get a chain map inhomogeneousChains A ⟶ inhomogeneousChains B and hence maps on homology Hₙ(G, A) ⟶ Hₙ(H, B).

We also provide extra API for these maps in degrees 0, 1, 2.

Main definitions #

theorem groupHomology.congr {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} {f₁ f₂ : G →* H} (h : f₁ = f₂) {φ : A (Action.res (ModuleCat k) f₁).obj B} {T : Type u_1} (F : (f : G →* H) → (A (Action.res (ModuleCat k) f).obj B) → T) :
F f₁ φ = F f₂ (h φ)
noncomputable def groupHomology.chainsMap {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the chain map sending ∑ aᵢ·gᵢ : Gⁿ →₀ A to ∑ φ(aᵢ)·(f ∘ gᵢ) : Hⁿ →₀ B.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    theorem groupHomology.chainsMap_f {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (i : ) :
    theorem groupHomology.chainsMap_f_hom {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (i : ) :
    theorem groupHomology.chainsMap_f_single {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (n : ) (x : Fin nG) (a : A.V) :
    theorem groupHomology.chainsMap_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) :
    @[simp]
    theorem groupHomology.chainsMap_zero {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) :
    chainsMap f 0 = 0
    theorem groupHomology.chainsMap_f_map_mono {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (hf : Function.Injective f) [CategoryTheory.Mono φ] (i : ) :
    instance groupHomology.chainsMap_id_f_map_mono {k G : Type u} [CommRing k] [Group G] {A B : Rep k G} (φ : A B) [CategoryTheory.Mono φ] (i : ) :
    theorem groupHomology.chainsMap_f_map_epi {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (hf : Function.Surjective f) [CategoryTheory.Epi φ] (i : ) :
    instance groupHomology.chainsMap_id_f_map_epi {k G : Type u} [CommRing k] [Group G] {A B : Rep k G} (φ : A B) [CategoryTheory.Epi φ] (i : ) :
    @[reducible, inline]
    noncomputable abbrev groupHomology.cyclesMap {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (n : ) :
    cycles A n cycles B n

    Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map Zₙ(G, A) ⟶ Zₙ(H, B) sending ∑ aᵢ·gᵢ : Gⁿ →₀ A to ∑ φ(aᵢ)·(f ∘ gᵢ) : Hⁿ →₀ B.

    Equations
    Instances For
      theorem groupHomology.cyclesMap_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) (n : ) :
      theorem groupHomology.cyclesMap_comp_assoc {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) (n : ) {Z : ModuleCat k} (h : cycles C n Z) :
      @[reducible, inline]
      noncomputable abbrev groupHomology.map {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (n : ) :

      Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map Hₙ(G, A) ⟶ Hₙ(H, B) sending ∑ aᵢ·gᵢ : Gⁿ →₀ A to ∑ φ(aᵢ)·(f ∘ gᵢ) : Hⁿ →₀ B.

      Equations
      Instances For
        theorem groupHomology.π_map {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (n : ) :
        theorem groupHomology.map_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) (n : ) :
        theorem groupHomology.map_comp_assoc {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) (n : ) {Z : ModuleCat k} (h : groupHomology C n Z) :
        theorem groupHomology.map_id_comp {k G : Type u} [CommRing k] [Group G] {A B C : Rep k G} (φ : A B) (ψ : B C) (n : ) :
        @[reducible, inline]
        noncomputable abbrev groupHomology.chainsMap₁ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
        ModuleCat.of k (G →₀ A.V) ModuleCat.of k (H →₀ B.V)

        Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map sending ∑ aᵢ·gᵢ : G →₀ A to ∑ φ(aᵢ)·f(gᵢ) : H →₀ B.

        Equations
        Instances For
          @[reducible, inline]
          noncomputable abbrev groupHomology.chainsMap₂ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
          ModuleCat.of k (G × G →₀ A.V) ModuleCat.of k (H × H →₀ B.V)

          Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map sending ∑ aᵢ·(gᵢ₁, gᵢ₂) : G × G →₀ A to ∑ φ(aᵢ)·(f(gᵢ₁), f(gᵢ₂)) : H × H →₀ B.

          Equations
          Instances For
            @[reducible, inline]
            noncomputable abbrev groupHomology.chainsMap₃ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
            ModuleCat.of k (G × G × G →₀ A.V) ModuleCat.of k (H × H × H →₀ B.V)

            Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map sending ∑ aᵢ·(gᵢ₁, gᵢ₂, gᵢ₃) : G × G × G →₀ A to ∑ φ(aᵢ)·(f(gᵢ₁), f(gᵢ₂), f(gᵢ₃)) : H × H × H →₀ B.

            Equations
            Instances For
              @[simp]
              theorem groupHomology.H0π_comp_map {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
              noncomputable def groupHomology.mapShortComplexH1 {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

              Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map from the short complex (G × G →₀ A) --d₂₁--> (G →₀ A) --d₁₀--> A to (H × H →₀ B) --d₂₁--> (H →₀ B) --d₁₀--> B.

              Equations
              Instances For
                @[simp]
                theorem groupHomology.mapShortComplexH1_τ₂ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                @[simp]
                theorem groupHomology.mapShortComplexH1_τ₃ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                @[simp]
                theorem groupHomology.mapShortComplexH1_τ₁ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                @[simp]
                theorem groupHomology.mapShortComplexH1_zero {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) :
                theorem groupHomology.mapShortComplexH1_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) :
                @[reducible, inline]
                noncomputable abbrev groupHomology.mapCycles₁ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

                Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map Z₁(G, A) ⟶ Z₁(H, B).

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  theorem groupHomology.mapCycles₁_hom {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                  theorem groupHomology.mapCycles₁_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) :
                  @[simp]
                  theorem groupHomology.coe_mapCycles₁ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (x : (ModuleCat.of k (cycles₁ A))) :
                  @[simp]
                  theorem groupHomology.H1π_comp_map {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                  @[simp]
                  theorem groupHomology.map₁_one {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (φ : A (Action.res (ModuleCat k) 1).obj B) :
                  map 1 φ 1 = 0

                  Exactness of the corestriction-coinflation sequence in degree 1 #

                  Given a group homomorphism f : G →* H, the nth corestriction map is the map Hₙ(G, Res(f)(A)) ⟶ Hₙ(H, A) induced by f and the identity map on Res(f)(A). Similarly, given a normal subgroup S ≤ G, we define the nth coinflation map Hₙ(G, A) ⟶ Hₙ(G ⧸ S, A_S) as the map induced by the quotient maps G →* G ⧸ S and A →ₗ A_S.

                  In particular, for S ≤ G normal and A : Rep k G, the corestriction map Hₙ(S, Res(ι)(A)) ⟶ Hₙ(G, A) and the coinflation map Hₙ(G, A) ⟶ Hₙ(G ⧸ S, A_S) form a short complex, where ι : S →* G is the natural inclusion. In this section we define this short complex for degree 1, groupHomology.H1CoresCoinf A S, and prove it is exact.

                  We do this first when A is S-trivial, and then extend to the general case.

                  Given a G-representation A on which a normal subgroup S ≤ G acts trivially, this is the short complex H₁(S, A) ⟶ H₁(G, A) ⟶ H₁(G ⧸ S, A). (This is a simplified expression for the degree 1 corestriction-coinflation sequence when A is S-trivial.)

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

                    Given a G-representation A on which a normal subgroup S ≤ G acts trivially, the induced map H₁(G, A) ⟶ H₁(G ⧸ S, A) is an epimorphism.

                    Given a G-representation A on which a normal subgroup S ≤ G acts trivially, the short complex H₁(S, A) ⟶ H₁(G, A) ⟶ H₁(G ⧸ S, A) is exact.

                    noncomputable def groupHomology.H1CoresCoinf {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :

                    The short complex H₁(S, A) ⟶ H₁(G, A) ⟶ H₁(G ⧸ S, A_S). The first map is the "corestriction" map induced by the inclusion ι : S →* G and the identity on Res(ι)(A), and the second map is the "coinflation" map induced by the quotient maps G →* G ⧸ S and A →ₗ A_S.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem groupHomology.H1CoresCoinf_X₂ {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                      @[simp]
                      theorem groupHomology.H1CoresCoinf_X₁ {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                      @[simp]
                      theorem groupHomology.H1CoresCoinf_g {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :
                      @[simp]
                      theorem groupHomology.H1CoresCoinf_X₃ {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :

                      Given a G-representation A and a normal subgroup S ≤ G, let I(S)A denote the submodule of A spanned by elements of the form ρ(s)(a) - a for s : S, a : A. Then the image of C₁(G, I(S)A) in C₁(G, A)⧸B₁(G, A) is contained in the image of C₁(S, A).

                      Given a G-representation A and a normal subgroup S ≤ G, the map H₁(G, A) ⟶ H₁(G ⧸ S, A_S) induced by the quotient maps G →* G ⧸ S and A →ₗ A_S is an epimorphism.

                      theorem groupHomology.H1CoresCoinf_exact {k G : Type u} [CommRing k] [Group G] (A : Rep k G) (S : Subgroup G) [S.Normal] :

                      Given a G-representation A and a normal subgroup S ≤ G, the degree 1 corestriction-coinflation sequence H₁(S, A) ⟶ H₁(G, A) ⟶ H₁(G ⧸ S, A_S) is exact. simps squeezed for performance.

                      noncomputable def groupHomology.mapShortComplexH2 {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

                      Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map from the short complex (G × G × G →₀ A) --d₃₂--> (G × G →₀ A) --d₂₁--> (G →₀ A) to (H × H × H →₀ B) --d₃₂--> (H × H →₀ B) --d₂₁--> (H →₀ B).

                      Equations
                      Instances For
                        @[simp]
                        theorem groupHomology.mapShortComplexH2_τ₃ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                        @[simp]
                        theorem groupHomology.mapShortComplexH2_τ₁ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                        @[simp]
                        theorem groupHomology.mapShortComplexH2_τ₂ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                        @[simp]
                        theorem groupHomology.mapShortComplexH2_zero {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) :
                        theorem groupHomology.mapShortComplexH2_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) :
                        @[reducible, inline]
                        noncomputable abbrev groupHomology.mapCycles₂ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

                        Given a group homomorphism f : G →* H and a representation morphism φ : A ⟶ Res(f)(B), this is the induced map Z₂(G, A) ⟶ Z₂(H, B).

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          theorem groupHomology.mapCycles₂_hom {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :
                          theorem groupHomology.mapCycles₂_comp {k : Type u} [CommRing k] {G H K : Type u} [Group G] [Group H] [Group K] {A : Rep k G} {B : Rep k H} {C : Rep k K} (f : G →* H) (g : H →* K) (φ : A (Action.res (ModuleCat k) f).obj B) (ψ : B (Action.res (ModuleCat k) g).obj C) :
                          @[simp]
                          theorem groupHomology.coe_mapCycles₂ {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) (x : (ModuleCat.of k (cycles₂ A))) :
                          @[simp]
                          theorem groupHomology.H2π_comp_map {k G H : Type u} [CommRing k] [Group G] [Group H] {A : Rep k G} {B : Rep k H} (f : G →* H) (φ : A (Action.res (ModuleCat k) f).obj B) :

                          The functor sending a representation to its complex of inhomogeneous chains.

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

                            The functor sending a G-representation A to Hₙ(G, A).

                            Equations
                            Instances For
                              @[simp]
                              theorem groupHomology.functor_map (k G : Type u) [CommRing k] [Group G] (n : ) {A B : Rep k G} (φ : A B) :
                              (functor k G n).map φ = map (MonoidHom.id G) φ n
                              @[simp]
                              theorem groupHomology.functor_obj (k G : Type u) [CommRing k] [Group G] (n : ) (A : Rep k G) :
                              (functor k G n).obj A = groupHomology A n
                              noncomputable def groupHomology.coresNatTrans (k : Type u) {G H : Type u} [CommRing k] [Group G] [Group H] (f : G →* H) (n : ) :
                              (Action.res (ModuleCat k) f).comp (functor k G n) functor k H n

                              Given a group homomorphism f : G →* H this sends A : Rep k H to the nth "corestriction" map Hₙ(G, Res(f)(A)) ⟶ Hₙ(H, A) induced by f and the identity map on Res(f)(A).

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                @[simp]
                                theorem groupHomology.coresNatTrans_app (k : Type u) {G H : Type u} [CommRing k] [Group G] [Group H] (f : G →* H) (n : ) (X : Action (ModuleCat k) H) :
                                noncomputable def groupHomology.coinfNatTrans (k : Type u) {G : Type u} [CommRing k] [Group G] (S : Subgroup G) [S.Normal] (n : ) :

                                Given a normal subgroup S ≤ G, this sends A : Rep k G to the nth "coinflation" map Hₙ(G, A) ⟶ Hₙ(G ⧸ S, A_S) induced by the quotient maps G →* G ⧸ S and A →ₗ A_S.

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