Matrices #
This file defines basic properties of matrices up to the module structure.
Matrices with rows indexed by m
, columns indexed by n
, and entries of type α
are represented
with Matrix m n α
. For the typical approach of counting rows and columns,
Matrix (Fin m) (Fin n) α
can be used.
Main definitions #
Matrix.transpose
: transpose of a matrix, turning rows into columns and vice versaMatrix.submatrix
: take a submatrix by reindexing rows and columnsMatrix.module
: matrices are a module over the ring of entries
Notation #
The locale Matrix
gives the following notation:
ᵀ
forMatrix.transpose
See Mathlib/Data/Matrix/ConjTranspose.lean
for
ᴴ
forMatrix.conjTranspose
Implementation notes #
For convenience, Matrix m n α
is defined as m → n → α
, as this allows elements of the matrix
to be accessed with A i j
. However, it is not advisable to construct matrices using terms of the
form fun i j ↦ _
or even (fun i j ↦ _ : Matrix m n α)
, as these are not recognized by Lean
as having the right type. Instead, Matrix.of
should be used.
Cast a function into a matrix.
The two sides of the equivalence are definitionally equal types. We want to use an explicit cast
to distinguish the types because Matrix
has different instances to pi types (such as Pi.mul
,
which performs elementwise multiplication, vs Matrix.mul
).
If you are defining a matrix, in terms of its entries, use of (fun i j ↦ _)
. The
purpose of this approach is to ensure that terms of the form (fun i j ↦ _) * (fun i j ↦ _)
do not
appear, as the type of *
can be misleading.
Porting note: In Lean 3, it is also safe to use pattern matching in a definition as | i j := _
,
which can only be unfolded when fully-applied. https://github.com/leanprover/lean4/issues/2042 means this does not
(currently) work in Lean 4.
Equations
- Matrix.of = Equiv.refl (m → n → α)
Instances For
M.map f
is the matrix obtained by applying f
to each entry of the matrix M
.
This is available in bundled forms as:
AddMonoidHom.mapMatrix
LinearMap.mapMatrix
RingHom.mapMatrix
AlgHom.mapMatrix
Equiv.mapMatrix
AddEquiv.mapMatrix
LinearEquiv.mapMatrix
RingEquiv.mapMatrix
AlgEquiv.mapMatrix
Equations
- M.map f = Matrix.of fun (i : m) (j : n) => f (M i j)
Instances For
The transpose of a matrix.
Equations
- Matrix.«term_ᵀ» = Lean.ParserDescr.trailingNode `Matrix.«term_ᵀ» 1024 1024 (Lean.ParserDescr.symbol "ᵀ")
Instances For
Equations
- Matrix.addSemigroup = Pi.addSemigroup
Equations
- Matrix.addCommSemigroup = Pi.addCommSemigroup
Equations
- Matrix.addZeroClass = Pi.addZeroClass
Equations
- Matrix.addCommMonoid = Pi.addCommMonoid
Equations
- Matrix.addCommGroup = Pi.addCommGroup
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- Matrix.distribMulAction = Pi.distribMulAction R
simp-normal form pulls of
to the outside.
The scalar action via mul.toOppositeSMul
is transformed by the same map as the
elements of the matrix, when f
preserves multiplication.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Given maps (r_reindex : l → m)
and (c_reindex : o → n)
reindexing the rows and columns of
a matrix M : Matrix m n α
, the matrix M.submatrix r_reindex c_reindex : Matrix l o α
is defined
by (M.submatrix r_reindex c_reindex) i j = M (r_reindex i) (c_reindex j)
for (i,j) : l × o
.
Note that the total number of row and columns does not have to be preserved.
Equations
- A.submatrix r_reindex c_reindex = Matrix.of fun (i : l) (j : o) => A (r_reindex i) (c_reindex j)
Instances For
The natural map that reindexes a matrix's rows and columns with equivalent types is an equivalence.
Equations
- Matrix.reindex eₘ eₙ = { toFun := fun (M : Matrix m n α) => M.submatrix ⇑eₘ.symm ⇑eₙ.symm, invFun := fun (M : Matrix l o α) => M.submatrix ⇑eₘ ⇑eₙ, left_inv := ⋯, right_inv := ⋯ }