Further results on (semi)linear maps #
def
LinearMap.ltoFun
(R : Type u_6)
(M : Type u_7)
(N : Type u_8)
(A : Type u_9)
[Semiring R]
[Semiring A]
[AddCommMonoid M]
[Module R M]
[AddCommMonoid N]
[Module R N]
[Module A N]
[SMulCommClass R A N]
:
A-linearly coerce an R-linear map from M to N to a function, when N has
commuting R-module and A-module structures.
Equations
Instances For
@[simp]
theorem
LinearMap.ltoFun_apply
{R : Type u_6}
{M : Type u_7}
{N : Type u_8}
{A : Type u_9}
[Semiring R]
[Semiring A]
[AddCommMonoid M]
[Module R M]
[AddCommMonoid N]
[Module R N]
[Module A N]
[SMulCommClass R A N]
{f : M →ₗ[R] N}
:
instance
LinearMap.instSMulDomMulAct
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
:
theorem
DomMulAct.smul_linearMap_apply
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
(a : S'ᵈᵐᵃ)
(f : M →ₛₗ[σ₁₂] M')
(x : M)
:
@[simp]
theorem
DomMulAct.mk_smul_linearMap_apply
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
(a : S')
(f : M →ₛₗ[σ₁₂] M')
(x : M)
:
theorem
DomMulAct.coe_smul_linearMap
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
(a : S'ᵈᵐᵃ)
(f : M →ₛₗ[σ₁₂] M')
:
instance
LinearMap.instSMulCommClassDomMulAct
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
{T' : Type u_7}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
[Monoid T']
[DistribMulAction T' M]
[SMulCommClass R T' M]
[SMulCommClass S' T' M]
:
SMulCommClass S'ᵈᵐᵃ T'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M')
instance
LinearMap.instDistribMulActionDomMulActOfSMulCommClass
{R : Type u_1}
{R' : Type u_2}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
{S' : Type u_6}
[Monoid S']
[DistribMulAction S' M]
[SMulCommClass R S' M]
:
DistribMulAction S'ᵈᵐᵃ (M →ₛₗ[σ₁₂] M')
Equations
- LinearMap.instDistribMulActionDomMulActOfSMulCommClass = { toSMul := LinearMap.instSMulDomMulAct, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, smul_add := ⋯ }
instance
LinearMap.instNoZeroSMulDivisors
{R : Type u_1}
{R' : Type u_2}
{S : Type u_3}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
[Semiring S]
[Module S M']
[SMulCommClass R' S M']
[NoZeroSMulDivisors S M']
:
NoZeroSMulDivisors S (M →ₛₗ[σ₁₂] M')
instance
LinearMap.instModuleDomMulActOfSMulCommClass
{R : Type u_1}
{R' : Type u_2}
{S : Type u_3}
{M : Type u_4}
{M' : Type u_5}
[Semiring R]
[Semiring R']
[AddCommMonoid M]
[AddCommMonoid M']
[Module R M]
[Module R' M']
{σ₁₂ : R →+* R'}
[Semiring S]
[Module S M]
[SMulCommClass R S M]
:
Equations
- LinearMap.instModuleDomMulActOfSMulCommClass = { toDistribMulAction := LinearMap.instDistribMulActionDomMulActOfSMulCommClass, add_smul := ⋯, zero_smul := ⋯ }