Documentation

Mathlib.LinearAlgebra.Matrix.Trace

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

def Matrix.trace {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommMonoid R] (A : Matrix n n R) :
R

The trace of a square matrix. For more bundled versions, see:

Equations
    Instances For
      @[simp]
      theorem Matrix.trace_diagonal {R : Type u_6} [AddCommMonoid R] {o : Type u_8} [Fintype o] [DecidableEq o] (d : oR) :
      (diagonal d).trace = i : o, d i
      @[simp]
      theorem Matrix.trace_zero (n : Type u_3) (R : Type u_6) [Fintype n] [AddCommMonoid R] :
      trace 0 = 0
      @[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) :
      A.trace = 0
      @[simp]
      theorem Matrix.trace_add {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommMonoid R] (A B : Matrix n n R) :
      (A + B).trace = A.trace + B.trace
      @[simp]
      theorem Matrix.trace_smul {n : Type u_3} {α : Type u_5} {R : Type u_6} [Fintype n] [AddCommMonoid R] [DistribSMul α R] (r : α) (A : Matrix n n R) :
      (r A).trace = r A.trace
      @[simp]
      theorem Matrix.trace_transpose {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommMonoid 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) :
      def Matrix.traceAddMonoidHom (n : Type u_3) (R : Type u_6) [Fintype n] [AddCommMonoid R] :
      Matrix n n R →+ R

      Matrix.trace as an AddMonoidHom

      Equations
        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 n n R →ₗ[α] R

          Matrix.trace as a LinearMap

          Equations
            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) :
              (traceLinearMap n α R) A = A.trace
              @[simp]
              theorem Matrix.trace_list_sum {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommMonoid R] (l : List (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)) :
              @[simp]
              theorem Matrix.trace_sum {ι : Type u_1} {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommMonoid R] (s : Finset ι) (f : ιMatrix n n R) :
              (∑ is, f i).trace = is, (f i).trace
              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) :
              f A.trace = ((↑f).mapMatrix A).trace
              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 : pMatrix n n R) :
              (blockDiagonal M).trace = i : p, (M i).trace
              theorem Matrix.trace_blockDiagonal' {p : Type u_4} {R : Type u_6} [Fintype p] [AddCommMonoid R] [DecidableEq p] {m : pType u_8} [(i : p) → Fintype (m i)] (M : (i : p) → Matrix (m i) (m i) R) :
              (blockDiagonal' M).trace = i : p, (M i).trace
              @[simp]
              theorem Matrix.trace_sub {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommGroup R] (A B : Matrix n n R) :
              (A - B).trace = A.trace - B.trace
              @[simp]
              theorem Matrix.trace_neg {n : Type u_3} {R : Type u_6} [Fintype n] [AddCommGroup R] (A : Matrix n n R) :
              @[simp]
              theorem Matrix.trace_one {n : Type u_3} {R : Type u_6} [Fintype n] [DecidableEq n] [AddCommMonoidWithOne R] :
              @[simp]
              theorem Matrix.trace_transpose_mul {m : Type u_2} {n : Type u_3} {R : Type u_6} [Fintype m] [Fintype n] [AddCommMonoid R] [Mul R] (A : Matrix m n R) (B : Matrix n m R) :
              theorem Matrix.trace_mul_comm {m : Type u_2} {n : Type u_3} {R : Type u_6} [Fintype m] [Fintype n] [AddCommMonoid R] [CommMagma R] (A : Matrix m n R) (B : Matrix n m R) :
              (A * B).trace = (B * A).trace
              theorem Matrix.trace_mul_cycle {m : Type u_2} {n : Type u_3} {p : Type u_4} {R : Type u_6} [Fintype m] [Fintype n] [Fintype p] [NonUnitalCommSemiring R] (A : Matrix m n R) (B : Matrix n p R) (C : Matrix p m R) :
              (A * B * C).trace = (C * A * B).trace
              theorem Matrix.trace_mul_cycle' {m : Type u_2} {n : Type u_3} {p : Type u_4} {R : Type u_6} [Fintype m] [Fintype n] [Fintype p] [NonUnitalCommSemiring R] (A : Matrix m n R) (B : Matrix n p R) (C : Matrix p m R) :
              (A * (B * C)).trace = (C * (A * B)).trace
              @[simp]
              theorem Matrix.trace_replicateCol_mul_replicateRow {n : Type u_3} {R : Type u_6} [Fintype n] {ι : Type u_8} [Unique ι] [NonUnitalNonAssocSemiring R] (a b : nR) :
              @[deprecated Matrix.trace_replicateCol_mul_replicateRow (since := "2025-03-20")]
              theorem Matrix.trace_col_mul_row {n : Type u_3} {R : Type u_6} [Fintype n] {ι : Type u_8} [Unique ι] [NonUnitalNonAssocSemiring R] (a b : nR) :

              Alias of Matrix.trace_replicateCol_mul_replicateRow.

              theorem Matrix.trace_units_conj {m : Type u_2} {R : Type u_6} [Fintype m] [DecidableEq m] [CommSemiring R] (M : (Matrix m m R)ˣ) (N : Matrix m m R) :
              (M * N * M⁻¹).trace = N.trace
              theorem Matrix.trace_units_conj' {m : Type u_2} {R : Type u_6} [Fintype m] [DecidableEq m] [CommSemiring R] (M : (Matrix m m R)ˣ) (N : Matrix m m R) :
              (M⁻¹ * N * M).trace = N.trace

              Special cases for Fin n for low values of n #

              @[simp]
              theorem Matrix.trace_fin_zero {R : Type u_6} [AddCommMonoid R] (A : Matrix (Fin 0) (Fin 0) R) :
              A.trace = 0
              theorem Matrix.trace_fin_one {R : Type u_6} [AddCommMonoid R] (A : Matrix (Fin 1) (Fin 1) R) :
              A.trace = A 0 0
              theorem Matrix.trace_fin_two {R : Type u_6} [AddCommMonoid R] (A : Matrix (Fin 2) (Fin 2) R) :
              A.trace = A 0 0 + A 1 1
              theorem Matrix.trace_fin_three {R : Type u_6} [AddCommMonoid R] (A : Matrix (Fin 3) (Fin 3) R) :
              A.trace = A 0 0 + A 1 1 + A 2 2
              @[simp]
              theorem Matrix.trace_fin_one_of {R : Type u_6} [AddCommMonoid R] (a : R) :
              !![a].trace = a
              @[simp]
              theorem Matrix.trace_fin_two_of {R : Type u_6} [AddCommMonoid R] (a b c d : R) :
              !![a, b; c, d].trace = a + d
              @[simp]
              theorem Matrix.trace_fin_three_of {R : Type u_6} [AddCommMonoid R] (a b c d e f g h i : R) :
              !![a, b, c; d, e, f; g, h, i].trace = a + e + i
              @[simp]
              theorem Matrix.trace_single_eq_of_ne {n : Type u_10} {α : Type u_12} [DecidableEq n] [Fintype n] [AddCommMonoid α] (i j : n) (c : α) (h : i j) :
              (single i j c).trace = 0
              @[deprecated Matrix.trace_single_eq_of_ne (since := "2025-05-05")]
              theorem Matrix.StdBasisMatrix.trace_zero {n : Type u_10} {α : Type u_12} [DecidableEq n] [Fintype n] [AddCommMonoid α] (i j : n) (c : α) (h : i j) :
              (single i j c).trace = 0

              Alias of Matrix.trace_single_eq_of_ne.

              @[simp]
              theorem Matrix.trace_single_eq_same {n : Type u_10} {α : Type u_12} [DecidableEq n] [Fintype n] [AddCommMonoid α] (i : n) (c : α) :
              (single i i c).trace = c
              @[deprecated Matrix.trace_single_eq_same (since := "2025-05-05")]
              theorem Matrix.StdBasisMatrix.trace_eq {n : Type u_10} {α : Type u_12} [DecidableEq n] [Fintype n] [AddCommMonoid α] (i : n) (c : α) :
              (single i i c).trace = c

              Alias of Matrix.trace_single_eq_same.