Lifting monoid algebras #
This file defines liftNC
. For the definition of MonoidAlgebra.lift
, see
Mathlib/Algebra/MonoidAlgebra/Basic.lean
.
Main results #
MonoidAlgebra.liftNC
,AddMonoidAlgebra.liftNC
: lift a homomorphismf : k →+ R
and a functiong : G → R
to a homomorphismMonoidAlgebra k G →+ R
.
Multiplicative monoids #
A non-commutative version of MonoidAlgebra.lift
: given an additive homomorphism f : k →+ R
and a homomorphism g : G → R
, returns the additive homomorphism from
MonoidAlgebra k G
such that liftNC f g (single a b) = f b * g a
. If f
is a ring homomorphism
and the range of either f
or g
is in center of R
, then the result is a ring homomorphism. If
R
is a k
-algebra and f = algebraMap k R
, then the result is an algebra homomorphism called
MonoidAlgebra.lift
.
Equations
- MonoidAlgebra.liftNC f g = Finsupp.liftAddHom fun (x : G) => (AddMonoidHom.mulRight (g x)).comp f
Instances For
Semiring structure #
liftNC
as a RingHom
, for when f x
and g y
commute
Equations
- MonoidAlgebra.liftNCRingHom f g h_comm = { toFun := (↑(MonoidAlgebra.liftNC ↑f ⇑g)).toFun, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Additive monoids #
A non-commutative version of AddMonoidAlgebra.lift
: given an additive homomorphism
f : k →+ R
and a map g : Multiplicative G → R
, returns the additive
homomorphism from k[G]
such that liftNC f g (single a b) = f b * g a
. If f
is a ring homomorphism and the range of either f
or g
is in center of R
, then the result is a
ring homomorphism. If R
is a k
-algebra and f = algebraMap k R
, then the result is an algebra
homomorphism called AddMonoidAlgebra.lift
.
Equations
- AddMonoidAlgebra.liftNC f g = Finsupp.liftAddHom fun (x : G) => (AddMonoidHom.mulRight (g (Multiplicative.ofAdd x))).comp f
Instances For
Semiring structure #
liftNC
as a RingHom
, for when f
and g
commute
Equations
- AddMonoidAlgebra.liftNCRingHom f g h_comm = { toFun := (↑(AddMonoidAlgebra.liftNC ↑f ⇑g)).toFun, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }