Documentation

Mathlib.Algebra.Lie.Weights.Linear

Lie modules with linear weights #

Given a Lie module M over a nilpotent Lie algebra L with coefficients in R, one frequently studies M via its weights. These are functions χ : L → R whose corresponding weight space LieModule.weightSpace M χ, is non-trivial. If L is Abelian or if R has characteristic zero (and M is finite-dimensional) then such χ are necessarily R-linear. However in general non-linear weights do exist. For example if we take:

then there is a single weight and it is non-linear. (See remark following Proposition 9 of chapter VII, §1.3 in N. Bourbaki, Chapters 7--9.)

We thus introduce a typeclass LieModule.LinearWeights to encode the fact that a Lie module does have linear weights and provide typeclass instances in the two important cases that L is Abelian or R has characteristic zero.

Main definitions #

class LieModule.LinearWeights (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] :

A typeclass encoding the fact that a given Lie module has linear weights, vanishing on the derived ideal.

Instances
    @[simp]
    theorem LieModule.weight.toLinear_apply (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] [LieModule.LinearWeights R L M] [NoZeroSMulDivisors R M] [IsNoetherian R M] (χ : { x : LR // x LieModule.weight R L M }) :
    ∀ (a : L), (LieModule.weight.toLinear R L M χ) a = χ a
    def LieModule.weight.toLinear (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] [LieModule.LinearWeights R L M] [NoZeroSMulDivisors R M] [IsNoetherian R M] (χ : { x : LR // x LieModule.weight R L M }) :

    A weight of a Lie module, bundled as a linear map.

    Equations
    Instances For
      @[simp]
      theorem LieModule.weight.toLinear_apply_lie (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] [LieModule.LinearWeights R L M] [NoZeroSMulDivisors R M] [IsNoetherian R M] (χ : { x : LR // x LieModule.weight R L M }) (x : L) (y : L) :
      χ x, y = 0

      For an Abelian Lie algebra, the weights of any Lie module are linear.

      Equations
      • =

      In characteristic zero, the weights of any finite-dimensional Lie module are linear and vanish on the derived ideal.

      Equations
      • =
      def LieModule.shiftedWeightSpace (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] (χ : LR) :

      A type synonym for the χ-weight space but with the action of x : L on m : weightSpace M χ, shifted to act as ⁅x, m⁆ - χ x • m.

      Equations
      Instances For
        @[simp]
        theorem LieModule.shiftedWeightSpace.coe_lie_shiftedWeightSpace_apply (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] [LieModule.LinearWeights R L M] (χ : LR) (x : L) (m : (LieModule.shiftedWeightSpace R L M χ)) :
        x, m = x, m - χ x m
        @[simp]
        theorem LieModule.shiftedWeightSpace.shift_symm_apply (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] (χ : LR) :
        @[simp]
        theorem LieModule.shiftedWeightSpace.shift_apply (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] (χ : LR) :
        ∀ (a : (LieModule.weightSpace M χ)), (LieModule.shiftedWeightSpace.shift R L M χ) a = a
        def LieModule.shiftedWeightSpace.shift (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] (χ : LR) :

        Forgetting the action of L, the spaces weightSpace M χ and shiftedWeightSpace R L M χ are equivalent.

        Equations
        Instances For
          theorem LieModule.exists_forall_lie_eq_smul_of_weightSpace_ne_bot (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [LieAlgebra R L] [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] [LieAlgebra.IsNilpotent R L] [LieModule.LinearWeights R L M] (χ : LR) [IsNoetherian R M] (hχ : LieModule.weightSpace M χ ) :
          ∃ (m : M), m 0 ∀ (x : L), x, m = χ x m

          Given a Lie module M of a Lie algebra L with coefficients in R, if a function χ : L → R has a simultaneous generalized eigenvector for the action of L then it has a simultaneous true eigenvector, provided M is Noetherian and has linear weights.