Diagonal matrices #
This file contains the definition and basic results about diagonal matrices.
Main results #
- Matrix.IsDiag: a proposition that states a given square matrix- Ais diagonal.
Tags #
diag, diagonal, matrix
@[simp]
theorem
Matrix.IsDiag.diagonal_diag
{α : Type u_1}
{n : Type u_4}
[Zero α]
[DecidableEq n]
{A : Matrix n n α}
(h : A.IsDiag)
 :
Diagonal matrices are generated by the Matrix.diagonal of their Matrix.diag.
theorem
Matrix.isDiag_iff_diagonal_diag
{α : Type u_1}
{n : Type u_4}
[Zero α]
[DecidableEq n]
(A : Matrix n n α)
 :
Matrix.IsDiag.diagonal_diag as an iff.
theorem
Matrix.isDiag_of_subsingleton
{α : Type u_1}
{n : Type u_4}
[Zero α]
[Subsingleton n]
(A : Matrix n n α)
 :
A.IsDiag
Every matrix indexed by a subsingleton is diagonal.
@[simp]
Every identity matrix is diagonal.
theorem
Matrix.IsDiag.neg
{α : Type u_1}
{n : Type u_4}
[SubtractionMonoid α]
{A : Matrix n n α}
(ha : A.IsDiag)
 :
@[simp]
theorem
Matrix.isDiag_neg_iff
{α : Type u_1}
{n : Type u_4}
[SubtractionMonoid α]
{A : Matrix n n α}
 :
theorem
Matrix.IsDiag.add
{α : Type u_1}
{n : Type u_4}
[AddZeroClass α]
{A B : Matrix n n α}
(ha : A.IsDiag)
(hb : B.IsDiag)
 :
theorem
Matrix.IsDiag.sub
{α : Type u_1}
{n : Type u_4}
[SubtractionMonoid α]
{A B : Matrix n n α}
(ha : A.IsDiag)
(hb : B.IsDiag)
 :
@[simp]
theorem
Matrix.isDiag_smul_one
{α : Type u_1}
(n : Type u_6)
[MulZeroOneClass α]
[DecidableEq n]
(k : α)
 :
theorem
Matrix.IsDiag.conjTranspose
{α : Type u_1}
{n : Type u_4}
[NonUnitalNonAssocSemiring α]
[StarRing α]
{A : Matrix n n α}
(ha : A.IsDiag)
 :
@[simp]
theorem
Matrix.isDiag_conjTranspose_iff
{α : Type u_1}
{n : Type u_4}
[NonUnitalNonAssocSemiring α]
[StarRing α]
{A : Matrix n n α}
 :
theorem
Matrix.IsDiag.kronecker
{α : Type u_1}
{n : Type u_4}
{m : Type u_5}
[MulZeroClass α]
{A : Matrix m m α}
{B : Matrix n n α}
(hA : A.IsDiag)
(hB : B.IsDiag)
 :
(kroneckerMap (fun (x1 x2 : α) => x1 * x2) A B).IsDiag
(A ⊗ B).IsDiag if both A and B are diagonal.
theorem
Matrix.IsDiag.fromBlocks
{α : Type u_1}
{n : Type u_4}
{m : Type u_5}
[Zero α]
{A : Matrix m m α}
{D : Matrix n n α}
(ha : A.IsDiag)
(hd : D.IsDiag)
 :
(Matrix.fromBlocks A 0 0 D).IsDiag
The block matrix A.fromBlocks 0 0 D is diagonal if A and D are diagonal.
theorem
Matrix.IsDiag.fromBlocks_of_isSymm
{α : Type u_1}
{n : Type u_4}
{m : Type u_5}
[Zero α]
{A : Matrix m m α}
{C : Matrix n m α}
{D : Matrix n n α}
(h : (Matrix.fromBlocks A 0 C D).IsSymm)
(ha : A.IsDiag)
(hd : D.IsDiag)
 :
(Matrix.fromBlocks A 0 C D).IsDiag
A symmetric block matrix A.fromBlocks B C D is diagonal
if A and D are diagonal and B is 0.
theorem
Matrix.mul_transpose_self_isDiag_iff_hasOrthogonalRows
{α : Type u_1}
{n : Type u_4}
{m : Type u_5}
[Fintype n]
[Mul α]
[AddCommMonoid α]
{A : Matrix m n α}
 :
theorem
Matrix.transpose_mul_self_isDiag_iff_hasOrthogonalCols
{α : Type u_1}
{n : Type u_4}
{m : Type u_5}
[Fintype m]
[Mul α]
[AddCommMonoid α]
{A : Matrix m n α}
 :