Actions by matrices on vectors through *ᵥ
and ᵥ*
, cast as Module
s #
This file provides the left- and right- module structures of square matrices on vectors, via
Matrix.mulVec
and Matrix.vecMul
.
*ᵥ
as a left-module #
instance
Matrix.instModuleForall
{n : Type u_1}
{R : Type u_2}
[Fintype n]
[DecidableEq n]
[Semiring R]
:
Equations
- Matrix.instModuleForall = { smul := Matrix.mulVec, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, smul_add := ⋯, add_smul := ⋯, zero_smul := ⋯ }
instance
Matrix.instSMulCommClassForall
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[SMulCommClass R S R]
:
SMulCommClass (Matrix n n R) S (n → R)
instance
Matrix.instSMulCommClassForall_1
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[SMulCommClass S R R]
:
SMulCommClass S (Matrix n n R) (n → R)
instance
Matrix.instIsScalarTowerForall
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[IsScalarTower S R R]
:
IsScalarTower S (Matrix n n R) (n → R)
*ᵥ
as a right-module #
instance
Matrix.instModuleMulOppositeForall
{n : Type u_1}
{R : Type u_2}
[Fintype n]
[DecidableEq n]
[Semiring R]
:
Equations
- One or more equations did not get rendered due to their size.
@[simp]
theorem
Matrix.op_smul_eq_vecMul
{n : Type u_1}
{R : Type u_2}
[Fintype n]
[DecidableEq n]
[Semiring R]
(A : (Matrix n n R)ᵐᵒᵖ)
(v : n → R)
:
instance
Matrix.instSMulCommClassMulOppositeForallOfIsScalarTower
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[IsScalarTower S R R]
:
SMulCommClass (Matrix n n R)ᵐᵒᵖ S (n → R)
instance
Matrix.instSMulCommClassMulOppositeForallOfIsScalarTower_1
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[IsScalarTower S R R]
:
SMulCommClass S (Matrix n n R)ᵐᵒᵖ (n → R)
instance
Matrix.instIsScalarTowerMulOppositeForallOfSMulCommClass
{n : Type u_1}
{R : Type u_2}
{S : Type u_3}
[Fintype n]
[DecidableEq n]
[Semiring R]
[DistribSMul S R]
[SMulCommClass S R R]
:
IsScalarTower S (Matrix n n R)ᵐᵒᵖ (n → R)