mathlib documentation

linear_algebra.matrix

Linear maps and matrices

This file defines the maps to send matrices to a linear map, and to send linear maps between modules with a finite bases to matrices. This defines a linear equivalence between linear maps between finite-dimensional vector spaces and matrices indexed by the respective bases.

It also defines the trace of an endomorphism, and the determinant of a family of vectors with respect to some basis.

Some results are proved about the linear map corresponding to a diagonal matrix (range, ker and rank).

Main definitions

In the list below, and in all this file, R is a commutative ring (semiring is sometimes enough), M and its variations are R-modules, ι, κ, n and m are finite types used for indexing.

Tags

linear_map, matrix, linear_equiv, diagonal, det, trace

@[instance]
def matrix.fintype {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] [decidable_eq m] [decidable_eq n] (R : Type u_1) [fintype R] :
fintype (matrix m n R)

Equations
def matrix.eval {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] :
matrix m n R →ₗ[R] (n → R) →ₗ[R] m → R

Evaluation of matrices gives a linear map from matrix m n R to linear maps (n → R) →ₗ[R] (m → R).

Equations
def matrix.to_lin {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] (a : matrix m n R) :
(n → R) →ₗ[R] m → R

Evaluation of matrices gives a map from matrix m n R to linear maps (n → R) →ₗ[R] (m → R).

Equations
theorem matrix.to_lin_of_equiv {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] {p : Type u_1} {q : Type u_4} [fintype p] [fintype q] (e₁ : m p) (e₂ : n q) (f : matrix p q R) :
matrix.to_lin (λ (i : m) (j : n), f (e₁ i) (e₂ j)) = ((linear_map.fun_congr_left R R e₂).arrow_congr (linear_map.fun_congr_left R R e₁)) f.to_lin

theorem matrix.to_lin_add {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] (M N : matrix m n R) :
(M + N).to_lin = M.to_lin + N.to_lin

@[simp]
theorem matrix.to_lin_zero {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] :
0.to_lin = 0

@[simp]
theorem matrix.to_lin_neg {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] (M : matrix m n R) :

@[instance]
def matrix.to_lin.is_add_monoid_hom {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] :

@[simp]
theorem matrix.to_lin_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] (M : matrix m n R) (v : n → R) :
(M.to_lin) v = M.mul_vec v

theorem matrix.mul_to_lin {l : Type u_1} {m : Type u_2} {n : Type u_3} [fintype l] [fintype m] [fintype n] {R : Type v} [comm_ring R] (M : matrix m n R) (N : matrix n l R) :

@[simp]
theorem matrix.to_lin_one {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] :

def linear_map.to_matrixₗ {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] :
((n → R) →ₗ[R] m → R) →ₗ[R] matrix m n R

The linear map from linear maps (n → R) →ₗ[R] (m → R) to matrix m n R.

Equations
def linear_map.to_matrix {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (a : (n → R) →ₗ[R] m → R) :
matrix m n R

The map from linear maps (n → R) →ₗ[R] (m → R) to matrix m n R.

Equations
@[simp]
theorem linear_map.to_matrix_id {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] :

theorem linear_map.to_matrix_of_equiv {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] {p : Type u_1} {q : Type u_4} [decidable_eq n] [decidable_eq q] [fintype p] [fintype q] (e₁ : m p) (e₂ : n q) (f : (q → R) →ₗ[R] p → R) (i : m) (j : n) :

theorem to_matrix_to_lin {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] {f : (n → R) →ₗ[R] m → R} :

to_lin is the left inverse of to_matrix.

theorem to_lin_to_matrix {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] {M : matrix m n R} :

to_lin is the right inverse of to_matrix.

def linear_equiv_matrix' {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] :
((n → R) →ₗ[R] m → R) ≃ₗ[R] matrix m n R

Linear maps (n → R) →ₗ[R] (m → R) are linearly equivalent to matrix m n R.

Equations
@[simp]
theorem linear_equiv_matrix'_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (f : (n → R) →ₗ[R] m → R) :

@[simp]
theorem linear_equiv_matrix'_symm_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (M : matrix m n R) :

def linear_equiv_matrix {R : Type v} [comm_ring R] {ι : Type u_4} {κ : Type u_5} {M₁ : Type u_6} {M₂ : Type u_7} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [fintype ι] [decidable_eq ι] [fintype κ] {v₁ : ι → M₁} {v₂ : κ → M₂} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) :
(M₁ →ₗ[R] M₂) ≃ₗ[R] matrix κ ι R

