Documentation

Mathlib.Algebra.Lie.SkewAdjoint

Lie algebras of skew-adjoint endomorphisms of a bilinear form #

When a module carries a bilinear form, the Lie algebra of endomorphisms of the module contains a distinguished Lie subalgebra: the skew-adjoint endomorphisms. Such subalgebras are important because they provide a simple, explicit construction of the so-called classical Lie algebras.

This file defines the Lie subalgebra of skew-adjoint endomorphisms cut out by a bilinear form on a module and proves some basic related results. It also provides the corresponding definitions and results for the Lie algebra of square matrices.

Main definitions #

Tags #

lie algebra, skew-adjoint, bilinear form

Given an R-module M, equipped with a bilinear form, the skew-adjoint endomorphisms form a Lie subalgebra of the Lie algebra of endomorphisms.

Equations
Instances For

    An equivalence of modules with bilinear forms gives equivalence of Lie algebras of skew-adjoint endomorphisms.

    Equations
    Instances For
      @[simp]
      theorem skewAdjointLieSubalgebraEquiv_apply {R : Type u} {M : Type v} [CommRing R] [AddCommGroup M] [Module R M] (B : LinearMap.BilinForm R M) {N : Type w} [AddCommGroup N] [Module R N] (e : N ≃ₗ[R] M) (f : (skewAdjointLieSubalgebra (LinearMap.compl₁₂ B e e))) :
      ((skewAdjointLieSubalgebraEquiv B e) f) = e.lieConj f
      @[simp]
      theorem skewAdjointLieSubalgebraEquiv_symm_apply {R : Type u} {M : Type v} [CommRing R] [AddCommGroup M] [Module R M] (B : LinearMap.BilinForm R M) {N : Type w} [AddCommGroup N] [Module R N] (e : N ≃ₗ[R] M) (f : (skewAdjointLieSubalgebra B)) :
      ((skewAdjointLieSubalgebraEquiv B e).symm f) = e.symm.lieConj f
      theorem Matrix.lie_transpose {R : Type u} {n : Type w} [CommRing R] [DecidableEq n] [Fintype n] (A : Matrix n n R) (B : Matrix n n R) :
      A, B.transpose = B.transpose, A.transpose
      def skewAdjointMatricesLieSubalgebra {R : Type u} {n : Type w} [CommRing R] [DecidableEq n] [Fintype n] (J : Matrix n n R) :

      The Lie subalgebra of skew-adjoint square matrices corresponding to a square matrix J.

      Equations
      Instances For

        An invertible matrix P gives a Lie algebra equivalence between those endomorphisms that are skew-adjoint with respect to a square matrix J and those with respect to PᵀJP.

        Equations
        Instances For
          def skewAdjointMatricesLieSubalgebraEquivTranspose {R : Type u} {n : Type w} [CommRing R] [DecidableEq n] [Fintype n] (J : Matrix n n R) {m : Type w} [DecidableEq m] [Fintype m] (e : Matrix n n R ≃ₐ[R] Matrix m m R) (h : ∀ (A : Matrix n n R), (e A).transpose = e A.transpose) :

          An equivalence of matrix algebras commuting with the transpose endomorphisms restricts to an equivalence of Lie algebras of skew-adjoint matrices.

          Equations
          Instances For
            @[simp]
            theorem skewAdjointMatricesLieSubalgebraEquivTranspose_apply {R : Type u} {n : Type w} [CommRing R] [DecidableEq n] [Fintype n] (J : Matrix n n R) {m : Type w} [DecidableEq m] [Fintype m] (e : Matrix n n R ≃ₐ[R] Matrix m m R) (h : ∀ (A : Matrix n n R), (e A).transpose = e A.transpose) (A : (skewAdjointMatricesLieSubalgebra J)) :