mathlib3 documentation

representation_theory.group_cohomology.basic

The group cohomology of a k-linear G-representation #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

Let k be a commutative ring and G a group. This file defines the group cohomology of A : Rep k G to be the cohomology of the complex $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ with differential $d^n$ sending $f: G^n \to A$ to the function mapping $(g_0, \dots, g_n)$ to $$\rho(g_0)(f(g_1, \dots, g_n)) + \sum_{i = 0}^{n - 1} (-1)^{i + 1}\cdot f(g_0, \dots, g_ig_{i + 1}, \dots, g_n)$$ $$+ (-1)^{n + 1}\cdot f(g_0, \dots, g_{n - 1})$$ (where ρ is the representation attached to A).

We have a k-linear isomorphism $\mathrm{Fun}(G^n, A) \cong \mathrm{Hom}(k[G^{n + 1}], A)$, where the righthand side is morphisms in Rep k G, and the representation on $k[G^{n + 1}]$ is induced by the diagonal action of G. If we conjugate the $n$th differential in $\mathrm{Hom}(P, A)$ by this isomorphism, where P is the standard resolution of k as a trivial k-linear G-representation, then the resulting map agrees with the differential $d^n$ defined above, a fact we prove.

This gives us for free a proof that our $d^n$ squares to zero. It also gives us an isomorphism $\mathrm{H}^n(G, A) \cong \mathrm{Ext}^n(k, A),$ where $\mathrm{Ext}$ is taken in the category Rep k G.

Main definitions #

Implementation notes #

Group cohomology is typically stated for G-modules, or equivalently modules over the group ring ℤ[G]. However, can be generalized to any commutative ring k, which is what we use. Moreover, we express k[G]-module structures on a module k-module A using the Rep definition. We avoid using instances module (monoid_algebra k G) A so that we do not run into possible scalar action diamonds.

TODO #

Longer term:

@[reducible]
noncomputable def group_cohomology.linear_yoneda_obj_resolution {k G : Type u} [comm_ring k] [monoid G] (A : Rep k G) :

The complex Hom(P, A), where P is the standard resolution of k as a trivial k-linear G-representation.

@[simp]
theorem inhomogeneous_cochains.d_apply {k G : Type u} [comm_ring k] [monoid G] (n : ) (A : Rep k G) (f : (fin n G) A) (g : fin (n + 1) G) :
(inhomogeneous_cochains.d n A) f g = ((A.ρ) (g 0)) (f (λ (i : fin n), g i.succ)) + finset.univ.sum (λ (j : fin (n + 1)), (-1) ^ (j + 1) f (j.contract_nth has_mul.mul g))
def inhomogeneous_cochains.d {k G : Type u} [comm_ring k] [monoid G] (n : ) (A : Rep k G) :
((fin n G) A) →ₗ[k] (fin (n + 1) G) A

The differential in the complex of inhomogeneous cochains used to calculate group cohomology.

Equations

The theorem that our isomorphism Fun(Gⁿ, A) ≅ Hom(k[Gⁿ⁺¹], A) (where the righthand side is morphisms in Rep k G) commutes with the differentials in the complex of inhomogeneous cochains and the homogeneous linear_yoneda_obj_resolution.

@[reducible]

Given a k-linear G-representation A, this is the complex of inhomogeneous cochains $$0 \to \mathrm{Fun}(G^0, A) \to \mathrm{Fun}(G^1, A) \to \mathrm{Fun}(G^2, A) \to \dots$$ which calculates the group cohomology of A.

Given a k-linear G-representation A, the complex of inhomogeneous cochains is isomorphic to Hom(P, A), where P is the standard resolution of k as a trivial G-representation.

Equations
noncomputable def group_cohomology {k G : Type u} [comm_ring k] [group G] (A : Rep k G) (n : ) :

The group cohomology of a k-linear G-representation A, as the cohomology of its complex of inhomogeneous cochains.

Equations
noncomputable def group_cohomology_iso_Ext {k G : Type u} [comm_ring k] [group G] (A : Rep k G) (n : ) :
group_cohomology A n ((Ext k (Rep k G) n).obj (opposite.op (Rep.trivial k G k))).obj A

The nth group cohomology of a k-linear G-representation A is isomorphic to Extⁿ(k, A) (taken in Rep k G), where k is a trivial k-linear G-representation.

Equations