Given bases of two modules M₁ and M₂ over a commutative ring R, we get a linear equivalence between linear maps M₁ →ₗ M₂ and matrices over R indexed by the bases.

Equations
theorem linear_equiv_matrix_apply {R : Type v} [comm_ring R] {ι : Type u_4} {κ : Type u_5} {M₁ : Type u_6} {M₂ : Type u_7} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [fintype ι] [decidable_eq ι] [fintype κ] {v₁ : ι → M₁} {v₂ : κ → M₂} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (f : M₁ →ₗ[R] M₂) (i : κ) (j : ι) :
(linear_equiv_matrix hv₁ hv₂) f i j = (hv₂.equiv_fun) (f (v₁ j)) i

theorem linear_equiv_matrix_symm_apply {R : Type v} [comm_ring R] {ι : Type u_4} {κ : Type u_5} {M₁ : Type u_6} {M₂ : Type u_7} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [fintype ι] [decidable_eq ι] [fintype κ] {v₁ : ι → M₁} {v₂ : κ → M₂} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (m : matrix κ ι R) (x : M₁) :
(((linear_equiv_matrix hv₁ hv₂).symm) m) x = (hv₂.equiv_fun.symm) ((m.to_lin) ((hv₁.equiv_fun) x))

theorem linear_equiv_matrix_apply' {R : Type v} [comm_ring R] {ι : Type u_4} {κ : Type u_5} {M₁ : Type u_6} {M₂ : Type u_7} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [fintype ι] [decidable_eq ι] [fintype κ] {v₁ : ι → M₁} {v₂ : κ → M₂} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (f : M₁ →ₗ[R] M₂) (i : κ) (j : ι) :
(linear_equiv_matrix hv₁ hv₂) f i j = ((hv₂.repr) (f (v₁ j))) i

@[simp]
theorem linear_equiv_matrix_id {R : Type v} [comm_ring R] {ι : Type u_4} {M₁ : Type u_6} [add_comm_group M₁] [module R M₁] [fintype ι] [decidable_eq ι] {v₁ : ι → M₁} (hv₁ : is_basis R v₁) :

@[simp]
theorem linear_equiv_matrix_symm_one {R : Type v} [comm_ring R] {ι : Type u_4} {M₁ : Type u_6} [add_comm_group M₁] [module R M₁] [fintype ι] [decidable_eq ι] {v₁ : ι → M₁} (hv₁ : is_basis R v₁) :

theorem linear_equiv_matrix_range {R : Type v} [comm_ring R] {ι : Type u_4} {κ : Type u_5} {M₁ : Type u_6} {M₂ : Type u_7} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [fintype ι] [decidable_eq ι] [fintype κ] {v₁ : ι → M₁} {v₂ : κ → M₂} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (f : M₁ →ₗ[R] M₂) (k : κ) (i : ι) :
(linear_equiv_matrix _ _) f v₂ k, _⟩ v₁ i, _⟩ = (linear_equiv_matrix hv₁ hv₂) f k i

theorem matrix.comp_to_matrix_mul {l : Type u_1} {m : Type u_2} {n : Type u_3} [fintype l] [fintype m] [fintype n] {R : Type v} [comm_ring R] [decidable_eq l] [decidable_eq m] (f : (m → R) →ₗ[R] n → R) (g : (l → R) →ₗ[R] m → R) :

