Scalar multiplication on Finsupp
#
This file defines the pointwise scalar multiplication on Finsupp
, assuming it preserves zero.
Main declarations #
Finsupp.smulZeroClass
: if the action ofR
onM
preserves0
, then it acts onα →₀ M
Implementation notes #
This file is intermediate between Finsupp.Defs
and Finsupp.Module
in that it covers scalar
multiplication but does not rely on the definition of Module
. Scalar multiplication is needed to
supply the nsmul
(and zsmul
) fields of (semi)ring structures which are fundamental for e.g.
Polynomial
, so we want to keep the imports requied for the Finsupp.smulZeroClass
instance
reasonably light.
This file is a noncomputable theory
and uses classical logic throughout.
instance
Finsupp.smulZeroClass
{α : Type u_1}
{M : Type u_5}
{R : Type u_11}
[Zero M]
[SMulZeroClass R M]
:
SMulZeroClass R (α →₀ M)
Equations
- Finsupp.smulZeroClass = SMulZeroClass.mk ⋯
Throughout this section, some Monoid
and Semiring
arguments are specified with {}
instead of
[]
. See note [implicit instance arguments].
instance
Finsupp.instSMulWithZero
{α : Type u_1}
{M : Type u_5}
{R : Type u_11}
[Zero R]
[Zero M]
[SMulWithZero R M]
:
SMulWithZero R (α →₀ M)
Equations
- Finsupp.instSMulWithZero = SMulWithZero.mk ⋯
instance
Finsupp.distribSMul
(α : Type u_1)
(M : Type u_5)
{R : Type u_11}
[AddZeroClass M]
[DistribSMul R M]
:
DistribSMul R (α →₀ M)
Equations
instance
Finsupp.isScalarTower
(α : Type u_1)
(M : Type u_5)
{R : Type u_11}
{S : Type u_12}
[Zero M]
[SMulZeroClass R M]
[SMulZeroClass S M]
[SMul R S]
[IsScalarTower R S M]
:
IsScalarTower R S (α →₀ M)
instance
Finsupp.smulCommClass
(α : Type u_1)
(M : Type u_5)
{R : Type u_11}
{S : Type u_12}
[Zero M]
[SMulZeroClass R M]
[SMulZeroClass S M]
[SMulCommClass R S M]
:
SMulCommClass R S (α →₀ M)
instance
Finsupp.isCentralScalar
(α : Type u_1)
(M : Type u_5)
{R : Type u_11}
[Zero M]
[SMulZeroClass R M]
[SMulZeroClass Rᵐᵒᵖ M]
[IsCentralScalar R M]
:
IsCentralScalar R (α →₀ M)
theorem
Finsupp.support_smul
{α : Type u_1}
{M : Type u_5}
{R : Type u_11}
[AddMonoid M]
[SMulZeroClass R M]
{b : R}
{g : α →₀ M}
:
@[simp]
theorem
Finsupp.smul_single
{α : Type u_1}
{M : Type u_5}
{R : Type u_11}
[Zero M]
[SMulZeroClass R M]
(c : R)
(a : α)
(b : M)
:
c • Finsupp.single a b = Finsupp.single a (c • b)
theorem
Finsupp.mapRange_smul
{α : Type u_1}
{M : Type u_5}
{N : Type u_7}
{R : Type u_11}
{x✝ : Monoid R}
[AddMonoid M]
[DistribMulAction R M]
[AddMonoid N]
[DistribMulAction R N]
{f : M → N}
{hf : f 0 = 0}
(c : R)
(v : α →₀ M)
(hsmul : ∀ (x : M), f (c • x) = c • f x)
:
Finsupp.mapRange f hf (c • v) = c • Finsupp.mapRange f hf v