Documentation

Mathlib.Algebra.Lie.Subalgebra

Lie subalgebras #

This file defines Lie subalgebras of a Lie algebra and provides basic related definitions and results.

Main definitions #

Tags #

lie algebra, lie subalgebra

structure LieSubalgebra (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] extends Submodule R L :

A Lie subalgebra of a Lie algebra is submodule that is closed under the Lie bracket. This is a sufficient condition for the subset itself to form a Lie algebra.

instance instZeroLieSubalgebra (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] :

The zero algebra is a subalgebra of any Lie algebra.

Equations
Equations
instance LieSubalgebra.instSetLike (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
instance LieSubalgebra.lieRing (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
LieRing L'

A Lie subalgebra forms a new Lie ring.

Equations
instance LieSubalgebra.instModuleSubtypeMemOfIsScalarTower (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] {R₁ : Type u_1} [Semiring R₁] [SMul R₁ R] [Module R₁ L] [IsScalarTower R₁ R L] (L' : LieSubalgebra R L) :
Module R₁ L'

A Lie subalgebra inherits module structures from L.

Equations
instance LieSubalgebra.instIsCentralScalarSubtypeMem (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] {R₁ : Type u_1} [Semiring R₁] [SMul R₁ R] [SMul R₁ᵐᵒᵖ R] [Module R₁ L] [Module R₁ᵐᵒᵖ L] [IsScalarTower R₁ R L] [IsScalarTower R₁ᵐᵒᵖ R L] [IsCentralScalar R₁ L] (L' : LieSubalgebra R L) :
IsCentralScalar R₁ L'
instance LieSubalgebra.instIsScalarTowerSubtypeMem (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] {R₁ : Type u_1} [Semiring R₁] [SMul R₁ R] [Module R₁ L] [IsScalarTower R₁ R L] (L' : LieSubalgebra R L) :
IsScalarTower R₁ R L'
instance LieSubalgebra.lieAlgebra (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
LieAlgebra R L'

A Lie subalgebra forms a new Lie algebra.

Equations
@[simp]
theorem LieSubalgebra.zero_mem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
0 L'
theorem LieSubalgebra.add_mem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x y : L} :
x L'y L'x + y L'
theorem LieSubalgebra.sub_mem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x y : L} :
x L'y L'x - y L'
theorem LieSubalgebra.smul_mem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) (t : R) {x : L} (h : x L') :
t x L'
theorem LieSubalgebra.lie_mem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x y : L} (hx : x L') (hy : y L') :
x, y L'
theorem LieSubalgebra.mem_carrier {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x : L} :
x L'.carrier x L'
@[simp]
theorem LieSubalgebra.mem_mk_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set L) (h₁ : ∀ {a b : L}, a Sb Sa + b S) (h₂ : S 0) (h₃ : ∀ (c : R) {x : L}, x { carrier := S, add_mem' := h₁, zero_mem' := h₂ }.carrierc x { carrier := S, add_mem' := h₁, zero_mem' := h₂ }.carrier) (h₄ : ∀ {x y : L}, x { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carriery { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carrierx, y { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carrier) {x : L} :
x { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃, lie_mem' := h₄ } x S
@[simp]
theorem LieSubalgebra.mem_toSubmodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x : L} :
@[deprecated LieSubalgebra.mem_toSubmodule (since := "2024-12-30")]
theorem LieSubalgebra.mem_coe_submodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x : L} :

Alias of LieSubalgebra.mem_toSubmodule.

theorem LieSubalgebra.mem_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {x : L} :
x L' x L'
@[simp]
theorem LieSubalgebra.coe_bracket {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) (x y : L') :
x, y = x, y
theorem LieSubalgebra.ext_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) (x y : L') :
x = y x = y
theorem LieSubalgebra.coe_zero_iff_zero {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) (x : L') :
x = 0 x = 0
theorem LieSubalgebra.ext {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) (h : ∀ (x : L), x L₁' x L₂') :
L₁' = L₂'
theorem LieSubalgebra.ext_iff' {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) :
L₁' = L₂' ∀ (x : L), x L₁' x L₂'
@[simp]
theorem LieSubalgebra.mk_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set L) (h₁ : ∀ {a b : L}, a Sb Sa + b S) (h₂ : S 0) (h₃ : ∀ (c : R) {x : L}, x { carrier := S, add_mem' := h₁, zero_mem' := h₂ }.carrierc x { carrier := S, add_mem' := h₁, zero_mem' := h₂ }.carrier) (h₄ : ∀ {x y : L}, x { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carriery { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carrierx, y { carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃ }.carrier) :
{ carrier := S, add_mem' := h₁, zero_mem' := h₂, smul_mem' := h₃, lie_mem' := h₄ } = S
theorem LieSubalgebra.toSubmodule_mk {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (p : Submodule R L) (h : ∀ {x y : L}, x p.carriery p.carrierx, y p.carrier) :
{ toSubmodule := p, lie_mem' := h }.toSubmodule = p
@[deprecated LieSubalgebra.toSubmodule_mk (since := "2024-12-30")]
theorem LieSubalgebra.coe_to_submodule_mk {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (p : Submodule R L) (h : ∀ {x y : L}, x p.carriery p.carrierx, y p.carrier) :
{ toSubmodule := p, lie_mem' := h }.toSubmodule = p

Alias of LieSubalgebra.toSubmodule_mk.

theorem LieSubalgebra.coe_set_eq {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) :
L₁' = L₂' L₁' = L₂'
@[deprecated LieSubalgebra.toSubmodule_injective (since := "2024-12-30")]

Alias of LieSubalgebra.toSubmodule_injective.

@[simp]
theorem LieSubalgebra.toSubmodule_inj {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) :
L₁'.toSubmodule = L₂'.toSubmodule L₁' = L₂'
@[deprecated LieSubalgebra.toSubmodule_inj (since := "2024-12-30")]
theorem LieSubalgebra.coe_to_submodule_inj {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) :
L₁'.toSubmodule = L₂'.toSubmodule L₁' = L₂'

Alias of LieSubalgebra.toSubmodule_inj.

@[deprecated LieSubalgebra.toSubmodule_inj (since := "2024-12-29")]
theorem LieSubalgebra.toSubmodule_eq_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L₁' L₂' : LieSubalgebra R L) :
L₁'.toSubmodule = L₂'.toSubmodule L₁' = L₂'

Alias of LieSubalgebra.toSubmodule_inj.

theorem LieSubalgebra.coe_toSubmodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
L'.toSubmodule = L'
@[deprecated LieSubalgebra.coe_toSubmodule (since := "2024-12-30")]
theorem LieSubalgebra.coe_to_submodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
L'.toSubmodule = L'

Alias of LieSubalgebra.coe_toSubmodule.

instance LieSubalgebra.instBracketSubtypeMem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] :
Bracket (↥L') M
Equations
@[simp]
theorem LieSubalgebra.coe_bracket_of_module {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] (x : L') (m : M) :
x, m = x, m
instance LieSubalgebra.instIsLieTowerSubtypeMem {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] :
IsLieTower (↥L') L M
instance LieSubalgebra.lieRingModule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] :
LieRingModule (↥L') M

Given a Lie algebra L containing a Lie subalgebra L' ⊆ L, together with a Lie ring module M of L, we may regard M as a Lie ring module of L' by restriction.

Equations
instance LieSubalgebra.lieModule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] [Module R M] [LieModule R L M] :
LieModule R (↥L') M

Given a Lie algebra L containing a Lie subalgebra L' ⊆ L, together with a Lie module M of L, we may regard M as a Lie module of L' by restriction.

def LieModuleHom.restrictLie {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {M : Type w} [AddCommGroup M] [LieRingModule L M] {N : Type w₁} [AddCommGroup N] [LieRingModule L N] [Module R N] [Module R M] (f : M →ₗ⁅R,L N) (L' : LieSubalgebra R L) :
M →ₗ⁅R,L' N

An L-equivariant map of Lie modules M → N is L'-equivariant for any Lie subalgebra L' ⊆ L.

Equations
@[simp]
theorem LieModuleHom.coe_restrictLie {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) {M : Type w} [AddCommGroup M] [LieRingModule L M] {N : Type w₁} [AddCommGroup N] [LieRingModule L N] [Module R N] [Module R M] (f : M →ₗ⁅R,L N) :
(f.restrictLie L') = f
def LieSubalgebra.incl {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
L' →ₗ⁅R L

The embedding of a Lie subalgebra into the ambient space as a morphism of Lie algebras.

Equations
@[simp]
theorem LieSubalgebra.coe_incl {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
def LieSubalgebra.incl' {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
L' →ₗ⁅R,L' L

The embedding of a Lie subalgebra into the ambient space as a morphism of Lie modules.

Equations
@[simp]
theorem LieSubalgebra.coe_incl' {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (L' : LieSubalgebra R L) :
def LieHom.range {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) :

The range of a morphism of Lie algebras is a Lie subalgebra.

Equations
@[simp]
theorem LieHom.range_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) :
f.range = Set.range f
@[simp]
theorem LieHom.mem_range {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (x : L₂) :
x f.range ∃ (y : L), f y = x
theorem LieHom.mem_range_self {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (x : L) :
f x f.range
def LieHom.rangeRestrict {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) :

We can restrict a morphism to a (surjective) map to its range.

Equations
@[simp]
theorem LieHom.rangeRestrict_apply {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (x : L) :
f.rangeRestrict x = f x,
theorem LieHom.surjective_rangeRestrict {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) :
noncomputable def LieHom.equivRangeOfInjective {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (h : Function.Injective f) :

A Lie algebra is equivalent to its range under an injective Lie algebra morphism.

Equations
@[simp]
theorem LieHom.equivRangeOfInjective_apply {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (h : Function.Injective f) (x : L) :
(f.equivRangeOfInjective h) x = f x,
theorem Submodule.exists_lieSubalgebra_coe_eq_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (p : Submodule R L) :
(∃ (K : LieSubalgebra R L), K.toSubmodule = p) ∀ (x y : L), x py px, y p
@[simp]
theorem LieSubalgebra.incl_range {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K : LieSubalgebra R L) :
def LieSubalgebra.map {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (K : LieSubalgebra R L) :

The image of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the codomain.

Equations
@[simp]
theorem LieSubalgebra.mem_map {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (K : LieSubalgebra R L) (x : L₂) :
x map f K yK, f y = x
theorem LieSubalgebra.mem_map_submodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (K : LieSubalgebra R L) (e : L ≃ₗ⁅R L₂) (x : L₂) :
def LieSubalgebra.comap {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) (K₂ : LieSubalgebra R L₂) :

The preimage of a Lie subalgebra under a Lie algebra morphism is a Lie subalgebra of the domain.

Equations
Equations
  • One or more equations did not get rendered due to their size.
theorem LieSubalgebra.le_def {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) :
K K' K K'
@[deprecated LieSubalgebra.toSubmodule_le_toSubmodule (since := "2024-12-30")]

Alias of LieSubalgebra.toSubmodule_le_toSubmodule.

instance LieSubalgebra.instBot {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
@[simp]
theorem LieSubalgebra.bot_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
= {0}
@[deprecated LieSubalgebra.bot_toSubmodule (since := "2024-12-30")]

Alias of LieSubalgebra.bot_toSubmodule.

@[simp]
theorem LieSubalgebra.mem_bot {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (x : L) :
x x = 0
instance LieSubalgebra.instTop {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
@[simp]
theorem LieSubalgebra.top_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
@[deprecated LieSubalgebra.top_toSubmodule (since := "2024-12-30")]

Alias of LieSubalgebra.top_toSubmodule.

@[simp]
theorem LieSubalgebra.mem_top {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (x : L) :
theorem LieHom.range_eq_map {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] (f : L →ₗ⁅R L₂) :
instance LieSubalgebra.instMin {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
instance LieSubalgebra.instInfSet {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem LieSubalgebra.inf_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) :
(KK') = K K'
@[simp]
theorem LieSubalgebra.sInf_toSubmodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set (LieSubalgebra R L)) :
(sInf S).toSubmodule = sInf {x : Submodule R L | sS, s.toSubmodule = x}
@[deprecated LieSubalgebra.sInf_toSubmodule (since := "2024-12-30")]
theorem LieSubalgebra.sInf_coe_to_submodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set (LieSubalgebra R L)) :
(sInf S).toSubmodule = sInf {x : Submodule R L | sS, s.toSubmodule = x}

Alias of LieSubalgebra.sInf_toSubmodule.

@[simp]
theorem LieSubalgebra.sInf_coe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set (LieSubalgebra R L)) :
(sInf S) = sS, s
theorem LieSubalgebra.sInf_glb {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (S : Set (LieSubalgebra R L)) :
IsGLB S (sInf S)

The set of Lie subalgebras of a Lie algebra form a complete lattice.

We provide explicit values for the fields bot, top, inf to get more convenient definitions than we would otherwise obtain from completeLatticeOfInf.

Equations
  • One or more equations did not get rendered due to their size.
instance LieSubalgebra.instAdd {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
instance LieSubalgebra.instZero {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] :
Equations
Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem LieSubalgebra.add_eq_sup {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) :
K + K' = KK'
@[simp]
theorem LieSubalgebra.inf_toSubmodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) :
(KK').toSubmodule = K.toSubmoduleK'.toSubmodule
@[deprecated LieSubalgebra.inf_toSubmodule (since := "2024-12-30")]
theorem LieSubalgebra.inf_coe_to_submodule {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) :
(KK').toSubmodule = K.toSubmoduleK'.toSubmodule

Alias of LieSubalgebra.inf_toSubmodule.

@[simp]
theorem LieSubalgebra.mem_inf {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K K' : LieSubalgebra R L) (x : L) :
x KK' x K x K'
theorem LieSubalgebra.eq_bot_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K : LieSubalgebra R L) :
K = xK, x = 0
def LieSubalgebra.inclusion {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :
K →ₗ⁅R K'

Given two nested Lie subalgebras K ⊆ K', the inclusion K ↪ K' is a morphism of Lie algebras.

Equations
@[simp]
theorem LieSubalgebra.coe_inclusion {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') (x : K) :
((inclusion h) x) = x
theorem LieSubalgebra.inclusion_apply {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') (x : K) :
(inclusion h) x = x,
theorem LieSubalgebra.inclusion_injective {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :
def LieSubalgebra.ofLe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :

Given two nested Lie subalgebras K ⊆ K', we can view K as a Lie subalgebra of K', regarded as Lie algebra in its own right.

Equations
@[simp]
theorem LieSubalgebra.mem_ofLe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') (x : K') :
x ofLe h x K
theorem LieSubalgebra.ofLe_eq_comap_incl {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :
ofLe h = comap K'.incl K
@[simp]
theorem LieSubalgebra.coe_ofLe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :
noncomputable def LieSubalgebra.equivOfLe {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') :
K ≃ₗ⁅R (ofLe h)

Given nested Lie subalgebras K ⊆ K', there is a natural equivalence from K to its image in K'.

Equations
@[simp]
theorem LieSubalgebra.equivOfLe_apply {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {K K' : LieSubalgebra R L} (h : K K') (x : K) :
(equivOfLe h) x = (inclusion h) x,
theorem LieSubalgebra.map_le_iff_le_comap {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] {f : L →ₗ⁅R L₂} {K : LieSubalgebra R L} {K' : LieSubalgebra R L₂} :
map f K K' K comap f K'
theorem LieSubalgebra.gc_map_comap {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {L₂ : Type w} [LieRing L₂] [LieAlgebra R L₂] {f : L →ₗ⁅R L₂} :
def LieSubalgebra.lieSpan (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] (s : Set L) :

The Lie subalgebra of a Lie algebra L generated by a subset s ⊆ L.

Equations
theorem LieSubalgebra.mem_lieSpan {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {s : Set L} {x : L} :
x lieSpan R L s ∀ (K : LieSubalgebra R L), s Kx K
theorem LieSubalgebra.subset_lieSpan {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {s : Set L} :
s (lieSpan R L s)
theorem LieSubalgebra.lieSpan_le {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {s : Set L} {K : LieSubalgebra R L} :
lieSpan R L s K s K
theorem LieSubalgebra.lieSpan_mono {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {s t : Set L} (h : s t) :
lieSpan R L s lieSpan R L t
theorem LieSubalgebra.lieSpan_eq {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (K : LieSubalgebra R L) :
lieSpan R L K = K
theorem LieSubalgebra.coe_lieSpan_submodule_eq_iff {R : Type u} {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {p : Submodule R L} :
(lieSpan R L p).toSubmodule = p ∃ (K : LieSubalgebra R L), K.toSubmodule = p

lieSpan forms a Galois insertion with the coercion from LieSubalgebra to Set.

Equations
@[simp]
theorem LieSubalgebra.span_empty (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] :
@[simp]
theorem LieSubalgebra.span_univ (R : Type u) (L : Type v) [CommRing R] [LieRing L] [LieAlgebra R L] :
theorem LieSubalgebra.span_union (R : Type u) {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] (s t : Set L) :
lieSpan R L (s t) = lieSpan R L slieSpan R L t
theorem LieSubalgebra.span_iUnion (R : Type u) {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {ι : Sort u_1} (s : ιSet L) :
lieSpan R L (⋃ (i : ι), s i) = ⨆ (i : ι), lieSpan R L (s i)
theorem LieSubalgebra.lieSpan_induction (R : Type u) {L : Type v} [CommRing R] [LieRing L] [LieAlgebra R L] {s : Set L} {p : (x : L) → x lieSpan R L sProp} (mem : ∀ (x : L) (h : x s), p x ) (zero : p 0 ) (add : ∀ (x y : L) (hx : x lieSpan R L s) (hy : y lieSpan R L s), p x hxp y hyp (x + y) ) (smul : ∀ (a : R) (x : L) (hx : x lieSpan R L s), p x hxp (a x) ) {x : L} (lie : ∀ (x y : L) (hx : x lieSpan R L s) (hy : y lieSpan R L s), p x hxp y hyp x, y ) (hx : x lieSpan R L s) :
p x hx

An induction principle for span membership. If p holds for 0 and all elements of s, and is preserved under addition, scalar multiplication and the Lie bracket, then p holds for all elements of the Lie algebra spanned by s.

noncomputable def LieEquiv.ofInjective {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (f : L₁ →ₗ⁅R L₂) (h : Function.Injective f) :
L₁ ≃ₗ⁅R f.range

An injective Lie algebra morphism is an equivalence onto its range.

Equations
@[simp]
theorem LieEquiv.ofInjective_apply {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (f : L₁ →ₗ⁅R L₂) (h : Function.Injective f) (x : L₁) :
((ofInjective f h) x) = f x
def LieEquiv.ofEq {R : Type u} {L₁ : Type v} [CommRing R] [LieRing L₁] [LieAlgebra R L₁] (L₁' L₁'' : LieSubalgebra R L₁) (h : L₁' = L₁'') :
L₁' ≃ₗ⁅R L₁''

Lie subalgebras that are equal as sets are equivalent as Lie algebras.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem LieEquiv.ofEq_apply {R : Type u} {L₁ : Type v} [CommRing R] [LieRing L₁] [LieAlgebra R L₁] (L L' : LieSubalgebra R L₁) (h : L = L') (x : L) :
((ofEq L L' h) x) = x
def LieEquiv.lieSubalgebraMap {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (L₁'' : LieSubalgebra R L₁) (e : L₁ ≃ₗ⁅R L₂) :
L₁'' ≃ₗ⁅R (LieSubalgebra.map e.toLieHom L₁'')

An equivalence of Lie algebras restricts to an equivalence from any Lie subalgebra onto its image.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem LieEquiv.lieSubalgebraMap_apply {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (L₁'' : LieSubalgebra R L₁) (e : L₁ ≃ₗ⁅R L₂) (x : L₁'') :
((lieSubalgebraMap L₁'' e) x) = e x
def LieEquiv.ofSubalgebras {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (L₁' : LieSubalgebra R L₁) (L₂' : LieSubalgebra R L₂) (e : L₁ ≃ₗ⁅R L₂) (h : LieSubalgebra.map e.toLieHom L₁' = L₂') :
L₁' ≃ₗ⁅R L₂'

An equivalence of Lie algebras restricts to an equivalence from any Lie subalgebra onto its image.

Equations
  • One or more equations did not get rendered due to their size.
@[simp]
theorem LieEquiv.ofSubalgebras_apply {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (L₁' : LieSubalgebra R L₁) (L₂' : LieSubalgebra R L₂) (e : L₁ ≃ₗ⁅R L₂) (h : LieSubalgebra.map e.toLieHom L₁' = L₂') (x : L₁') :
((ofSubalgebras L₁' L₂' e h) x) = e x
@[simp]
theorem LieEquiv.ofSubalgebras_symm_apply {R : Type u} {L₁ : Type v} {L₂ : Type w} [CommRing R] [LieRing L₁] [LieRing L₂] [LieAlgebra R L₁] [LieAlgebra R L₂] (L₁' : LieSubalgebra R L₁) (L₂' : LieSubalgebra R L₂) (e : L₁ ≃ₗ⁅R L₂) (h : LieSubalgebra.map e.toLieHom L₁' = L₂') (x : L₂') :
((ofSubalgebras L₁' L₂' e h).symm x) = e.symm x