Lie algebras from Cartan matrices #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
Split semi-simple Lie algebras are uniquely determined by their Cartan matrix. Indeed, if A
is
an l × l
Cartan matrix, the corresponding Lie algebra may be obtained as the Lie algebra on
3l
generators: $H_1, H_2, \ldots H_l, E_1, E_2, \ldots, E_l, F_1, F_2, \ldots, F_l$
subject to the following relations:
$$
\begin{align}
[H_i, H_j] &= 0\\
[E_i, F_i] &= H_i\\
[E_i, F_j] &= 0 \quad\text{if $i \ne j$}\\
[H_i, E_j] &= A_{ij}E_j\\
[H_i, F_j] &= -A_{ij}F_j\\
ad(E_i)^{1 - A_{ij}}(E_j) &= 0 \quad\text{if $i \ne j$}\\
ad(F_i)^{1 - A_{ij}}(F_j) &= 0 \quad\text{if $i \ne j$}\\
\end{align}
$$
In this file we provide the above construction. It is defined for any square matrix of integers but the results for non-Cartan matrices should be regarded as junk.
Recall that a Cartan matrix is a square matrix of integers A
such that:
- For diagonal values we have:
A i i = 2
. - For off-diagonal values (
i ≠ j
) we have:A i j ∈ {-3, -2, -1, 0}
. A i j = 0 ↔ A j i = 0
.- There exists a diagonal matrix
D
over ℝ such thatD ⬝ A ⬝ D⁻¹
is symmetric positive definite.
Alternative construction #
This construction is sometimes performed within the free unital associative algebra
free_algebra R X
, rather than within the free Lie algebra free_lie_algebra R X
, as we do here.
However the difference is illusory since the construction stays inside the Lie subalgebra of
free_algebra R X
generated by X
, and this is naturally isomorphic to free_lie_algebra R X
(though the proof of this seems to require Poincaré–Birkhoff–Witt).
Definitions of exceptional Lie algebras #
This file also contains the Cartan matrices of the exceptional Lie algebras. By using these in the above construction, it thus provides definitions of the exceptional Lie algebras. These definitions make sense over any commutative ring. When the ring is ℝ, these are the split real forms of the exceptional semisimple Lie algebras.
References #
-
N. Bourbaki, Lie Groups and Lie Algebras, Chapters 4--6 plates V -- IX, pages 275--290
-
N. Bourbaki, Lie Groups and Lie Algebras, Chapters 7--9 chapter VIII, §4.3
-
J.P. Serre, Complex Semisimple Lie Algebras chapter VI, appendix
Main definitions #
matrix.to_lie_algebra
cartan_matrix.E₆
cartan_matrix.E₇
cartan_matrix.E₈
cartan_matrix.F₄
cartan_matrix.G₂
lie_algebra.e₆
lie_algebra.e₇
lie_algebra.e₈
lie_algebra.f₄
lie_algebra.g₂
Tags #
lie algebra, semi-simple, cartan matrix
- H : Π {B : Type v} [_inst_2 : decidable_eq B] [_inst_3 : fintype B], B → cartan_matrix.generators B
- E : Π {B : Type v} [_inst_2 : decidable_eq B] [_inst_3 : fintype B], B → cartan_matrix.generators B
- F : Π {B : Type v} [_inst_2 : decidable_eq B] [_inst_3 : fintype B], B → cartan_matrix.generators B
The generators of the free Lie algebra from which we construct the Lie algebra of a Cartan matrix as a quotient.
Instances for cartan_matrix.generators
- cartan_matrix.generators.has_sizeof_inst
- cartan_matrix.generators.inhabited
The terms correpsonding to the ⁅H, H⁆
-relations.
Equations
- cartan_matrix.relations.HH R = function.uncurry (λ (i j : B), ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.H) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.H) j⁆)
The terms correpsonding to the ⁅E, F⁆
-relations.
Equations
- cartan_matrix.relations.EF R = function.uncurry (λ (i j : B), ite (i = j) (⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.E) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.F) i⁆ - (free_lie_algebra.of R ∘ cartan_matrix.generators.H) i) ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.E) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.F) j⁆)
The terms correpsonding to the ⁅H, E⁆
-relations.
Equations
- cartan_matrix.relations.HE R A = function.uncurry (λ (i j : B), ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.H) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.E) j⁆ - A i j • (free_lie_algebra.of R ∘ cartan_matrix.generators.E) j)
The terms correpsonding to the ⁅H, F⁆
-relations.
Equations
- cartan_matrix.relations.HF R A = function.uncurry (λ (i j : B), ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.H) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.F) j⁆ + A i j • (free_lie_algebra.of R ∘ cartan_matrix.generators.F) j)
The terms correpsonding to the ad E
-relations.
Note that we use int.to_nat
so that we can take the power and that we do not bother
restricting to the case i ≠ j
since these relations are zero anyway. We also defensively
ensure this with ad_E_of_eq_eq_zero
.
Equations
- cartan_matrix.relations.ad_E R A = function.uncurry (λ (i j : B), ⇑(⇑(lie_algebra.ad R (free_lie_algebra R (cartan_matrix.generators B))) ((free_lie_algebra.of R ∘ cartan_matrix.generators.E) i) ^ (-A i j).to_nat) ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.E) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.E) j⁆)
The terms correpsonding to the ad F
-relations.
See also ad_E
docstring.
Equations
- cartan_matrix.relations.ad_F R A = function.uncurry (λ (i j : B), ⇑(⇑(lie_algebra.ad R (free_lie_algebra R (cartan_matrix.generators B))) ((free_lie_algebra.of R ∘ cartan_matrix.generators.F) i) ^ (-A i j).to_nat) ⁅(free_lie_algebra.of R ∘ cartan_matrix.generators.F) i, (free_lie_algebra.of R ∘ cartan_matrix.generators.F) j⁆)
The union of all the relations as a subset of the free Lie algebra.
Equations
- cartan_matrix.relations.to_set R A = set.range (cartan_matrix.relations.HH R) ∪ set.range (cartan_matrix.relations.EF R) ∪ set.range (cartan_matrix.relations.HE R A) ∪ set.range (cartan_matrix.relations.HF R A) ∪ set.range (cartan_matrix.relations.ad_E R A) ∪ set.range (cartan_matrix.relations.ad_F R A)
The ideal of the free Lie algebra generated by the relations.
Equations
The Lie algebra corresponding to a Cartan matrix.
Note that it is defined for any matrix of integers. Its value for non-Cartan matrices should be regarded as junk.
Equations
Instances for matrix.to_lie_algebra
The Cartan matrix of type e₆. See [Bou02] plate V, page 277.
The corresponding Dynkin diagram is:
o
|
o --- o --- o --- o --- o
The Cartan matrix of type e₇. See [Bou02] plate VI, page 281.
The corresponding Dynkin diagram is:
o
|
o --- o --- o --- o --- o --- o
The Cartan matrix of type e₈. See [Bou02] plate VII, page 285.
The corresponding Dynkin diagram is:
o
|
o --- o --- o --- o --- o --- o --- o
The Cartan matrix of type g₂. See [Bou02] plate IX, page 290.
The corresponding Dynkin diagram is:
o ≡>≡ o
Actually we are using the transpose of Bourbaki's matrix. This is to make this matrix consistent
with cartan_matrix.F₄
, in the sense that all non-zero values below the diagonal are -1.
The exceptional split Lie algebra of type e₆.
The exceptional split Lie algebra of type e₇.
The exceptional split Lie algebra of type e₈.
The exceptional split Lie algebra of type f₄.
The exceptional split Lie algebra of type g₂.