Documentation

Mathlib.Algebra.NonAssoc.PreLie.Basic

Pre-Lie rings and algebras #

In this file we introduce left and right pre-Lie rings, defined as a NonUnitalNonAssocRing where the associator associator x y z := (x * y) * z - x * (y * z) is left or right symmetric, respectively.

We prove that every Left(Right)PreLieRing L is a Right(Left)PreLieRing L with the opposite mul. The equivalence is simple given by op : L ≃* Lᵐᵒᵖ.

Everything holds for the algebra versions where L is also an R-Module over a commutative ring R.

Main definitions #

All are a defined as a NonUnitalNonAssocRing whose associator satisfies an identity.

Main results #

Implementation notes #

There are left and right versions of the structures, equivalent via ᵐᵒᵖ. Perhaps one could be favored but there is no real reason to.

References #

F. Chapoton, M. Livernet, Pre-Lie algebras and the rooted trees operad D. Manchon, A short survey on pre-Lie algebras J.-M. Oudom, D. Guin, On the Lie enveloping algebra of a pre-Lie algebra https://ncatlab.org/nlab/show/pre-Lie+algebra

class LeftPreLieRing (L : Type u_1) extends NonUnitalNonAssocRing L :
Type u_1

LeftPreLieRings are NonUnitalNonAssocRings such that the associator is symmetric in the first two variables.

Instances
    class RightPreLieRing (L : Type u_1) extends NonUnitalNonAssocRing L :
    Type u_1

    RightPreLieRings are NonUnitalNonAssocRings such that the associator is symmetric in the last two variables.

    Instances
      class LeftPreLieAlgebra (R : Type u_1) [CommRing R] (L : Type u_2) [LeftPreLieRing L] extends Module R L, IsScalarTower R L L, SMulCommClass R L L :
      Type (max u_1 u_2)

      A LeftPreLieAlgebra is a LeftPreLieRing with an action of a CommRing satisfying r • x * y = r • (x * y) and x * (r • y) = r • (x * y).

      Instances
        theorem LeftPreLieAlgebra.ext_iff {R : Type u_1} {inst✝ : CommRing R} {L : Type u_2} {inst✝¹ : LeftPreLieRing L} {x y : LeftPreLieAlgebra R L} :
        theorem LeftPreLieAlgebra.ext {R : Type u_1} {inst✝ : CommRing R} {L : Type u_2} {inst✝¹ : LeftPreLieRing L} {x y : LeftPreLieAlgebra R L} (smul : SMul.smul = SMul.smul) :
        x = y
        class RightPreLieAlgebra (R : Type u_1) [CommRing R] (L : Type u_2) [RightPreLieRing L] extends Module R L, IsScalarTower R L L, SMulCommClass R L L :
        Type (max u_1 u_2)

        A RightPreLieAlgebra is a RightPreLieRing with an action of a CommRing satisfying r • x * y = r • (x * y) and x * (r • y) = r • (x * y).

        Instances
          theorem RightPreLieAlgebra.ext {R : Type u_1} {inst✝ : CommRing R} {L : Type u_2} {inst✝¹ : RightPreLieRing L} {x y : RightPreLieAlgebra R L} (smul : SMul.smul = SMul.smul) :
          x = y
          theorem RightPreLieAlgebra.ext_iff {R : Type u_1} {inst✝ : CommRing R} {L : Type u_2} {inst✝¹ : RightPreLieRing L} {x y : RightPreLieAlgebra R L} :
          theorem LeftPreLieRing.assoc_symm {L : Type u_2} [LeftPreLieRing L] (x y z : L) :
          associator x y z = associator y x z

          Every left pre-Lie ring is a right pre-Lie ring with the opposite multiplication

          Equations

          Every left pre-Lie algebra is a right pre-Lie algebra with the opposite multiplication

          Equations
          theorem RightPreLieRing.assoc_symm {L : Type u_2} [RightPreLieRing L] (x y z : L) :
          associator x y z = associator x z y

          Every left pre-Lie ring is a right pre-Lie ring with the opposite multiplication

          Equations

          Every left pre-Lie algebra is a right pre-Lie algebra with the opposite multiplication

          Equations