Documentation

Mathlib.Algebra.Lie.Semisimple.Basic

Semisimple Lie algebras #

The famous Cartan-Dynkin-Killing classification of semisimple Lie algebras renders them one of the most important classes of Lie algebras. In this file we prove basic results about simple and semisimple Lie algebras.

Main declarations #

Tags #

lie algebra, radical, simple, semisimple

theorem LieModule.nontrivial_of_isIrreducible (R : Type u_1) (L : Type u_2) (M : Type u_3) [CommRing R] [LieRing L] [AddCommGroup M] [Module R M] [LieRingModule L M] [IsIrreducible R L M] :
theorem LieAlgebra.hasTrivialRadical_of_no_solvable_ideals {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] (h : ∀ (I : LieIdeal R L), IsSolvable II = ) :
theorem LieAlgebra.IsSimple.eq_top_of_isAtom {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] [IsSimple R L] (I : LieIdeal R L) (hI : IsAtom I) :
I =
@[simp]
theorem LieAlgebra.IsSimple.isAtom_iff_eq_top {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] [IsSimple R L] (I : LieIdeal R L) :
theorem LieAlgebra.IsSemisimple.isSimple_of_isAtom {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] [IsSemisimple R L] (I : LieIdeal R L) (hI : IsAtom I) :
IsSimple R I
@[instance 100]
Equations
@[instance 100]
noncomputable instance LieAlgebra.IsSemisimple.instBooleanAlgebra {R : Type u_1} {L : Type u_2} [CommRing R] [LieRing L] [LieAlgebra R L] [IsSemisimple R L] :
Equations
@[instance 100]

A semisimple Lie algebra has trivial radical.

@[instance 100]
instance LieAlgebra.IsSimple.instIsSemisimple (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] [IsSimple R L] :

A simple Lie algebra is semisimple.

An abelian Lie algebra with trivial radical is trivial.

theorem LieAlgebra.abelian_radical_iff_solvable_is_abelian (R : Type u_1) (L : Type u_2) [CommRing R] [LieRing L] [LieAlgebra R L] [IsNoetherian R L] :
IsLieAbelian (radical R L) ∀ (I : LieIdeal R L), IsSolvable IIsLieAbelian I

The two properties shown to be equivalent here are possible definitions for a Lie algebra to be reductive.

Note that there is absolutely no agreement on what the label 'reductive' should mean when the coefficients are not a field of characteristic zero.