Trace of a matrix #
This file defines the trace of a matrix, the map sending a matrix to the sum of its diagonal entries.
See also LinearAlgebra.Trace
for the trace of an endomorphism.
Tags #
matrix, trace, diagonal
The trace of a square matrix. For more bundled versions, see:
Instances For
theorem
Matrix.trace_diagonal
{R : Type u_6}
[AddCommMonoid R]
{o : Type u_8}
[Fintype o]
[DecidableEq o]
(d : o → R)
:
@[simp]
@[simp]
theorem
Matrix.trace_eq_zero_of_isEmpty
{n : Type u_3}
{R : Type u_6}
[Fintype n]
[AddCommMonoid R]
[IsEmpty n]
(A : Matrix n n R)
:
@[simp]
theorem
Matrix.trace_smul
{n : Type u_3}
{α : Type u_5}
{R : Type u_6}
[Fintype n]
[AddCommMonoid R]
[Monoid α]
[DistribMulAction α R]
(r : α)
(A : Matrix n n R)
:
@[simp]
theorem
Matrix.trace_conjTranspose
{n : Type u_3}
{R : Type u_6}
[Fintype n]
[AddCommMonoid R]
[StarAddMonoid R]
(A : Matrix n n R)
:
Matrix.trace
as an AddMonoidHom
Equations
- Matrix.traceAddMonoidHom n R = { toFun := Matrix.trace, map_zero' := ⋯, map_add' := ⋯ }
Instances For
@[simp]
theorem
Matrix.traceAddMonoidHom_apply
(n : Type u_3)
(R : Type u_6)
[Fintype n]
[AddCommMonoid R]
(A : Matrix n n R)
:
def
Matrix.traceLinearMap
(n : Type u_3)
(α : Type u_5)
(R : Type u_6)
[Fintype n]
[AddCommMonoid R]
[Semiring α]
[Module α R]
:
Matrix.trace
as a LinearMap
Equations
- Matrix.traceLinearMap n α R = { toFun := Matrix.trace, map_add' := ⋯, map_smul' := ⋯ }
Instances For
@[simp]
theorem
Matrix.traceLinearMap_apply
(n : Type u_3)
(α : Type u_5)
(R : Type u_6)
[Fintype n]
[AddCommMonoid R]
[Semiring α]
[Module α R]
(A : Matrix n n R)
:
@[simp]
theorem
Matrix.trace_multiset_sum
{n : Type u_3}
{R : Type u_6}
[Fintype n]
[AddCommMonoid R]
(s : Multiset (Matrix n n R))
:
theorem
AddMonoidHom.map_trace
{n : Type u_3}
{R : Type u_6}
{S : Type u_7}
[Fintype n]
[AddCommMonoid R]
[AddCommMonoid S]
{F : Type u_8}
[FunLike F R S]
[AddMonoidHomClass F R S]
(f : F)
(A : Matrix n n R)
:
theorem
Matrix.trace_blockDiagonal
{n : Type u_3}
{p : Type u_4}
{R : Type u_6}
[Fintype n]
[Fintype p]
[AddCommMonoid R]
[DecidableEq p]
(M : p → Matrix n n R)
:
theorem
Matrix.trace_blockDiagonal'
{p : Type u_4}
{R : Type u_6}
[Fintype p]
[AddCommMonoid R]
[DecidableEq p]
{m : p → Type u_8}
[(i : p) → Fintype (m i)]
(M : (i : p) → Matrix (m i) (m i) R)
:
@[simp]
theorem
Matrix.trace_one
{n : Type u_3}
{R : Type u_6}
[Fintype n]
[DecidableEq n]
[AddCommMonoidWithOne R]
:
@[simp]
@[simp]
@[simp]
@[simp]
@[simp]
theorem
Matrix.StdBasisMatrix.trace_zero
{n : Type u_10}
{α : Type u_12}
[DecidableEq n]
[Fintype n]
[AddCommMonoid α]
(i j : n)
(c : α)
(h : j ≠ i)
:
@[simp]
theorem
Matrix.StdBasisMatrix.trace_eq
{n : Type u_10}
{α : Type u_12}
[DecidableEq n]
[Fintype n]
[AddCommMonoid α]
(i : n)
(c : α)
: