Matrices with entries in a C⋆-algebra #
This file creates a type copy of Matrix m n A
called CStarMatrix m n A
meant for matrices with
entries in a C⋆-algebra A
. Its action on C⋆ᵐᵒᵈ (n → A)
(via Matrix.mulVec
) gives
it the operator norm, and this norm makes CStarMatrix n n A
a C⋆-algebra.
Main declarations #
CStarMatrix m n A
: the type copyCStarMatrix.instNonUnitalCStarAlgebra
: square matrices with entries in a non-unital C⋆-algebra form a non-unital C⋆-algebraCStarMatrix.instCStarAlgebra
: square matrices with entries in a unital C⋆-algebra form a unital C⋆-algebra
Implementation notes #
The norm on this type induces the product uniformity and bornology, but these are not defeq to
Pi.uniformSpace
and Pi.instBornology
. Hence, we prove the equality to the Pi instances and
replace the uniformity and bornology by the Pi ones when registering the
NormedAddCommGroup (CStarMatrix m n A)
instance. See the docstring of the TopologyAux
section
below for more details.
Type copy Matrix m n A
meant for matrices with entries in a C⋆-algebra. This is
a C⋆-algebra when m = n
.
Equations
- CStarMatrix m n A = Matrix m n A
Instances For
The equivalence between Matrix m n A
and CStarMatrix m n A
.
Equations
- CStarMatrix.ofMatrix = Equiv.refl (Matrix m n A)
Instances For
M.map f
is the matrix obtained by applying f
to each entry of the matrix M
.
Equations
- M.map f = CStarMatrix.ofMatrix fun (i : m) (j : n) => f (M i j)
Instances For
The transpose of a matrix.
Equations
- M.transpose = CStarMatrix.ofMatrix fun (x : n) (y : m) => M y x
Instances For
The conjugate transpose of a matrix defined in term of star
.
Equations
- M.conjTranspose = M.transpose.map star
Instances For
Equations
- CStarMatrix.instStar = { star := fun (M : CStarMatrix n n A) => M.conjTranspose }
Equations
Equations
- CStarMatrix.instInhabited = inferInstanceAs (Inhabited (m → n → A))
Equations
- CStarMatrix.instFintypeOfDecidableEq α = inferInstanceAs (Fintype (m → n → α))
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
Equations
- CStarMatrix.instModule = Pi.module m (fun (a : m) => n → A) R
simp-normal form pulls of
to the outside, to match the Matrix
API.
Equations
The equivalence to matrices, bundled as a linear equivalence.
Equations
- CStarMatrix.ofMatrixₗ = LinearEquiv.refl R (Matrix m n A)
Instances For
Equations
- CStarMatrix.instOne = inferInstanceAs (One (Matrix n n A))
Equations
Equations
- CStarMatrix.instAddGroupWithOne = AddGroupWithOne.mk ⋯ SubNegMonoid.zsmul ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
Equations
- One or more equations did not get rendered due to their size.
Equations
- CStarMatrix.instMulOfFintypeOfAddCommMonoid = { mul := fun (M N : CStarMatrix n n A) => M * N }
Equations
Equations
Equations
Equations
Equations
Equations
Equations
- CStarMatrix.instSemiring = inferInstanceAs (Semiring (Matrix n n A))
Equations
- CStarMatrix.instRing = inferInstanceAs (Ring (Matrix n n A))
ofMatrix
bundled as a ring equivalence.
Equations
- CStarMatrix.ofMatrixRingEquiv = { toEquiv := CStarMatrix.ofMatrix, map_mul' := ⋯, map_add' := ⋯ }
Instances For
Equations
- CStarMatrix.instStarRing = inferInstanceAs (StarRing (Matrix n n A))
Equations
- CStarMatrix.instAlgebra = inferInstanceAs (Algebra R (Matrix n n A))
ofMatrix
bundled as a star algebra equivalence.
Equations
- CStarMatrix.ofMatrixStarAlgEquiv = { toRingEquiv := CStarMatrix.ofMatrixRingEquiv, map_star' := ⋯, map_smul' := ⋯ }
Instances For
Interpret a CStarMatrix m n A
as a continuous linear map acting on C⋆ᵐᵒᵈ (n → A)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Interpret a CStarMatrix m n A
as a continuous linear map acting on C⋆ᵐᵒᵈ (n → A)
. This
version is specialized to the case m = n
and is bundled as a non-unital algebra homomorphism.
Equations
- CStarMatrix.toCLMNonUnitalAlgHom = { toFun := CStarMatrix.toCLM.toFun, map_smul' := ⋯, map_zero' := ⋯, map_add' := ⋯, map_mul' := ⋯ }
Instances For
The operator norm on CStarMatrix m n A
.
Equations
- CStarMatrix.instNorm = { norm := fun (M : CStarMatrix m n A) => ‖CStarMatrix.toCLM M‖ }
Equations
- CStarMatrix.instUniformSpace = Pi.uniformSpace fun (a : m) => n → A
Equations
Equations
Equations
Matrices with entries in a C⋆-algebra form a C⋆-algebra.
Matrices with entries in a non-unital C⋆-algebra form a non-unital C⋆-algebra.
Equations
Equations
Equations
Matrices with entries in a unital C⋆-algebra form a unital C⋆-algebra.
Equations
ofMatrix
bundled as a continuous linear equivalence.
Equations
- CStarMatrix.ofMatrixL = { toLinearEquiv := CStarMatrix.ofMatrixₗ, continuous_toFun := ⋯, continuous_invFun := ⋯ }