Geck's construction of a Lie algebra associated to a root system #
This file contains an implementation of Geck's construction of a semisimple Lie algebra from a reduced crystallographic root system. It follows Geck quite closely.
Main definitions: #
RootPairing.GeckConstruction.lieAlgebra
: the Geck construction of the Lie algebra associated to a root system with distinguished base.RootPairing.GeckConstruction.cartanSubalgebra
: a distinguished subalgebra corresponding to a Cartan subalgebra of the Geck construction.RootPairing.GeckConstruction.cartanSubalgebra_le_lieAlgebra
: the distinguished subalgebra is contained in the Geck construction.
Alternative approaches #
The are at least three ways to construct a Lie algebra from a root system:
- As a quotient of a free Lie algebra, using the Serre relations
- Directly defining the Lie bracket on $H ⊕ K^∣Φ|$
- The Geck construction
We comment on these as follows:
- This construction takes just a matrix as input. It yields a semisimple Lie algebra iff the
matrix is a Cartan matrix but it is quite a lot of work to prove this. On the other hand, it also
allows construction of Kac-Moody Lie algebras. It has been implemented as
Matrix.ToLieAlgebra
but as of May 2025, almost nothing has been proved about it in Mathlib. - This construction takes a root system with base as input, together with sufficient additional data to determine a collection of extraspecial pairs of roots. The additional data for the extraspecial pairs is required to pin down certain signs when defining the Lie bracket. (These signs can be interpreted as a set-theoretic splitting of Tits's extension of the Weyl group by an elementary 2-group of order $2^l$ where $l$ is the rank.)
- This construction takes a root system with base as input and is implemented here.
There seems to be no known construction of a Lie algebra from a root system without first choosing a base: https://mathoverflow.net/questions/495434/
Part of an sl₂
triple used in Geck's construction of a Lie algebra from a root system.
Equations
- RootPairing.GeckConstruction.h i = Matrix.fromBlocks 0 0 0 (Matrix.diagonal fun (x : ι) => ↑(P.pairingIn ℤ x ↑i))
Instances For
Part of an sl₂
triple used in Geck's construction of a Lie algebra from a root system.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Part of an sl₂
triple used in Geck's construction of a Lie algebra from a root system.
Equations
- One or more equations did not get rendered due to their size.
Instances For
An involutive matrix which can transfer results between RootPairing.GeckConstruction.e
and
RootPairing.GeckConstruction.f
.
Equations
- RootPairing.GeckConstruction.ω b = Matrix.fromBlocks 1 0 0 (Matrix.of fun (i j : ι) => if i = -j then 1 else 0)
Instances For
Geck's construction of the Lie algebra associated to a root system with distinguished base.
Note that it is convenient to include range h
in the Lie span, to make it elementary that it
contains RootPairing.GeckConstruction.cartanSubalgebra
, and not depend on
RootPairing.GeckConstruction.lie_e_f_same
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A distinguished subalgebra corresponding to a Cartan subalgebra of the Geck construction.
See also RootPairing.GeckConstruction.cartanSubalgebra'
.
Equations
- RootPairing.GeckConstruction.cartanSubalgebra b = { toSubmodule := Submodule.span R (Set.range RootPairing.GeckConstruction.h), lie_mem' := ⋯ }
Instances For
A distinguished Cartan subalgebra of the Geck construction.
Equations
Instances For
The element h i
, as a term of the Cartan subalgebra cartanSubalgebra' b
.
Equations
Instances For
Geck's name for the "left" basis elements of b.support ⊕ ι
.
Equations
Instances For
Geck's name for the "right" basis elements of b.support ⊕ ι
.
Equations
- RootPairing.GeckConstruction.v b i = Pi.single (Sum.inr i) 1
Instances For
The conjugation x ↦ ωxω
as an equivalence of Lie algebras.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The equivalence x ↦ ωxω
as an operation on Lie submodules of the Geck construction.
Equations
- RootPairing.GeckConstruction.ωConjLieSubmodule N = { toSubmodule := Submodule.comap (Matrix.toLin' (RootPairing.GeckConstruction.ω b)) ↑N, lie_mem := ⋯ }