Documentation

Mathlib.LinearAlgebra.Matrix.Determinant

Determinant of a matrix #

This file defines the determinant of a matrix, Matrix.det, and its essential properties.

Main definitions #

Main results #

Implementation notes #

It is possible to configure simp to compute determinants. See the file test/matrix.lean for some examples.

def Matrix.detRowAlternating {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] :
(nR) [⋀^n]→ₗ[R] R

det is an AlternatingMap in the rows of the matrix.

Equations
Instances For
    @[reducible, inline]
    abbrev Matrix.det {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) :
    R

    The determinant of a matrix given by the Leibniz formula.

    Equations
    • M.det = Matrix.detRowAlternating M
    Instances For
      theorem Matrix.det_apply {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) :
      M.det = Finset.univ.sum fun (σ : Equiv.Perm n) => Equiv.Perm.sign σ Finset.univ.prod fun (i : n) => M (σ i) i
      theorem Matrix.det_apply' {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) :
      M.det = Finset.univ.sum fun (σ : Equiv.Perm n) => (Equiv.Perm.sign σ) * Finset.univ.prod fun (i : n) => M (σ i) i
      @[simp]
      theorem Matrix.det_diagonal {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {d : nR} :
      (Matrix.diagonal d).det = Finset.univ.prod fun (i : n) => d i
      theorem Matrix.det_zero {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] :
      @[simp]
      theorem Matrix.det_one {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] :
      theorem Matrix.det_isEmpty {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] [IsEmpty n] {A : Matrix n n R} :
      A.det = 1
      @[simp]
      theorem Matrix.coe_det_isEmpty {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] [IsEmpty n] :
      Matrix.det = Function.const (Matrix n n R) 1
      theorem Matrix.det_eq_one_of_card_eq_zero {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} (h : Fintype.card n = 0) :
      A.det = 1
      @[simp]
      theorem Matrix.det_unique {R : Type v} [CommRing R] {n : Type u_3} [Unique n] [DecidableEq n] [Fintype n] (A : Matrix n n R) :
      A.det = A default default

      If n has only one element, the determinant of an n by n matrix is just that element. Although Unique implies DecidableEq and Fintype, the instances might not be syntactically equal. Thus, we need to fill in the args explicitly.

      theorem Matrix.det_eq_elem_of_subsingleton {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] [Subsingleton n] (A : Matrix n n R) (k : n) :
      A.det = A k k
      theorem Matrix.det_eq_elem_of_card_eq_one {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} (h : Fintype.card n = 1) (k : n) :
      A.det = A k k
      theorem Matrix.det_mul_aux {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {M : Matrix n n R} {N : Matrix n n R} {p : nn} (H : ¬Function.Bijective p) :
      (Finset.univ.sum fun (σ : Equiv.Perm n) => (Equiv.Perm.sign σ) * Finset.univ.prod fun (x : n) => M (σ x) (p x) * N (p x) x) = 0
      @[simp]
      theorem Matrix.det_mul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (N : Matrix n n R) :
      (M * N).det = M.det * N.det
      def Matrix.detMonoidHom {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] :
      Matrix n n R →* R

      The determinant of a matrix, as a monoid homomorphism.

      Equations
      • Matrix.detMonoidHom = { toFun := Matrix.det, map_one' := , map_mul' := }
      Instances For
        @[simp]
        theorem Matrix.coe_detMonoidHom {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] :
        Matrix.detMonoidHom = Matrix.det
        theorem Matrix.det_mul_comm {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : Matrix m m R) (N : Matrix m m R) :
        (M * N).det = (N * M).det

        On square matrices, mul_comm applies under det.

        theorem Matrix.det_mul_left_comm {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : Matrix m m R) (N : Matrix m m R) (P : Matrix m m R) :
        (M * (N * P)).det = (N * (M * P)).det

        On square matrices, mul_left_comm applies under det.

        theorem Matrix.det_mul_right_comm {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : Matrix m m R) (N : Matrix m m R) (P : Matrix m m R) :
        (M * N * P).det = (M * P * N).det

        On square matrices, mul_right_comm applies under det.

        theorem Matrix.det_units_conj {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : (Matrix m m R)ˣ) (N : Matrix m m R) :
        (M * N * M⁻¹).det = N.det
        theorem Matrix.det_units_conj' {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : (Matrix m m R)ˣ) (N : Matrix m m R) :
        (M⁻¹ * N * M).det = N.det
        @[simp]
        theorem Matrix.det_transpose {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) :
        M.transpose.det = M.det

        Transposing a matrix preserves the determinant.

        theorem Matrix.det_permute {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (σ : Equiv.Perm n) (M : Matrix n n R) :
        (M.submatrix (σ) id).det = (Equiv.Perm.sign σ) * M.det

        Permuting the columns changes the sign of the determinant.

        theorem Matrix.det_permute' {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (σ : Equiv.Perm n) (M : Matrix n n R) :
        (M.submatrix id σ).det = (Equiv.Perm.sign σ) * M.det

        Permuting the rows changes the sign of the determinant.

        @[simp]
        theorem Matrix.det_submatrix_equiv_self {m : Type u_1} {n : Type u_2} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (e : n m) (A : Matrix m m R) :
        (A.submatrix e e).det = A.det

        Permuting rows and columns with the same equivalence has no effect.

        theorem Matrix.det_reindex_self {m : Type u_1} {n : Type u_2} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (e : m n) (A : Matrix m m R) :
        ((Matrix.reindex e e) A).det = A.det

        Reindexing both indices along the same equivalence preserves the determinant.

        For the simp version of this lemma, see det_submatrix_equiv_self; this one is unsuitable because Matrix.reindex_apply unfolds reindex first.

        @[simp]
        theorem Matrix.det_permutation {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (σ : Equiv.Perm n) :
        (Equiv.toPEquiv σ).toMatrix.det = (Equiv.Perm.sign σ)

        The determinant of a permutation matrix equals its sign.

        theorem Matrix.det_smul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) (c : R) :
        (c A).det = c ^ Fintype.card n * A.det
        @[simp]
        theorem Matrix.det_smul_of_tower {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {α : Type u_3} [Monoid α] [DistribMulAction α R] [IsScalarTower α R R] [SMulCommClass α R R] (c : α) (A : Matrix n n R) :
        (c A).det = c ^ Fintype.card n A.det
        theorem Matrix.det_neg {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) :
        (-A).det = (-1) ^ Fintype.card n * A.det
        theorem Matrix.det_neg_eq_smul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) :
        (-A).det = (-1) ^ Fintype.card n A.det

        A variant of Matrix.det_neg with scalar multiplication by Units instead of multiplication by R.

        theorem Matrix.det_mul_row {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (v : nR) (A : Matrix n n R) :
        (Matrix.of fun (i j : n) => v j * A i j).det = (Finset.univ.prod fun (i : n) => v i) * A.det

        Multiplying each row by a fixed v i multiplies the determinant by the product of the vs.

        theorem Matrix.det_mul_column {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (v : nR) (A : Matrix n n R) :
        (Matrix.of fun (i j : n) => v i * A i j).det = (Finset.univ.prod fun (i : n) => v i) * A.det

        Multiplying each column by a fixed v j multiplies the determinant by the product of the vs.

        @[simp]
        theorem Matrix.det_pow {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (M : Matrix m m R) (n : ) :
        (M ^ n).det = M.det ^ n
        theorem RingHom.map_det {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {S : Type w} [CommRing S] (f : R →+* S) (M : Matrix n n R) :
        f M.det = (f.mapMatrix M).det
        theorem RingEquiv.map_det {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {S : Type w} [CommRing S] (f : R ≃+* S) (M : Matrix n n R) :
        f M.det = (f.mapMatrix M).det
        theorem AlgHom.map_det {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {S : Type w} [CommRing S] [Algebra R S] {T : Type z} [CommRing T] [Algebra R T] (f : S →ₐ[R] T) (M : Matrix n n S) :
        f M.det = (f.mapMatrix M).det
        theorem AlgEquiv.map_det {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {S : Type w} [CommRing S] [Algebra R S] {T : Type z} [CommRing T] [Algebra R T] (f : S ≃ₐ[R] T) (M : Matrix n n S) :
        f M.det = (f.mapMatrix M).det
        @[simp]
        theorem Matrix.det_conjTranspose {m : Type u_1} [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] [StarRing R] (M : Matrix m m R) :
        M.conjTranspose.det = star M.det

        det_zero section #

        Prove that a matrix with a repeated column has determinant equal to zero.

        theorem Matrix.det_eq_zero_of_row_eq_zero {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} (i : n) (h : ∀ (j : n), A i j = 0) :
        A.det = 0
        theorem Matrix.det_eq_zero_of_column_eq_zero {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} (j : n) (h : ∀ (i : n), A i j = 0) :
        A.det = 0
        theorem Matrix.det_zero_of_row_eq {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {M : Matrix n n R} {i : n} {j : n} (i_ne_j : i j) (hij : M i = M j) :
        M.det = 0

        If a matrix has a repeated row, the determinant will be zero.

        theorem Matrix.det_zero_of_column_eq {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {M : Matrix n n R} {i : n} {j : n} (i_ne_j : i j) (hij : ∀ (k : n), M k i = M k j) :
        M.det = 0

        If a matrix has a repeated column, the determinant will be zero.

        theorem Matrix.det_updateRow_add {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (u : nR) (v : nR) :
        (M.updateRow j (u + v)).det = (M.updateRow j u).det + (M.updateRow j v).det
        theorem Matrix.det_updateColumn_add {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (u : nR) (v : nR) :
        (M.updateColumn j (u + v)).det = (M.updateColumn j u).det + (M.updateColumn j v).det
        theorem Matrix.det_updateRow_smul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (s : R) (u : nR) :
        (M.updateRow j (s u)).det = s * (M.updateRow j u).det
        theorem Matrix.det_updateColumn_smul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (s : R) (u : nR) :
        (M.updateColumn j (s u)).det = s * (M.updateColumn j u).det
        theorem Matrix.det_updateRow_smul' {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (s : R) (u : nR) :
        ((s M).updateRow j u).det = s ^ (Fintype.card n - 1) * (M.updateRow j u).det
        theorem Matrix.det_updateColumn_smul' {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (M : Matrix n n R) (j : n) (s : R) (u : nR) :
        ((s M).updateColumn j u).det = s ^ (Fintype.card n - 1) * (M.updateColumn j u).det

        det_eq section #

        Lemmas showing the determinant is invariant under a variety of operations.

        theorem Matrix.det_eq_of_eq_mul_det_one {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} {B : Matrix n n R} (C : Matrix n n R) (hC : C.det = 1) (hA : A = B * C) :
        A.det = B.det
        theorem Matrix.det_eq_of_eq_det_one_mul {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} {B : Matrix n n R} (C : Matrix n n R) (hC : C.det = 1) (hA : A = C * B) :
        A.det = B.det
        theorem Matrix.det_updateRow_add_self {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) {i : n} {j : n} (hij : i j) :
        (A.updateRow i (A i + A j)).det = A.det
        theorem Matrix.det_updateColumn_add_self {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) {i : n} {j : n} (hij : i j) :
        (A.updateColumn i fun (k : n) => A k i + A k j).det = A.det
        theorem Matrix.det_updateRow_add_smul_self {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) {i : n} {j : n} (hij : i j) (c : R) :
        (A.updateRow i (A i + c A j)).det = A.det
        theorem Matrix.det_updateColumn_add_smul_self {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] (A : Matrix n n R) {i : n} {j : n} (hij : i j) (c : R) :
        (A.updateColumn i fun (k : n) => A k i + c A k j).det = A.det
        theorem Matrix.det_eq_of_forall_row_eq_smul_add_const_aux {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} {B : Matrix n n R} {s : Finset n} (c : nR) :
        (is, c i = 0)ks, (∀ (i j : n), A i j = B i j + c i * B k j)A.det = B.det
        theorem Matrix.det_eq_of_forall_row_eq_smul_add_const {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {A : Matrix n n R} {B : Matrix n n R} (c : nR) (k : n) (hk : c k = 0) (A_eq : ∀ (i j : n), A i j = B i j + c i * B k j) :
        A.det = B.det

        If you add multiples of row B k to other rows, the determinant doesn't change.

        theorem Matrix.det_eq_of_forall_row_eq_smul_add_pred_aux {R : Type v} [CommRing R] {n : } (k : Fin (n + 1)) (c : Fin nR) (_hc : ∀ (i : Fin n), k < i.succc i = 0) {M : Matrix (Fin n.succ) (Fin n.succ) R} {N : Matrix (Fin n.succ) (Fin n.succ) R} (_h0 : ∀ (j : Fin n.succ), M 0 j = N 0 j) (_hsucc : ∀ (i : Fin n) (j : Fin n.succ), M i.succ j = N i.succ j + c i * M i.castSucc j) :
        M.det = N.det
        theorem Matrix.det_eq_of_forall_row_eq_smul_add_pred {R : Type v} [CommRing R] {n : } {A : Matrix (Fin (n + 1)) (Fin (n + 1)) R} {B : Matrix (Fin (n + 1)) (Fin (n + 1)) R} (c : Fin nR) (A_zero : ∀ (j : Fin (n + 1)), A 0 j = B 0 j) (A_succ : ∀ (i : Fin n) (j : Fin (n + 1)), A i.succ j = B i.succ j + c i * A i.castSucc j) :
        A.det = B.det

        If you add multiples of previous rows to the next row, the determinant doesn't change.

        theorem Matrix.det_eq_of_forall_col_eq_smul_add_pred {R : Type v} [CommRing R] {n : } {A : Matrix (Fin (n + 1)) (Fin (n + 1)) R} {B : Matrix (Fin (n + 1)) (Fin (n + 1)) R} (c : Fin nR) (A_zero : ∀ (i : Fin (n + 1)), A i 0 = B i 0) (A_succ : ∀ (i : Fin (n + 1)) (j : Fin n), A i j.succ = B i j.succ + c j * A i j.castSucc) :
        A.det = B.det

        If you add multiples of previous columns to the next columns, the determinant doesn't change.

        @[simp]
        theorem Matrix.det_blockDiagonal {n : Type u_2} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] {o : Type u_3} [Fintype o] [DecidableEq o] (M : oMatrix n n R) :
        (Matrix.blockDiagonal M).det = Finset.univ.prod fun (k : o) => (M k).det
        @[simp]
        theorem Matrix.det_fromBlocks_zero₂₁ {m : Type u_1} {n : Type u_2} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (A : Matrix m m R) (B : Matrix m n R) (D : Matrix n n R) :
        (A.fromBlocks B 0 D).det = A.det * D.det

        The determinant of a 2×2 block matrix with the lower-left block equal to zero is the product of the determinants of the diagonal blocks. For the generalization to any number of blocks, see Matrix.det_of_upper_triangular.

        @[simp]
        theorem Matrix.det_fromBlocks_zero₁₂ {m : Type u_1} {n : Type u_2} [DecidableEq n] [Fintype n] [DecidableEq m] [Fintype m] {R : Type v} [CommRing R] (A : Matrix m m R) (C : Matrix n m R) (D : Matrix n n R) :
        (A.fromBlocks 0 C D).det = A.det * D.det

        The determinant of a 2×2 block matrix with the upper-right block equal to zero is the product of the determinants of the diagonal blocks. For the generalization to any number of blocks, see Matrix.det_of_lower_triangular.

        theorem Matrix.det_succ_column_zero {R : Type v} [CommRing R] {n : } (A : Matrix (Fin n.succ) (Fin n.succ) R) :
        A.det = Finset.univ.sum fun (i : Fin n.succ) => (-1) ^ i * A i 0 * (A.submatrix i.succAbove Fin.succ).det

        Laplacian expansion of the determinant of an n+1 × n+1 matrix along column 0.

        theorem Matrix.det_succ_row_zero {R : Type v} [CommRing R] {n : } (A : Matrix (Fin n.succ) (Fin n.succ) R) :
        A.det = Finset.univ.sum fun (j : Fin n.succ) => (-1) ^ j * A 0 j * (A.submatrix Fin.succ j.succAbove).det

        Laplacian expansion of the determinant of an n+1 × n+1 matrix along row 0.

        theorem Matrix.det_succ_row {R : Type v} [CommRing R] {n : } (A : Matrix (Fin n.succ) (Fin n.succ) R) (i : Fin n.succ) :
        A.det = Finset.univ.sum fun (j : Fin n.succ) => (-1) ^ (i + j) * A i j * (A.submatrix i.succAbove j.succAbove).det

        Laplacian expansion of the determinant of an n+1 × n+1 matrix along row i.

        theorem Matrix.det_succ_column {R : Type v} [CommRing R] {n : } (A : Matrix (Fin n.succ) (Fin n.succ) R) (j : Fin n.succ) :
        A.det = Finset.univ.sum fun (i : Fin n.succ) => (-1) ^ (i + j) * A i j * (A.submatrix i.succAbove j.succAbove).det

        Laplacian expansion of the determinant of an n+1 × n+1 matrix along column j.

        @[simp]
        theorem Matrix.det_fin_zero {R : Type v} [CommRing R] {A : Matrix (Fin 0) (Fin 0) R} :
        A.det = 1

        Determinant of 0x0 matrix

        theorem Matrix.det_fin_one {R : Type v} [CommRing R] (A : Matrix (Fin 1) (Fin 1) R) :
        A.det = A 0 0

        Determinant of 1x1 matrix

        theorem Matrix.det_fin_one_of {R : Type v} [CommRing R] (a : R) :
        (Matrix.of ![![a]]).det = a
        theorem Matrix.det_fin_two {R : Type v} [CommRing R] (A : Matrix (Fin 2) (Fin 2) R) :
        A.det = A 0 0 * A 1 1 - A 0 1 * A 1 0

        Determinant of 2x2 matrix

        @[simp]
        theorem Matrix.det_fin_two_of {R : Type v} [CommRing R] (a : R) (b : R) (c : R) (d : R) :
        (Matrix.of ![![a, b], ![c, d]]).det = a * d - b * c
        theorem Matrix.det_fin_three {R : Type v} [CommRing R] (A : Matrix (Fin 3) (Fin 3) R) :
        A.det = A 0 0 * A 1 1 * A 2 2 - A 0 0 * A 1 2 * A 2 1 - A 0 1 * A 1 0 * A 2 2 + A 0 1 * A 1 2 * A 2 0 + A 0 2 * A 1 0 * A 2 1 - A 0 2 * A 1 1 * A 2 0

        Determinant of 3x3 matrix