Documentation

Mathlib.Algebra.Lie.Cochain

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 #

TODO #

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
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
      @[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) :
      (a x) x = 0
      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) :
      -(a x) y = (a y) x
      @[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) :
      a x = a x
      @[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) :
      ((a + b) x) y = (a x) y + (b x) y
      @[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) :
      ((r a) x) y = r (a x) y
      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] :
      oneCochain R L M →ₗ[R] (twoCochain 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) :
        (((d₁₂ R L M) f) x) y = x, f y - y, f x - f x, y
        @[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) :
        (((d₁₂ R L M) f) x) y = x, f y - y, f x - f x, y
        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) :
        (((d₁₂ R L M) f) x) y = -f x, y
        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) :
          ((((d₂₃ R L M) a) x) y) z = x, (a y) z - y, (a x) z + z, (a x) y - (a x, y) z + (a x, z) y - (a y, z) x
          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] :
          d₂₃ R L M ∘ₗ d₁₂ R L M = 0
          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)) :
            a twoCocycle R L M (d₂₃ R L M) a = 0
            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)) :
            a twoCocycle R L M ∀ (x y z : L), (a x) y, z = (a x, y) z + (a y) x, z