Documentation

Mathlib.Data.Matrix.RowCol

Row and column matrices #

This file provides results about row and column matrices

Main definitions #

def Matrix.col {m : Type u_2} {α : Type v} (w : mα) :

Matrix.col u is the column matrix whose entries are given by u.

Equations
Instances For
    @[simp]
    theorem Matrix.col_apply {m : Type u_2} {α : Type v} (w : mα) (i : m) (j : Unit) :
    Matrix.col w i j = w i
    def Matrix.row {n : Type u_3} {α : Type v} (v : nα) :

    Matrix.row u is the row matrix whose entries are given by u.

    Equations
    Instances For
      @[simp]
      theorem Matrix.row_apply {n : Type u_3} {α : Type v} (v : nα) (i : Unit) (j : n) :
      Matrix.row v i j = v j
      theorem Matrix.col_injective {m : Type u_2} {α : Type v} :
      @[simp]
      theorem Matrix.col_inj {m : Type u_2} {α : Type v} {v : mα} {w : mα} :
      @[simp]
      theorem Matrix.col_zero {m : Type u_2} {α : Type v} [Zero α] :
      @[simp]
      theorem Matrix.col_eq_zero {m : Type u_2} {α : Type v} [Zero α] (v : mα) :
      Matrix.col v = 0 v = 0
      @[simp]
      theorem Matrix.col_add {m : Type u_2} {α : Type v} [Add α] (v : mα) (w : mα) :
      @[simp]
      theorem Matrix.col_smul {m : Type u_2} {R : Type u_5} {α : Type v} [SMul R α] (x : R) (v : mα) :
      theorem Matrix.row_injective {n : Type u_3} {α : Type v} :
      @[simp]
      theorem Matrix.row_inj {n : Type u_3} {α : Type v} {v : nα} {w : nα} :
      @[simp]
      theorem Matrix.row_zero {n : Type u_3} {α : Type v} [Zero α] :
      @[simp]
      theorem Matrix.row_eq_zero {n : Type u_3} {α : Type v} [Zero α] (v : nα) :
      Matrix.row v = 0 v = 0
      @[simp]
      theorem Matrix.row_add {m : Type u_2} {α : Type v} [Add α] (v : mα) (w : mα) :
      @[simp]
      theorem Matrix.row_smul {m : Type u_2} {R : Type u_5} {α : Type v} [SMul R α] (x : R) (v : mα) :
      @[simp]
      theorem Matrix.transpose_col {m : Type u_2} {α : Type v} (v : mα) :
      (Matrix.col v).transpose = Matrix.row v
      @[simp]
      theorem Matrix.transpose_row {m : Type u_2} {α : Type v} (v : mα) :
      (Matrix.row v).transpose = Matrix.col v
      @[simp]
      theorem Matrix.conjTranspose_col {m : Type u_2} {α : Type v} [Star α] (v : mα) :
      (Matrix.col v).conjTranspose = Matrix.row (star v)
      @[simp]
      theorem Matrix.conjTranspose_row {m : Type u_2} {α : Type v} [Star α] (v : mα) :
      (Matrix.row v).conjTranspose = Matrix.col (star v)
      theorem Matrix.row_vecMul {m : Type u_2} {n : Type u_3} {α : Type v} [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : mα) :
      theorem Matrix.col_vecMul {m : Type u_2} {n : Type u_3} {α : Type v} [Fintype m] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : mα) :
      Matrix.col (Matrix.vecMul v M) = (Matrix.row v * M).transpose
      theorem Matrix.col_mulVec {m : Type u_2} {n : Type u_3} {α : Type v} [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : nα) :
      Matrix.col (M.mulVec v) = M * Matrix.col v
      theorem Matrix.row_mulVec {m : Type u_2} {n : Type u_3} {α : Type v} [Fintype n] [NonUnitalNonAssocSemiring α] (M : Matrix m n α) (v : nα) :
      Matrix.row (M.mulVec v) = (M * Matrix.col v).transpose
      @[simp]
      theorem Matrix.row_mul_col_apply {m : Type u_2} {α : Type v} [Fintype m] [Mul α] [AddCommMonoid α] (v : mα) (w : mα) (i : Unit) (j : Unit) :
      @[simp]
      theorem Matrix.diag_col_mul_row {n : Type u_3} {α : Type v} [Mul α] [AddCommMonoid α] (a : nα) (b : nα) :
      (Matrix.col a * Matrix.row b).diag = a * b
      theorem Matrix.vecMulVec_eq {m : Type u_2} {n : Type u_3} {α : Type v} [Mul α] [AddCommMonoid α] (w : mα) (v : nα) :

      Updating rows and columns #

      def Matrix.updateRow {m : Type u_2} {n : Type u_3} {α : Type v} [DecidableEq m] (M : Matrix m n α) (i : m) (b : nα) :
      Matrix m n α

      Update, i.e. replace the ith row of matrix A with the values in b.

      Equations
      Instances For
        def Matrix.updateColumn {m : Type u_2} {n : Type u_3} {α : Type v} [DecidableEq n] (M : Matrix m n α) (j : n) (b : mα) :
        Matrix m n α

        Update, i.e. replace the jth column of matrix A with the values in b.

        Equations
        Instances For
          @[simp]
          theorem Matrix.updateRow_self {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {b : nα} [DecidableEq m] :
          M.updateRow i b i = b
          @[simp]
          theorem Matrix.updateColumn_self {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {j : n} {c : mα} [DecidableEq n] :
          M.updateColumn j c i j = c i
          @[simp]
          theorem Matrix.updateRow_ne {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {b : nα} [DecidableEq m] {i' : m} (i_ne : i' i) :
          M.updateRow i b i' = M i'
          @[simp]
          theorem Matrix.updateColumn_ne {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {j : n} {c : mα} [DecidableEq n] {j' : n} (j_ne : j' j) :
          M.updateColumn j c i j' = M i j'
          theorem Matrix.updateRow_apply {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {j : n} {b : nα} [DecidableEq m] {i' : m} :
          M.updateRow i b i' j = if i' = i then b j else M i' j
          theorem Matrix.updateColumn_apply {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {j : n} {c : mα} [DecidableEq n] {j' : n} :
          M.updateColumn j c i j' = if j' = j then c i else M i j'
          @[simp]
          theorem Matrix.updateColumn_subsingleton {m : Type u_2} {n : Type u_3} {R : Type u_5} [Subsingleton n] (A : Matrix m n R) (i : n) (b : mR) :
          A.updateColumn i b = (Matrix.col b).submatrix id (Function.const n ())
          @[simp]
          theorem Matrix.updateRow_subsingleton {m : Type u_2} {n : Type u_3} {R : Type u_5} [Subsingleton m] (A : Matrix m n R) (i : m) (b : nR) :
          A.updateRow i b = (Matrix.row b).submatrix (Function.const m ()) id
          theorem Matrix.map_updateRow {m : Type u_2} {n : Type u_3} {α : Type v} {β : Type w} {M : Matrix m n α} {i : m} {b : nα} [DecidableEq m] (f : αβ) :
          (M.updateRow i b).map f = (M.map f).updateRow i (f b)
          theorem Matrix.map_updateColumn {m : Type u_2} {n : Type u_3} {α : Type v} {β : Type w} {M : Matrix m n α} {j : n} {c : mα} [DecidableEq n] (f : αβ) :
          (M.updateColumn j c).map f = (M.map f).updateColumn j (f c)
          theorem Matrix.updateRow_transpose {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {j : n} {c : mα} [DecidableEq n] :
          M.transpose.updateRow j c = (M.updateColumn j c).transpose
          theorem Matrix.updateColumn_transpose {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {b : nα} [DecidableEq m] :
          M.transpose.updateColumn i b = (M.updateRow i b).transpose
          theorem Matrix.updateRow_conjTranspose {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {j : n} {c : mα} [DecidableEq n] [Star α] :
          M.conjTranspose.updateRow j (star c) = (M.updateColumn j c).conjTranspose
          theorem Matrix.updateColumn_conjTranspose {m : Type u_2} {n : Type u_3} {α : Type v} {M : Matrix m n α} {i : m} {b : nα} [DecidableEq m] [Star α] :
          M.conjTranspose.updateColumn i (star b) = (M.updateRow i b).conjTranspose
          @[simp]
          theorem Matrix.updateRow_eq_self {m : Type u_2} {n : Type u_3} {α : Type v} [DecidableEq m] (A : Matrix m n α) (i : m) :
          A.updateRow i (A i) = A
          @[simp]
          theorem Matrix.updateColumn_eq_self {m : Type u_2} {n : Type u_3} {α : Type v} [DecidableEq n] (A : Matrix m n α) (i : n) :
          (A.updateColumn i fun (j : m) => A j i) = A
          theorem Matrix.diagonal_updateColumn_single {n : Type u_3} {α : Type v} [DecidableEq n] [Zero α] (v : nα) (i : n) (x : α) :
          (Matrix.diagonal v).updateColumn i (Pi.single i x) = Matrix.diagonal (Function.update v i x)
          theorem Matrix.diagonal_updateRow_single {n : Type u_3} {α : Type v} [DecidableEq n] [Zero α] (v : nα) (i : n) (x : α) :

          Updating rows and columns commutes in the obvious way with reindexing the matrix.

          theorem Matrix.updateRow_submatrix_equiv {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) (r : oα) (e : l m) (f : o n) :
          (A.submatrix e f).updateRow i r = (A.updateRow (e i) fun (j : n) => r (f.symm j)).submatrix e f
          theorem Matrix.submatrix_updateRow_equiv {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) (r : nα) (e : l m) (f : o n) :
          (A.updateRow i r).submatrix e f = (A.submatrix e f).updateRow (e.symm i) fun (i : o) => r (f i)
          theorem Matrix.updateColumn_submatrix_equiv {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : lα) (e : l m) (f : o n) :
          (A.submatrix e f).updateColumn j c = (A.updateColumn (f j) fun (i : m) => c (e.symm i)).submatrix e f
          theorem Matrix.submatrix_updateColumn_equiv {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : mα) (e : l m) (f : o n) :
          (A.updateColumn j c).submatrix e f = (A.submatrix e f).updateColumn (f.symm j) fun (i : l) => c (e i)

          reindex versions of the above submatrix lemmas for convenience.

          theorem Matrix.updateRow_reindex {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : l) (r : oα) (e : m l) (f : n o) :
          ((Matrix.reindex e f) A).updateRow i r = (Matrix.reindex e f) (A.updateRow (e.symm i) fun (j : n) => r (f j))
          theorem Matrix.reindex_updateRow {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq l] [DecidableEq m] (A : Matrix m n α) (i : m) (r : nα) (e : m l) (f : n o) :
          (Matrix.reindex e f) (A.updateRow i r) = ((Matrix.reindex e f) A).updateRow (e i) fun (i : o) => r (f.symm i)
          theorem Matrix.updateColumn_reindex {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : o) (c : lα) (e : m l) (f : n o) :
          ((Matrix.reindex e f) A).updateColumn j c = (Matrix.reindex e f) (A.updateColumn (f.symm j) fun (i : m) => c (e i))
          theorem Matrix.reindex_updateColumn {l : Type u_1} {m : Type u_2} {n : Type u_3} {o : Type u_4} {α : Type v} [DecidableEq o] [DecidableEq n] (A : Matrix m n α) (j : n) (c : mα) (e : m l) (f : n o) :
          (Matrix.reindex e f) (A.updateColumn j c) = ((Matrix.reindex e f) A).updateColumn (f j) fun (i : l) => c (e.symm i)