theorem matrix.linear_equiv_matrix_comp {R : Type u_4} {ι : Type u_5} {κ : Type u_6} {μ : Type u_7} {M₁ : Type u_8} {M₂ : Type u_9} {M₃ : Type u_10} [comm_ring R] [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [add_comm_group M₃] [module R M₃] [fintype ι] [decidable_eq κ] [fintype κ] [fintype μ] {v₁ : ι → M₁} {v₂ : κ → M₂} {v₃ : μ → M₃} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (hv₃ : is_basis R v₃) [decidable_eq ι] (f : M₂ →ₗ[R] M₃) (g : M₁ →ₗ[R] M₂) :
(linear_equiv_matrix hv₁ hv₃) (f.comp g) = (linear_equiv_matrix hv₂ hv₃) f (linear_equiv_matrix hv₁ hv₂) g

theorem matrix.linear_equiv_matrix_mul {R : Type u_4} {ι : Type u_5} {M₁ : Type u_8} [comm_ring R] [add_comm_group M₁] [module R M₁] [fintype ι] {v₁ : ι → M₁} (hv₁ : is_basis R v₁) [decidable_eq ι] (f g : M₁ →ₗ[R] M₁) :
(linear_equiv_matrix hv₁ hv₁) (f * g) = ((linear_equiv_matrix hv₁ hv₁) f) * (linear_equiv_matrix hv₁ hv₁) g

theorem matrix.linear_equiv_matrix_symm_mul {R : Type u_4} {ι : Type u_5} {κ : Type u_6} {μ : Type u_7} {M₁ : Type u_8} {M₂ : Type u_9} {M₃ : Type u_10} [comm_ring R] [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] [add_comm_group M₃] [module R M₃] [fintype ι] [decidable_eq κ] [fintype κ] [fintype μ] {v₁ : ι → M₁} {v₂ : κ → M₂} {v₃ : μ → M₃} (hv₁ : is_basis R v₁) (hv₂ : is_basis R v₂) (hv₃ : is_basis R v₃) [decidable_eq μ] (A : matrix ι κ R) (B : matrix κ μ R) :
((linear_equiv_matrix hv₃ hv₁).symm) (A B) = (((linear_equiv_matrix hv₂ hv₁).symm) A).comp (((linear_equiv_matrix hv₃ hv₂).symm) B)

def is_basis.to_matrix {ι : Type u_4} {R : Type u_6} {M : Type u_7} [fintype ι] [decidable_eq ι] [comm_ring R] [add_comm_group M] [module R M] {e : ι → M} (he : is_basis R e) (v : ι → M) :
matrix ι ι R

From a basis e : ι → M and a family of vectors v : ι → M, make the matrix whose columns are the vectors v i written in the basis e.

Equations
theorem is_basis.to_matrix_apply {ι : Type u_4} {R : Type u_6} {M : Type u_7} [fintype ι] [decidable_eq ι] [comm_ring R] [add_comm_group M] [module R M] {e : ι → M} (he : is_basis R e) (v : ι → M) (i j : ι) :
he.to_matrix v i j = (he.equiv_fun) (v j) i

@[simp]
theorem is_basis.to_matrix_self {ι : Type u_4} {R : Type u_6} {M : Type u_7} [fintype ι] [decidable_eq ι] [comm_ring R] [add_comm_group M] [module R M] {e : ι → M} (he : is_basis R e) :
he.to_matrix e = 1

theorem is_basis.to_matrix_update {ι : Type u_4} {R : Type u_6} {M : Type u_7} [fintype ι] [decidable_eq ι] [comm_ring R] [add_comm_group M] [module R M] {e : ι → M} (he : is_basis R e) (v : ι → M) (i : ι) (x : M) :

def is_basis.to_matrix_equiv {ι : Type u_4} {R : Type u_6} {M : Type u_7} [fintype ι] [decidable_eq ι] [comm_ring R] [add_comm_group M] [module R M] {e : ι → M} (he : is_basis R e) :
(ι → M) ≃ₗ[R] matrix ι ι R

From a basis e : ι → M, build a linear equivalence between families of vectors v : ι → M, and matrices, making the matrix whose columns are the vectors v i written in the basis e.

Equations
theorem linear_equiv.is_unit_det {R : Type u_4} {ι : Type u_5} {M : Type u_6} {M' : Type u_7} [comm_ring R] [add_comm_group M] [module R M] [add_comm_group M'] [module R M'] [decidable_eq ι] [fintype ι] {v : ι → M} {v' : ι → M'} (f : M ≃ₗ[R] M') (hv : is_basis R v) (hv' : is_basis R v') :

def linear_equiv.of_is_unit_det {R : Type u_4} {ι : Type u_5} {M : Type u_6} {M' : Type u_7} [comm_ring R] [add_comm_group M] [module R M] [add_comm_group M'] [module R M'] [decidable_eq ι] [fintype ι] {v : ι → M} {v' : ι → M'} {f : M →ₗ[R] M'} {hv : is_basis R v} {hv' : is_basis R v'} (h : is_unit ((linear_equiv_matrix hv hv') f).det) :
M ≃ₗ[R] M'

Builds a linear equivalence from a linear map whose determinant in some bases is a unit.

Equations
def is_basis.det {R : Type u_4} {ι : Type u_5} {M : Type u_6} [comm_ring R] [add_comm_group M] [module R M] [decidable_eq ι] [fintype ι] {e : ι → M} (he : is_basis R e) :
multilinear_map R (λ (i : ι), M) R

The determinant of a family of vectors with respect to some basis, as a multilinear map.

Equations
theorem is_basis.det_apply {R : Type u_4} {ι : Type u_5} {M : Type u_6} [comm_ring R] [add_comm_group M] [module R M] [decidable_eq ι] [fintype ι] {e : ι → M} (he : is_basis R e) (v : ι → M) :
(he.det) v = (he.to_matrix v).det

theorem is_basis.det_self {R : Type u_4} {ι : Type u_5} {M : Type u_6} [comm_ring R] [add_comm_group M] [module R M] [decidable_eq ι] [fintype ι] {e : ι → M} (he : is_basis R e) :
(he.det) e = 1

theorem is_basis.iff_det {R : Type u_4} {ι : Type u_5} {M : Type u_6} [comm_ring R] [add_comm_group M] [module R M] [decidable_eq ι] [fintype ι] {e : ι → M} (he : is_basis R e) {v : ι → M} :
is_basis R v is_unit ((he.det) v)

@[simp]
theorem linear_equiv_matrix_transpose {K : Type u_4} {V₁ : Type u_5} {V₂ : Type u_6} {ι₁ : Type u_7} {ι₂ : Type u_8} [field K] [add_comm_group V₁] [vector_space K V₁] [add_comm_group V₂] [vector_space K V₂] [fintype ι₁] [fintype ι₂] [decidable_eq ι₁] [decidable_eq ι₂] {B₁ : ι₁ → V₁} (h₁ : is_basis K B₁) {B₂ : ι₂ → V₂} (h₂ : is_basis K B₂) (u : V₁ →ₗ[K] V₂) :

theorem linear_equiv_matrix_symm_transpose {K : Type u_4} {V₁ : Type u_5} {V₂ : Type u_6} {ι₁ : Type u_7} {ι₂ : Type u_8} [field K] [add_comm_group V₁] [vector_space K V₁] [add_comm_group V₂] [vector_space K V₂] [fintype ι₁] [fintype ι₂] [decidable_eq ι₁] [decidable_eq ι₂] {B₁ : ι₁ → V₁} (h₁ : is_basis K B₁) {B₂ : ι₂ → V₂} (h₂ : is_basis K B₂) (M : matrix ι₁ ι₂ K) :

def matrix.diag (n : Type u_3) [fintype n] (R : Type v) (M : Type w) [semiring R] [add_comm_monoid M] [semimodule R M] :
matrix n n M →ₗ[R] n → M

The diagonal of a square matrix.

Equations
@[simp]
theorem matrix.diag_apply {n : Type u_3} [fintype n] {R : Type v} {M : Type w} [semiring R] [add_comm_monoid M] [semimodule R M] (A : matrix n n M) (i : n) :
(matrix.diag n R M) A i = A i i

@[simp]
theorem matrix.diag_one {n : Type u_3} [fintype n] {R : Type v} [semiring R] [decidable_eq n] :
(matrix.diag n R R) 1 = λ (i : n), 1

@[simp]
theorem matrix.diag_transpose {n : Type u_3} [fintype n] {R : Type v} {M : Type w} [semiring R] [add_comm_monoid M] [semimodule R M] (A : matrix n n M) :
(matrix.diag n R M) A = (matrix.diag n R M) A

def matrix.trace (n : Type u_3) [fintype n] (R : Type v) (M : Type w) [semiring R] [add_comm_monoid M] [semimodule R M] :
matrix n n M →ₗ[R] M

The trace of a square matrix.

Equations
@[simp]
theorem matrix.trace_diag {n : Type u_3} [fintype n] {R : Type v} {M : Type w} [semiring R] [add_comm_monoid M] [semimodule R M] (A : matrix n n M) :
(matrix.trace n R M) A = ∑ (i : n), (matrix.diag n R M) A i

@[simp]
theorem matrix.trace_one {n : Type u_3} [fintype n] {R : Type v} [semiring R] [decidable_eq n] :

@[simp]
theorem matrix.trace_transpose {n : Type u_3} [fintype n] {R : Type v} {M : Type w} [semiring R] [add_comm_monoid M] [semimodule R M] (A : matrix n n M) :

@[simp]
theorem matrix.trace_transpose_mul {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [semiring R] (A : matrix m n R) (B : matrix n m R) :
(matrix.trace n R R) (A B) = (matrix.trace m R R) (A B)

theorem matrix.trace_mul_comm {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {S : Type v} [comm_ring S] (A : matrix m n S) (B : matrix n m S) :
(matrix.trace n S S) (B A) = (matrix.trace m S S) (A B)

theorem matrix.proj_diagonal {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (i : n) (w : n → R) :

theorem matrix.diagonal_comp_std_basis {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (w : n → R) (i : n) :
(matrix.diagonal w).to_lin.comp (linear_map.std_basis R (λ (_x : n), R) i) = w i linear_map.std_basis R (λ (_x : n), R) i

theorem matrix.diagonal_to_lin {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (w : n → R) :

def matrix.to_linear_equiv {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (P : matrix n n R) (h : is_unit P) :
(n → R) ≃ₗ[R] n → R

An invertible matrix yields a linear equivalence from the free module to itself.

Equations
@[simp]
theorem matrix.to_linear_equiv_apply {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (P : matrix n n R) (h : is_unit P) :

@[simp]
theorem matrix.to_linear_equiv_symm_apply {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] (P : matrix n n R) (h : is_unit P) :

theorem matrix.rank_vec_mul_vec {K : Type u} [field K] {m n : Type u} [fintype m] [fintype n] (w : m → K) (v : n → K) :

theorem matrix.ker_diagonal_to_lin {m : Type u_2} [fintype m] {K : Type u} [field K] [decidable_eq m] (w : m → K) :
(matrix.diagonal w).to_lin.ker = ⨆ (i : m) (H : i {i : m | w i = 0}), (linear_map.std_basis K (λ (i : m), K) i).range

theorem matrix.range_diagonal {m : Type u_2} [fintype m] {K : Type u} [field K] [decidable_eq m] (w : m → K) :
(matrix.diagonal w).to_lin.range = ⨆ (i : m) (H : i {i : m | w i 0}), (linear_map.std_basis K (λ (i : m), K) i).range

theorem matrix.rank_diagonal {m : Type u_2} [fintype m] {K : Type u} [field K] [decidable_eq m] [decidable_eq K] (w : m → K) :

@[instance]
def matrix.finite_dimensional {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [field R] :

@[simp]
theorem matrix.findim_matrix {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [field R] :

The dimension of the space of finite dimensional matrices is the product of the number of rows and columns.

def matrix.reindex {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} (eₘ : m m') (eₙ : n n') :
matrix m n R matrix m' n' R

The natural map that reindexes a matrix's rows and columns with equivalent types is an equivalence.

Equations
@[simp]
theorem matrix.reindex_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} (eₘ : m m') (eₙ : n n') (M : matrix m n R) :
(matrix.reindex eₘ eₙ) M = λ (i : m') (j : n'), M ((eₘ.symm) i) ((eₙ.symm) j)

@[simp]
theorem matrix.reindex_symm_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} (eₘ : m m') (eₙ : n n') (M : matrix m' n' R) :
((matrix.reindex eₘ eₙ).symm) M = λ (i : m) (j : n), M (eₘ i) (eₙ j)

def matrix.reindex_linear_equiv {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} [semiring R] (eₘ : m m') (eₙ : n n') :
matrix m n R ≃ₗ[R] matrix m' n' R

The natural map that reindexes a matrix's rows and columns with equivalent types is a linear equivalence.

Equations
@[simp]
theorem matrix.reindex_linear_equiv_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} [semiring R] (eₘ : m m') (eₙ : n n') (M : matrix m n R) :
(matrix.reindex_linear_equiv eₘ eₙ) M = λ (i : m') (j : n'), M ((eₘ.symm) i) ((eₙ.symm) j)

@[simp]
theorem matrix.reindex_linear_equiv_symm_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} [semiring R] (eₘ : m m') (eₙ : n n') (M : matrix m' n' R) :
((matrix.reindex_linear_equiv eₘ eₙ).symm) M = λ (i : m) (j : n), M (eₘ i) (eₙ j)

theorem matrix.reindex_mul {l : Type u_1} {m : Type u_2} {n : Type u_3} [fintype l] [fintype m] [fintype n] {l' : Type u_4} {m' : Type u_5} {n' : Type u_6} [fintype l'] [fintype m'] [fintype n'] {R : Type v} [semiring R] (eₘ : m m') (eₙ : n n') (eₗ : l l') (M : matrix m n R) (N : matrix n l R) :

def matrix.reindex_alg_equiv {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_semiring R] [decidable_eq m] [decidable_eq n] (e : m n) :
matrix m m R ≃ₐ[R] matrix n n R

For square matrices, the natural map that reindexes a matrix's rows and columns with equivalent types is an equivalence of algebras.

Equations
@[simp]
theorem matrix.reindex_alg_equiv_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_semiring R] [decidable_eq m] [decidable_eq n] (e : m n) (M : matrix m m R) :
(matrix.reindex_alg_equiv e) M = λ (i j : n), M ((e.symm) i) ((e.symm) j)

@[simp]
theorem matrix.reindex_alg_equiv_symm_apply {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {R : Type v} [comm_semiring R] [decidable_eq m] [decidable_eq n] (e : m n) (M : matrix n n R) :
((matrix.reindex_alg_equiv e).symm) M = λ (i j : m), M (e i) (e j)

theorem matrix.reindex_transpose {m : Type u_2} {n : Type u_3} [fintype m] [fintype n] {m' : Type u_5} {n' : Type u_6} [fintype m'] [fintype n'] {R : Type v} (eₘ : m m') (eₙ : n n') (M : matrix m n R) :
((matrix.reindex eₘ eₙ) M) = (matrix.reindex eₙ eₘ) M

def linear_map.trace_aux (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [decidable_eq ι] [fintype ι] {b : ι → M} (hb : is_basis R b) :
(M →ₗ[R] M) →ₗ[R] R

The trace of an endomorphism given a basis.

Equations
@[simp]
theorem linear_map.trace_aux_def (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [decidable_eq ι] [fintype ι] {b : ι → M} (hb : is_basis R b) (f : M →ₗ[R] M) :

theorem linear_map.trace_aux_eq' (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [decidable_eq ι] [fintype ι] {b : ι → M} (hb : is_basis R b) {κ : Type w} [decidable_eq κ] [fintype κ] {c : κ → M} (hc : is_basis R c) :

theorem linear_map.trace_aux_range (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [fintype ι] {b : ι → M} (hb : is_basis R b) :

theorem linear_map.trace_aux_eq (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type u_1} [decidable_eq ι] [fintype ι] {b : ι → M} (hb : is_basis R b) {κ : Type u_2} [decidable_eq κ] [fintype κ] {c : κ → M} (hc : is_basis R c) :

where ι and κ can reside in different universes

def linear_map.trace (R : Type u) [comm_ring R] (M : Type v) [add_comm_group M] [module R M] :
(M →ₗ[R] M) →ₗ[R] R

Trace of an endomorphism independent of basis.

Equations
theorem linear_map.trace_eq_matrix_trace (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] {ι : Type w} [fintype ι] {b : ι → M} (hb : is_basis R b) (f : M →ₗ[R] M) :

theorem linear_map.trace_mul_comm (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] (f g : M →ₗ[R] M) :
(linear_map.trace R M) (f * g) = (linear_map.trace R M) (g * f)

@[instance]
def linear_map.finite_dimensional {K : Type u_4} [field K] {V : Type u_5} [add_comm_group V] [vector_space K V] [finite_dimensional K V] {W : Type u_6} [add_comm_group W] [vector_space K W] [finite_dimensional K W] :

@[simp]

The dimension of the space of linear transformations is the product of the dimensions of the domain and codomain.

def alg_equiv_matrix' {n : Type u_3} [fintype n] {R : Type v} [comm_ring R] [decidable_eq n] :
module.End R (n → R) ≃ₐ[R] matrix n n R

The natural equivalence between linear endomorphisms of finite free modules and square matrices is compatible with the algebra structures.

Equations
def linear_equiv.alg_conj {R : Type v} [comm_ring R] {M₁ : Type u_1} {M₂ : Type (max u_2 u_3)} [add_comm_group M₁] [module R M₁] [add_comm_group M₂] [module R M₂] (e : M₁ ≃ₗ[R] M₂) :

A linear equivalence of two modules induces an equivalence of algebras of their endomorphisms.

Equations
def alg_equiv_matrix {n : Type u_3} [fintype n] {R : Type v} {M : Type w} [comm_ring R] [add_comm_group M] [module R M] [decidable_eq n] {b : n → M} (h : is_basis R b) :

A basis of a module induces an equivalence of algebras from the endomorphisms of the module to square matrices.

Equations