Matrices over star rings. #
Notation #
The locale Matrix
gives the following notation:
ᴴ
forMatrix.conjTranspose
The conjugate transpose of a matrix defined in term of star
.
Equations
- Matrix.«term_ᴴ» = Lean.ParserDescr.trailingNode `Matrix.«term_ᴴ» 1024 1024 (Lean.ParserDescr.symbol "ᴴ")
Instances For
@[simp]
theorem
Matrix.conjTranspose_stdBasisMatrix
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[DecidableEq m]
[AddMonoid α]
[StarAddMonoid α]
(i : m)
(j : n)
(a : α)
:
@[simp]
theorem
Matrix.diagonal_conjTranspose
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[AddMonoid α]
[StarAddMonoid α]
(v : n → α)
:
@[simp]
theorem
Matrix.diag_conjTranspose
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
(A : Matrix n n α)
:
theorem
Matrix.star_mulVec
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[NonUnitalSemiring α]
[Fintype n]
[StarRing α]
(M : Matrix m n α)
(v : n → α)
:
theorem
Matrix.star_vecMul
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[NonUnitalSemiring α]
[Fintype m]
[StarRing α]
(M : Matrix m n α)
(v : m → α)
:
theorem
Matrix.mulVec_conjTranspose
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[NonUnitalSemiring α]
[Fintype m]
[StarRing α]
(A : Matrix m n α)
(x : m → α)
:
theorem
Matrix.vecMul_conjTranspose
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[NonUnitalSemiring α]
[Fintype n]
[StarRing α]
(A : Matrix m n α)
(x : n → α)
:
@[simp]
theorem
Matrix.conjTranspose_conjTranspose
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[InvolutiveStar α]
(M : Matrix m n α)
:
theorem
Matrix.conjTranspose_injective
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[InvolutiveStar α]
:
@[simp]
theorem
Matrix.conjTranspose_inj
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[InvolutiveStar α]
{A B : Matrix m n α}
:
@[simp]
theorem
Matrix.conjTranspose_eq_diagonal
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[AddMonoid α]
[StarAddMonoid α]
{M : Matrix n n α}
{v : n → α}
:
@[simp]
theorem
Matrix.conjTranspose_zero
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
:
@[simp]
theorem
Matrix.conjTranspose_eq_zero
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
{M : Matrix m n α}
:
@[simp]
theorem
Matrix.conjTranspose_one
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
:
@[simp]
theorem
Matrix.conjTranspose_eq_one
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
{M : Matrix n n α}
:
@[simp]
theorem
Matrix.conjTranspose_natCast
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
(d : ℕ)
:
@[simp]
theorem
Matrix.conjTranspose_eq_natCast
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
{M : Matrix n n α}
{d : ℕ}
:
@[simp]
theorem
Matrix.conjTranspose_ofNat
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
(d : ℕ)
[d.AtLeastTwo]
:
@[simp]
theorem
Matrix.conjTranspose_eq_ofNat
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Semiring α]
[StarRing α]
{M : Matrix n n α}
{d : ℕ}
[d.AtLeastTwo]
:
@[simp]
theorem
Matrix.conjTranspose_intCast
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Ring α]
[StarRing α]
(d : ℤ)
:
@[simp]
theorem
Matrix.conjTranspose_eq_intCast
{n : Type u_3}
{α : Type v}
[DecidableEq n]
[Ring α]
[StarRing α]
{M : Matrix n n α}
{d : ℤ}
:
@[simp]
theorem
Matrix.conjTranspose_add
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
(M N : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_sub
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddGroup α]
[StarAddMonoid α]
(M N : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[Star R]
[Star α]
[SMul R α]
[StarModule R α]
(c : R)
(M : Matrix m n α)
:
Note that StarModule
is quite a strong requirement; as such we also provide the following
variants which this lemma would not apply to:
@[simp]
theorem
Matrix.conjTranspose_nsmul
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
(c : ℕ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_zsmul
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddGroup α]
[StarAddMonoid α]
(c : ℤ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_natCast_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[Semiring R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
(c : ℕ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_ofNat_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[Semiring R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
(c : ℕ)
[c.AtLeastTwo]
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_intCast_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[Ring R]
[AddCommGroup α]
[StarAddMonoid α]
[Module R α]
(c : ℤ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_inv_natCast_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[DivisionSemiring R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
(c : ℕ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_inv_ofNat_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[DivisionSemiring R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
(c : ℕ)
[c.AtLeastTwo]
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_inv_intCast_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[DivisionRing R]
[AddCommGroup α]
[StarAddMonoid α]
[Module R α]
(c : ℤ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_ratCast_smul
{m : Type u_2}
{n : Type u_3}
{R : Type u_7}
{α : Type v}
[DivisionRing R]
[AddCommGroup α]
[StarAddMonoid α]
[Module R α]
(c : ℚ)
(M : Matrix m n α)
:
theorem
Matrix.conjTranspose_rat_smul
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddCommGroup α]
[StarAddMonoid α]
[Module ℚ α]
(c : ℚ)
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_mul
{l : Type u_1}
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[Fintype n]
[NonUnitalSemiring α]
[StarRing α]
(M : Matrix m n α)
(N : Matrix n l α)
:
@[simp]
theorem
Matrix.conjTranspose_neg
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddGroup α]
[StarAddMonoid α]
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTranspose_eq_transpose_of_trivial
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[Star α]
[TrivialStar α]
(A : Matrix m n α)
:
When star x = x
on the coefficients (such as the real numbers) conjTranspose
and transpose
are the same operation.
def
Matrix.conjTransposeAddEquiv
(m : Type u_2)
(n : Type u_3)
(α : Type v)
[AddMonoid α]
[StarAddMonoid α]
:
Matrix.conjTranspose
as an AddEquiv
Equations
- Matrix.conjTransposeAddEquiv m n α = { toFun := Matrix.conjTranspose, invFun := Matrix.conjTranspose, left_inv := ⋯, right_inv := ⋯, map_add' := ⋯ }
Instances For
@[simp]
theorem
Matrix.conjTransposeAddEquiv_apply
(m : Type u_2)
(n : Type u_3)
(α : Type v)
[AddMonoid α]
[StarAddMonoid α]
(M : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTransposeAddEquiv_symm
(m : Type u_2)
(n : Type u_3)
(α : Type v)
[AddMonoid α]
[StarAddMonoid α]
:
theorem
Matrix.conjTranspose_list_sum
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
(l : List (Matrix m n α))
:
theorem
Matrix.conjTranspose_multiset_sum
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddCommMonoid α]
[StarAddMonoid α]
(s : Multiset (Matrix m n α))
:
theorem
Matrix.conjTranspose_sum
{m : Type u_2}
{n : Type u_3}
{α : Type v}
[AddCommMonoid α]
[StarAddMonoid α]
{ι : Type u_10}
(s : Finset ι)
(M : ι → Matrix m n α)
:
def
Matrix.conjTransposeLinearEquiv
(m : Type u_2)
(n : Type u_3)
(R : Type u_7)
(α : Type v)
[CommSemiring R]
[StarRing R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
[StarModule R α]
:
Matrix.conjTranspose
as a LinearMap
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
theorem
Matrix.conjTransposeLinearEquiv_apply
(m : Type u_2)
(n : Type u_3)
(R : Type u_7)
(α : Type v)
[CommSemiring R]
[StarRing R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
[StarModule R α]
(a✝ : Matrix m n α)
:
@[simp]
theorem
Matrix.conjTransposeLinearEquiv_symm
(m : Type u_2)
(n : Type u_3)
(R : Type u_7)
(α : Type v)
[CommSemiring R]
[StarRing R]
[AddCommMonoid α]
[StarAddMonoid α]
[Module R α]
[StarModule R α]
:
@[simp]
theorem
Matrix.conjTranspose_pow
{m : Type u_2}
{α : Type v}
[Semiring α]
[StarRing α]
[Fintype m]
[DecidableEq m]
(M : Matrix m m α)
(k : ℕ)
:
theorem
Matrix.conjTranspose_list_prod
{m : Type u_2}
{α : Type v}
[Semiring α]
[StarRing α]
[Fintype m]
[DecidableEq m]
(l : List (Matrix m m α))
:
When α
has a star operation, square matrices Matrix n n α
have a star
operation equal to Matrix.conjTranspose
.
Equations
- Matrix.instStar = { star := Matrix.conjTranspose }
instance
Matrix.instInvolutiveStar
{n : Type u_3}
{α : Type v}
[InvolutiveStar α]
:
InvolutiveStar (Matrix n n α)
Equations
instance
Matrix.instStarAddMonoid
{n : Type u_3}
{α : Type v}
[AddMonoid α]
[StarAddMonoid α]
:
StarAddMonoid (Matrix n n α)
When α
is a *
-additive monoid, Matrix.star
is also a *
-additive monoid.
Equations
instance
Matrix.instStarModule
{n : Type u_3}
{α : Type v}
{β : Type w}
[Star α]
[Star β]
[SMul α β]
[StarModule α β]
:
StarModule α (Matrix n n β)
instance
Matrix.instStarRing
{n : Type u_3}
{α : Type v}
[Fintype n]
[NonUnitalSemiring α]
[StarRing α]
:
When α
is a *
-(semi)ring, Matrix.star
is also a *
-(semi)ring.