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 #
- Every
LeftPreLieRing
is aRightPreLieRing
with the opposite multiplication.
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
LeftPreLieRing
s are NonUnitalNonAssocRing
s such that the associator
is symmetric in the
first two variables.
Instances
RightPreLieRing
s are NonUnitalNonAssocRing
s such that the associator
is symmetric in the
last two variables.
Instances
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)
.
- smul : R → L → L
Instances
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)
.
- smul : R → L → L
Instances
Every left pre-Lie ring is a right pre-Lie ring with the opposite multiplication
Equations
- LeftPreLieRing.instRightPreLieRingMulOpposite = { toNonUnitalNonAssocRing := MulOpposite.instNonUnitalNonAssocRing, assoc_symm' := ⋯ }
Every left pre-Lie algebra is a right pre-Lie algebra with the opposite multiplication
Equations
- LeftPreLieAlgebra.instRightPreLieAlgebraMulOpposite = { toModule := MulOpposite.instModule R, toIsScalarTower := ⋯, toSMulCommClass := ⋯ }
Every left pre-Lie ring is a right pre-Lie ring with the opposite multiplication
Equations
- RightPreLieRing.instLeftPreLieRingMulOpposite = { toNonUnitalNonAssocRing := MulOpposite.instNonUnitalNonAssocRing, assoc_symm' := ⋯ }
Every left pre-Lie algebra is a right pre-Lie algebra with the opposite multiplication
Equations
- RightPreLieAlgebra.instLeftPreLieAlgebraMulOpposite = { toModule := MulOpposite.instModule R, toIsScalarTower := ⋯, toSMulCommClass := ⋯ }