Quotients of Lie algebras and Lie modules #
Given a Lie submodule of a Lie module, the quotient carries a natural Lie module structure. In the special case that the Lie module is the Lie algebra itself via the adjoint action, the submodule is a Lie ideal and the quotient carries a natural Lie algebra structure.
We define these quotient structures here. A notable omission at the time of writing (February 2021) is a statement and proof of the universal property of these quotients.
Main definitions #
Tags #
lie algebra, quotient
The quotient of a Lie module by a Lie submodule. It is a Lie module.
Equations
- LieSubmodule.instHasQuotient = { quotient' := fun (N : LieSubmodule R L M) => M ⧸ ↑N }
Equations
Equations
Equations
- LieSubmodule.Quotient.inhabited = { default := 0 }
Map sending an element of M to the corresponding element of M/N, when N is a
lie_submodule of the lie_module N.
Instances For
Given a Lie module M over a Lie algebra L, together with a Lie submodule N ⊆ M, there
is a natural linear map from L to the endomorphisms of M leaving N invariant.
Equations
- LieSubmodule.Quotient.lieSubmoduleInvariant = LinearMap.codRestrict ((↑N).compatibleMaps ↑N) ↑(LieModule.toEnd R L M) ⋯
Instances For
Given a Lie module M over a Lie algebra L, together with a Lie submodule N ⊆ M, there
is a natural Lie algebra morphism from L to the linear endomorphism of the quotient M/N.
Equations
- LieSubmodule.Quotient.actionAsEndoMap N = { toLinearMap := (↑N).mapQLinear ↑N ∘ₗ LieSubmodule.Quotient.lieSubmoduleInvariant, map_lie' := ⋯ }
Instances For
Given a Lie module M over a Lie algebra L, together with a Lie submodule N ⊆ M, there is
a natural bracket action of L on the quotient M/N.
Equations
- LieSubmodule.Quotient.actionAsEndoMapBracket N = { bracket := fun (x : L) (n : M ⧸ N) => ((LieSubmodule.Quotient.actionAsEndoMap N) x) n }
Equations
- LieSubmodule.Quotient.lieQuotientLieRingModule N = { toBracket := LieSubmodule.Quotient.actionAsEndoMapBracket N, add_lie := ⋯, lie_add := ⋯, leibniz_lie := ⋯ }
The quotient of a Lie module by a Lie submodule, is a Lie module.
Equations
- LieSubmodule.Quotient.lieQuotientHasBracket I = { bracket := fun (x y : L ⧸ I) => Quotient.liftOn₂' x y (fun (x' y' : L) => LieSubmodule.Quotient.mk ⁅x', y'⁆) ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- LieSubmodule.Quotient.lieQuotientLieAlgebra I = { toModule := LieSubmodule.Quotient.module, lie_smul := ⋯ }
LieSubmodule.Quotient.mk as a LieModuleHom.
Equations
- LieSubmodule.Quotient.mk' N = { toFun := LieSubmodule.Quotient.mk, map_add' := ⋯, map_smul' := ⋯, map_lie' := ⋯ }
Instances For
Two LieModuleHoms from a quotient lie module are equal if their compositions with
LieSubmodule.Quotient.mk' are equal.
See note [partially-applied ext lemmas].
The first isomorphism theorem for morphisms of Lie algebras.
Equations
- f.quotKerEquivRange = { toFun := ⇑(↑f).quotKerEquivRange, map_add' := ⋯, map_smul' := ⋯, map_lie' := ⋯, invFun := (↑f).quotKerEquivRange.invFun, left_inv := ⋯, right_inv := ⋯ }