Lie algebra cohomology in low degree #
This file defines low degree cochains of Lie algebras with coefficients given by a module. They are useful in the construction of central extensions, so we treat these easier cases separately from the general theory of Lie algebra cohomology.
Main definitions #
LieAlgebra.oneCochain
: an abbreviation for a linear map.LieAlgebra.twoCochain
: a submodule of bilinear maps, giving 2-cochains.LieAlgebra.d₁₂
: The coboundary map taking 1-cochains to 2-cochains.LieAlgebra.d₂₃
: A coboundary map taking 2-cochains to a space containing 3-cochains.LieAlgebra.twoCocycle
: The submodule of 2-cocycles.
TODO #
- coboundaries, cohomology
- comparison to the Chevalley-Eilenberg complex.
- construction and classification of central extensions
References #
@[reducible, inline]
noncomputable abbrev
LieModule.Cohomology.oneCochain
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
:
Type (max u_2 u_3)
Lie algebra 1-cochains over L
with coefficients in the module M
.
Equations
- LieModule.Cohomology.oneCochain R L M = (L →ₗ[R] M)
Instances For
noncomputable def
LieModule.Cohomology.twoCochain
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
:
Lie algebra 2-cochains over L
with coefficients in the module M
.
Equations
Instances For
noncomputable instance
LieModule.Cohomology.instFunLikeSubtypeLinearMapIdMemSubmoduleTwoCochain
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
:
FunLike (↥(twoCochain R L M)) L (L →ₗ[R] M)
Equations
- LieModule.Cohomology.instFunLikeSubtypeLinearMapIdMemSubmoduleTwoCochain = { coe := fun (a : ↥(LieModule.Cohomology.twoCochain R L M)) (x : L) => ↑a x, coe_injective' := ⋯ }
instance
LieModule.Cohomology.instLinearMapClassSubtypeLinearMapIdMemSubmoduleTwoCochain
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
:
LinearMapClass (↥(twoCochain R L M)) R L (L →ₗ[R] M)
@[simp]
theorem
LieModule.Cohomology.twoCochain_alt
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
(a : ↥(twoCochain R L M))
(x : L)
:
theorem
LieModule.Cohomology.twoCochain_skew
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
(a : ↥(twoCochain R L M))
(x y : L)
:
@[simp]
theorem
LieModule.Cohomology.twoCochain_val_apply
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
(a : ↥(twoCochain R L M))
(x : L)
:
@[simp]
theorem
LieModule.Cohomology.add_apply_apply
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
(a b : ↥(twoCochain R L M))
(x y : L)
:
@[simp]
theorem
LieModule.Cohomology.smul_apply_apply
{R : Type u_1}
[CommRing R]
{L : Type u_2}
[LieRing L]
[LieAlgebra R L]
{M : Type u_3}
[AddCommGroup M]
[Module R M]
(r : R)
(a : ↥(twoCochain R L M))
(x y : L)
:
noncomputable def
LieModule.Cohomology.d₁₂
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
:
The coboundary operator taking degree 1 cochains to degree 2 cochains.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
LieModule.Cohomology.d₁₂_apply_coe_apply_apply
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
(f : oneCochain R L M)
(x y : L)
:
@[simp]
theorem
LieModule.Cohomology.d₁₂_apply_apply
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
(f : oneCochain R L M)
(x y : L)
:
theorem
LieModule.Cohomology.d₁₂_apply_apply_ofTrivial
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
[IsTrivial L M]
(f : oneCochain R L M)
(x y : L)
:
noncomputable def
LieModule.Cohomology.d₂₃
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
:
The coboundary operator taking degree 2 cochains to a space containing degree 3 cochains.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
LieModule.Cohomology.d₂₃_apply
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
(a : ↥(twoCochain R L M))
(x y z : L)
:
theorem
LieModule.Cohomology.d₂₃_comp_d₁₂
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
:
noncomputable def
LieModule.Cohomology.twoCocycle
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
:
Submodule R ↥(twoCochain R L M)
A Lie 2-cocycle is a 2-cochain that is annihilated by the coboundary map.
Equations
Instances For
theorem
LieModule.Cohomology.mem_twoCocycle_iff
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
(a : ↥(twoCochain R L M))
:
theorem
LieModule.Cohomology.mem_twoCocycle_iff_of_trivial
(R : Type u_1)
[CommRing R]
(L : Type u_2)
[LieRing L]
[LieAlgebra R L]
(M : Type u_3)
[AddCommGroup M]
[Module R M]
[LieRingModule L M]
[LieModule R L M]
[IsTrivial L M]
(a : ↥(twoCochain R L M))
: