Introduce SMulWithZero #
In analogy with the usual monoid action on a Type M, we introduce an action of a
MonoidWithZero on a Type with 0.
In particular, for Types R and M, both containing 0, we define SMulWithZero R M to
be the typeclass where the products r • 0 and 0 • m vanish for all r : R and all m : M.
Moreover, in the case in which R is a MonoidWithZero, we introduce the typeclass
MulActionWithZero R M, mimicking group actions and having an absorbing 0 in R.
Thus, the action is required to be compatible with
- the unit of the monoid, acting as the identity;
- the zero of the MonoidWithZero, acting as zero;
- associativity of the monoid.
We also add an instance:
- any MonoidWithZerohas aMulActionWithZero R Racting on itself.
Main declarations #
- smulMonoidWithZeroHom: Scalar multiplication bundled as a morphism of monoids with zero.
SMulWithZero is a class consisting of a Type R with 0 ∈ R and a scalar multiplication
of R on a Type M with 0, such that the equality r • m = 0 holds if at least one among r
or m equals 0.
- smul : R → M → M
- Scalar multiplication by the scalar - 0is- 0.
Instances
Scalar multiplication by the scalar 0 is 0.
Equations
Like MulZeroClass.toSMulWithZero, but multiplies on the right.
Equations
Pullback a SMulWithZero structure along an injective zero-preserving homomorphism.
See note [reducible non-instances].
Equations
- Function.Injective.smulWithZero f hf smul = SMulWithZero.mk ⋯
Instances For
Pushforward a SMulWithZero structure along a surjective zero-preserving homomorphism.
See note [reducible non-instances].
Equations
- Function.Surjective.smulWithZero f hf smul = SMulWithZero.mk ⋯
Instances For
Compose a SMulWithZero with a ZeroHom, with action f r' • m
Equations
Instances For
Equations
- AddMonoid.natSMulWithZero = SMulWithZero.mk ⋯
Equations
- AddGroup.intSMulWithZero = SMulWithZero.mk ⋯
An action of a monoid with zero R on a Type M, also with 0, extends MulAction and
is compatible with 0 (both in R and in M), with 1 ∈ R, and with associativity of
multiplication on the monoid M.
- smul : R → M → M
- Scalar multiplication by any element send - 0to- 0.
- Scalar multiplication by the scalar - 0is- 0.
Instances
Scalar multiplication by any element send 0 to 0.
Scalar multiplication by the scalar 0 is 0.
Equations
See also Semiring.toModule
Equations
- MonoidWithZero.toMulActionWithZero R = let __src := MulZeroClass.toSMulWithZero R; let __src_1 := Monoid.toMulAction R; MulActionWithZero.mk ⋯ ⋯
Like MonoidWithZero.toMulActionWithZero, but multiplies on the right. See also
Semiring.toOppositeModule
Equations
- MonoidWithZero.toOppositeMulActionWithZero R = let __src := MulZeroClass.toOppositeSMulWithZero R; let __src_1 := Monoid.toOppositeMulAction; MulActionWithZero.mk ⋯ ⋯
Pullback a MulActionWithZero structure along an injective zero-preserving homomorphism.
See note [reducible non-instances].
Equations
- Function.Injective.mulActionWithZero f hf smul = let __src := Function.Injective.mulAction (⇑f) hf smul; let __src_1 := Function.Injective.smulWithZero f hf smul; MulActionWithZero.mk ⋯ ⋯
Instances For
Pushforward a MulActionWithZero structure along a surjective zero-preserving homomorphism.
See note [reducible non-instances].
Equations
- Function.Surjective.mulActionWithZero f hf smul = let __src := Function.Surjective.mulAction (⇑f) hf smul; let __src_1 := Function.Surjective.smulWithZero f hf smul; MulActionWithZero.mk ⋯ ⋯
Instances For
Compose a MulActionWithZero with a MonoidWithZeroHom, with action f r' • m
Equations
- MulActionWithZero.compHom M f = let __src := SMulWithZero.compHom M ↑f; MulActionWithZero.mk ⋯ ⋯
Instances For
Scalar multiplication as a monoid homomorphism with zero.
Equations
- smulMonoidWithZeroHom = let __src := smulMonoidHom; { toFun := (↑__src).toFun, map_zero' := ⋯, map_one' := ⋯, map_mul' := ⋯ }
Instances For
Equations
- NonUnitalNonAssocSemiring.toDistribSMul = DistribSMul.mk ⋯