Documentation

Mathlib.Algebra.Lie.Derivation.AdjointAction

Adjoint action of a Lie algebra on itself #

This file defines the adjoint action of a Lie algebra on itself, and establishes basic properties.

Main definitions #

Main statements #

@[simp]
theorem LieDerivation.ad_apply_apply (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] (a : L) (a : L) :
((LieDerivation.ad R L) a✝) a = a✝, a
def LieDerivation.ad (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] :

The adjoint action of a Lie algebra L on itself, seen as a morphism of Lie algebras from L to its derivations. Note the minus sign: this is chosen to so that ad ⁅x, y⁆ = ⁅ad x, ad y⁆.

Equations
Instances For
    @[simp]
    theorem LieDerivation.coe_ad_apply_eq_ad_apply {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] (x : L) :
    ((LieDerivation.ad R L) x) = (LieAlgebra.ad R L) x

    The definitions LieDerivation.ad and LieAlgebra.ad agree.

    theorem LieDerivation.ad_apply_lieDerivation {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] (x : L) (D : LieDerivation R L L) :
    (LieDerivation.ad R L) (D x) = -x, D
    theorem LieDerivation.lie_ad {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] (x : L) (D : LieDerivation R L L) :
    theorem LieDerivation.ad_ker_eq_center (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] :

    The kernel of the adjoint action on a Lie algebra is equal to its center.

    If the center of a Lie algebra is trivial, then the adjoint action is injective.

    theorem LieDerivation.lie_der_ad_eq_ad_der {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] (D : LieDerivation R L L) (x : L) :

    The commutator of a derivation D and a derivation of the form ad x is ad (D x).

    theorem LieDerivation.ad_isIdealMorphism (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] :
    (LieDerivation.ad R L).IsIdealMorphism

    The range of the adjoint action homomorphism from a Lie algebra L to the Lie algebra of its derivations is an ideal of the latter.

    theorem LieDerivation.mem_ad_idealRange_iff {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] {D : LieDerivation R L L} :
    D (LieDerivation.ad R L).idealRange ∃ (x : L), (LieDerivation.ad R L) x = D

    A derivation D belongs to the ideal range of the adjoint action iff it is of the form ad x for some x in the Lie algebra L.