Documentation

Mathlib.Analysis.CStarAlgebra.CStarMatrix

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 #

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.

def CStarMatrix (m : Type u_1) (n : Type u_2) (A : Type u_3) :
Type (max u_1 u_2 u_3)

Type copy Matrix m n A meant for matrices with entries in a C⋆-algebra. This is a C⋆-algebra when m = n.

Equations
Instances For
    def CStarMatrix.ofMatrix {m : Type u_7} {n : Type u_8} {A : Type u_9} :
    Matrix m n A CStarMatrix m n A

    The equivalence between Matrix m n A and CStarMatrix m n A.

    Equations
    Instances For
      @[simp]
      theorem CStarMatrix.ofMatrix_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} {M : Matrix m n A} {i : m} :
      ofMatrix M i = M i
      @[simp]
      theorem CStarMatrix.ofMatrix_symm_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} {M : CStarMatrix m n A} {i : m} :
      ofMatrix.symm M i = M i
      theorem CStarMatrix.ext_iff {m : Type u_1} {n : Type u_2} {A : Type u_5} {M N : CStarMatrix m n A} :
      (∀ (i : m) (j : n), M i j = N i j) M = N
      theorem CStarMatrix.ext {m : Type u_1} {n : Type u_2} {A : Type u_5} {M₁ M₂ : CStarMatrix m n A} (h : ∀ (i : m) (j : n), M₁ i j = M₂ i j) :
      M₁ = M₂
      def CStarMatrix.map {m : Type u_1} {n : Type u_2} {A : Type u_5} {B : Type u_6} (M : CStarMatrix m n A) (f : AB) :

      M.map f is the matrix obtained by applying f to each entry of the matrix M.

      Equations
      Instances For
        @[simp]
        theorem CStarMatrix.map_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} {B : Type u_6} {M : CStarMatrix m n A} {f : AB} {i : m} {j : n} :
        M.map f i j = f (M i j)
        @[simp]
        theorem CStarMatrix.map_id {m : Type u_1} {n : Type u_2} {A : Type u_5} (M : CStarMatrix m n A) :
        M.map id = M
        @[simp]
        theorem CStarMatrix.map_id' {m : Type u_1} {n : Type u_2} {A : Type u_5} (M : CStarMatrix m n A) :
        (M.map fun (x : A) => x) = M
        theorem CStarMatrix.map_map {m : Type u_1} {n : Type u_2} {A : Type u_5} {B : Type u_6} {C : Type u_7} {M : Matrix m n A} {f : AB} {g : BC} :
        (M.map f).map g = M.map (g f)
        theorem CStarMatrix.map_injective {m : Type u_1} {n : Type u_2} {A : Type u_5} {B : Type u_6} {f : AB} (hf : Function.Injective f) :
        Function.Injective fun (M : CStarMatrix m n A) => M.map f
        def CStarMatrix.transpose {m : Type u_1} {n : Type u_2} {A : Type u_5} (M : CStarMatrix m n A) :

        The transpose of a matrix.

        Equations
        Instances For
          @[simp]
          theorem CStarMatrix.transpose_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} (M : CStarMatrix m n A) (i : n) (j : m) :
          M.transpose i j = M j i
          def CStarMatrix.conjTranspose {m : Type u_1} {n : Type u_2} {A : Type u_5} [Star A] (M : CStarMatrix m n A) :

          The conjugate transpose of a matrix defined in term of star.

          Equations
          Instances For
            @[simp]
            theorem CStarMatrix.conjTranspose_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Star A] (M : CStarMatrix m n A) (i : n) (j : m) :
            M.conjTranspose i j = star (M j i)
            instance CStarMatrix.instStar {n : Type u_2} {A : Type u_5} [Star A] :
            Equations
            theorem CStarMatrix.star_eq_conjTranspose {n : Type u_2} {A : Type u_5} [Star A] {M : CStarMatrix n n A} :
            Equations
            instance CStarMatrix.instInhabited {m : Type u_1} {n : Type u_2} {A : Type u_5} [Inhabited A] :
            Equations
            instance CStarMatrix.instFinite {n : Type u_7} {m : Type u_8} [Finite m] [Finite n] (α : Type u_9) [Finite α] :
            instance CStarMatrix.instAdd {m : Type u_1} {n : Type u_2} {A : Type u_5} [Add A] :
            Equations
            instance CStarMatrix.instZero {m : Type u_1} {n : Type u_2} {A : Type u_5} [Zero A] :
            Equations
            instance CStarMatrix.instNeg {m : Type u_1} {n : Type u_2} {A : Type u_5} [Neg A] :
            Equations
            instance CStarMatrix.instSub {m : Type u_1} {n : Type u_2} {A : Type u_5} [Sub A] :
            Equations
            instance CStarMatrix.instUnique {m : Type u_1} {n : Type u_2} {A : Type u_5} [Unique A] :
            Equations
            instance CStarMatrix.instSubsingleton {m : Type u_1} {n : Type u_2} {A : Type u_5} [Subsingleton A] :
            instance CStarMatrix.instNontrivial {m : Type u_1} {n : Type u_2} {A : Type u_5} [Nonempty m] [Nonempty n] [Nontrivial A] :
            instance CStarMatrix.instSMul {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [SMul R A] :
            SMul R (CStarMatrix m n A)
            Equations
            instance CStarMatrix.instSMulCommClass {m : Type u_1} {n : Type u_2} {R : Type u_3} {S : Type u_4} {A : Type u_5} [SMul R A] [SMul S A] [SMulCommClass R S A] :
            instance CStarMatrix.instIsScalarTower {m : Type u_1} {n : Type u_2} {R : Type u_3} {S : Type u_4} {A : Type u_5} [SMul R S] [SMul R A] [SMul S A] [IsScalarTower R S A] :
            instance CStarMatrix.instIsCentralScalar {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [SMul R A] [SMul Rᵐᵒᵖ A] [IsCentralScalar R A] :
            instance CStarMatrix.instMulAction {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [Monoid R] [MulAction R A] :
            Equations
            instance CStarMatrix.instModule {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [Semiring R] [AddCommMonoid A] [Module R A] :
            Equations
            @[simp]
            theorem CStarMatrix.zero_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Zero A] (i : m) (j : n) :
            0 i j = 0
            @[simp]
            theorem CStarMatrix.add_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Add A] (M N : CStarMatrix m n A) (i : m) (j : n) :
            (M + N) i j = M i j + N i j
            @[simp]
            theorem CStarMatrix.smul_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} {B : Type u_6} [SMul B A] (r : B) (M : CStarMatrix m n A) (i : m) (j : n) :
            (r M) i j = r M i j
            @[simp]
            theorem CStarMatrix.sub_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Sub A] (M N : CStarMatrix m n A) (i : m) (j : n) :
            (M - N) i j = M i j - N i j
            @[simp]
            theorem CStarMatrix.neg_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Neg A] (M : CStarMatrix m n A) (i : m) (j : n) :
            (-M) i j = -M i j
            @[simp]
            theorem CStarMatrix.conjTranspose_zero {m : Type u_1} {n : Type u_2} {A : Type u_5} [AddMonoid A] [StarAddMonoid A] :

            simp-normal form pulls of to the outside, to match the Matrix API.

            @[simp]
            theorem CStarMatrix.of_zero {m : Type u_1} {n : Type u_2} {A : Type u_5} [Zero A] :
            @[simp]
            theorem CStarMatrix.of_add_of {m : Type u_1} {n : Type u_2} {A : Type u_5} [Add A] (f g : Matrix m n A) :
            @[simp]
            theorem CStarMatrix.of_sub_of {m : Type u_1} {n : Type u_2} {A : Type u_5} [Sub A] (f g : Matrix m n A) :
            @[simp]
            theorem CStarMatrix.neg_of {m : Type u_1} {n : Type u_2} {A : Type u_5} [Neg A] (f : Matrix m n A) :
            @[simp]
            theorem CStarMatrix.smul_of {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [SMul R A] (r : R) (f : Matrix m n A) :
            theorem CStarMatrix.star_apply {n : Type u_2} {A : Type u_5} [Star A] {f : CStarMatrix n n A} {i j : n} :
            star f i j = star (f j i)
            theorem CStarMatrix.star_apply_of_isSelfAdjoint {n : Type u_2} {A : Type u_5} [Star A] {f : CStarMatrix n n A} (hf : IsSelfAdjoint f) {i j : n} :
            star (f i j) = f j i
            Equations
            instance CStarMatrix.instStarModule {n : Type u_2} {R : Type u_3} {A : Type u_5} [Star R] [Star A] [SMul R A] [StarModule R A] :
            def CStarMatrix.ofMatrixₗ {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [AddCommMonoid A] [Semiring R] [Module R A] :

            The equivalence to matrices, bundled as a linear equivalence.

            Equations
            Instances For
              def CStarMatrix.mapₗ {m : Type u_1} {n : Type u_2} {R : Type u_3} {S : Type u_4} {A : Type u_5} {B : Type u_6} [Semiring R] [Semiring S] {σ : R →+* S} [AddCommMonoid A] [AddCommMonoid B] [Module R A] [Module S B] (f : A →ₛₗ[σ] B) :

              The semilinear map constructed by applying a semilinear map to all the entries of the matrix.

              Equations
              Instances For
                @[simp]
                theorem CStarMatrix.mapₗ_apply {m : Type u_1} {n : Type u_2} {R : Type u_3} {S : Type u_4} {A : Type u_5} {B : Type u_6} [Semiring R] [Semiring S] {σ : R →+* S} [AddCommMonoid A] [AddCommMonoid B] [Module R A] [Module S B] (f : A →ₛₗ[σ] B) (M : CStarMatrix m n A) :
                (mapₗ f) M = M.map f
                instance CStarMatrix.instOne {n : Type u_2} {A : Type u_5} [Zero A] [One A] [DecidableEq n] :
                Equations
                theorem CStarMatrix.one_apply {n : Type u_2} {A : Type u_5} [Zero A] [One A] [DecidableEq n] {i j : n} :
                1 i j = if i = j then 1 else 0
                @[simp]
                theorem CStarMatrix.one_apply_eq {n : Type u_2} {A : Type u_5} [Zero A] [One A] [DecidableEq n] (i : n) :
                1 i i = 1
                @[simp]
                theorem CStarMatrix.one_apply_ne {n : Type u_2} {A : Type u_5} [Zero A] [One A] [DecidableEq n] {i j : n} :
                i j1 i j = 0
                theorem CStarMatrix.one_apply_ne' {n : Type u_2} {A : Type u_5} [Zero A] [One A] [DecidableEq n] {i j : n} :
                j i1 i j = 0
                Equations
                Equations
                • One or more equations did not get rendered due to their size.
                Equations
                • One or more equations did not get rendered due to their size.
                Equations
                • One or more equations did not get rendered due to their size.
                @[defaultInstance 100]
                instance CStarMatrix.instHMulOfFintypeOfMulOfAddCommMonoid {m : Type u_1} {n : Type u_2} {A : Type u_5} {l : Type u_7} [Fintype m] [Mul A] [AddCommMonoid A] :
                HMul (CStarMatrix l m A) (CStarMatrix m n A) (CStarMatrix l n A)
                Equations
                • One or more equations did not get rendered due to their size.
                Equations
                theorem CStarMatrix.mul_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} {l : Type u_7} [Fintype m] [Mul A] [AddCommMonoid A] {M : CStarMatrix l m A} {N : CStarMatrix m n A} {i : l} {k : n} :
                (M * N) i k = j : m, M i j * N j k
                theorem CStarMatrix.mul_apply' {m : Type u_1} {n : Type u_2} {A : Type u_5} {l : Type u_7} [Fintype m] [Mul A] [AddCommMonoid A] {M : CStarMatrix l m A} {N : CStarMatrix m n A} {i : l} {k : n} :
                (M * N) i k = (fun (j : m) => M i j) ⬝ᵥ fun (j : m) => N j k
                @[simp]
                theorem CStarMatrix.smul_mul {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} {l : Type u_7} [Fintype n] [Monoid R] [AddCommMonoid A] [Mul A] [DistribMulAction R A] [IsScalarTower R A A] (a : R) (M : CStarMatrix m n A) (N : CStarMatrix n l A) :
                a M * N = a (M * N)
                theorem CStarMatrix.mul_smul {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} {l : Type u_7} [Fintype n] [Monoid R] [AddCommMonoid A] [Mul A] [DistribMulAction R A] [SMulCommClass R A A] (M : CStarMatrix m n A) (a : R) (N : CStarMatrix n l A) :
                M * a N = a (M * N)
                @[simp]
                theorem CStarMatrix.mul_zero {m : Type u_1} {n : Type u_2} {A : Type u_5} {o : Type u_7} [Fintype n] [NonUnitalNonAssocSemiring A] (M : CStarMatrix m n A) :
                M * 0 = 0
                @[simp]
                theorem CStarMatrix.zero_mul {m : Type u_1} {n : Type u_2} {A : Type u_5} {l : Type u_7} [Fintype m] [NonUnitalNonAssocSemiring A] (M : CStarMatrix m n A) :
                0 * M = 0
                theorem CStarMatrix.mul_add {m : Type u_1} {n : Type u_2} {A : Type u_5} {o : Type u_7} [Fintype n] [NonUnitalNonAssocSemiring A] (L : CStarMatrix m n A) (M N : CStarMatrix n o A) :
                L * (M + N) = L * M + L * N
                theorem CStarMatrix.add_mul {m : Type u_1} {n : Type u_2} {A : Type u_5} {l : Type u_7} [Fintype m] [NonUnitalNonAssocSemiring A] (L M : CStarMatrix l m A) (N : CStarMatrix m n A) :
                (L + M) * N = L * N + M * N
                instance CStarMatrix.instRing {n : Type u_2} {A : Type u_5} [Fintype n] [DecidableEq n] [Ring A] :
                Equations
                def CStarMatrix.ofMatrixRingEquiv {n : Type u_2} {A : Type u_5} [Fintype n] [Semiring A] :

                ofMatrix bundled as a ring equivalence.

                Equations
                Instances For
                  instance CStarMatrix.instAlgebra {n : Type u_2} {R : Type u_3} {A : Type u_5} [Fintype n] [DecidableEq n] [CommSemiring R] [Semiring A] [Algebra R A] :
                  Equations

                  ofMatrix bundled as a star algebra equivalence.

                  Equations
                  Instances For
                    def CStarMatrix.reindexₗ {m : Type u_1} {n : Type u_2} (R : Type u_3) (A : Type u_5) {l : Type u_7} {o : Type u_8} [Semiring R] [AddCommMonoid A] [Module R A] (eₘ : m l) (eₙ : n o) :

                    The natural map that reindexes a matrix's rows and columns with equivalent types is an equivalence.

                    Equations
                    Instances For
                      @[simp]
                      theorem CStarMatrix.reindexₗ_apply {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} {l : Type u_7} {o : Type u_8} [Semiring R] [AddCommMonoid A] [Module R A] {eₘ : m l} {eₙ : n o} {M : CStarMatrix m n A} {i : l} {j : o} :
                      (reindexₗ R A eₘ eₙ) M i j = (Matrix.reindex eₘ eₙ) M i j
                      def CStarMatrix.reindexₐ {m : Type u_1} {n : Type u_2} (R : Type u_7) (A : Type u_8) [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid A] [Mul A] [Module R A] [Star A] (e : m n) :

                      The natural map that reindexes a matrix's rows and columns with equivalent types is an equivalence.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem CStarMatrix.reindexₐ_apply {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid A] [Mul A] [Star A] [Module R A] {e : m n} {M : CStarMatrix m m A} {i j : n} :
                        (reindexₐ R A e) M i j = (Matrix.reindex e e) M i j
                        theorem CStarMatrix.mapₗ_reindexₐ {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} {B : Type u_6} [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid A] [Mul A] [Module R A] [Star A] [AddCommMonoid B] [Mul B] [Module R B] [Star B] {e : m n} {M : CStarMatrix m m A} (φ : A →ₗ[R] B) :
                        (reindexₐ R B e) ((mapₗ φ) M) = (mapₗ φ) ((reindexₐ R A e) M)
                        @[simp]
                        theorem CStarMatrix.reindexₐ_symm {m : Type u_1} {n : Type u_2} {R : Type u_3} {A : Type u_5} [Fintype m] [Fintype n] [Semiring R] [AddCommMonoid A] [Mul A] [Module R A] [Star A] {e : m n} :
                        def CStarMatrix.mapₙₐ {n : Type u_2} {R : Type u_3} {A : Type u_5} {B : Type u_6} [Fintype n] [Semiring R] [NonUnitalNonAssocSemiring A] [Module R A] [Star A] [NonUnitalNonAssocSemiring B] [Module R B] [Star B] (f : A →⋆ₙₐ[R] B) :

                        Applying a non-unital ⋆-algebra homomorphism to every entry of a matrix is itself a ⋆-algebra homomorphism on matrices.

                        Equations
                        Instances For
                          @[simp]
                          theorem CStarMatrix.mapₙₐ_apply {n : Type u_2} {R : Type u_3} {A : Type u_5} {B : Type u_6} [Fintype n] [Semiring R] [NonUnitalNonAssocSemiring A] [Module R A] [Star A] [NonUnitalNonAssocSemiring B] [Module R B] [Star B] (f : A →⋆ₙₐ[R] B) (M : CStarMatrix n n A) :
                          (mapₙₐ f) M = (mapₗ f) M
                          theorem CStarMatrix.algebraMap_apply {n : Type u_2} {R : Type u_3} {A : Type u_5} [Fintype n] [DecidableEq n] [CommSemiring R] [Semiring A] [Algebra R A] {r : R} {i j : n} :
                          (algebraMap R (CStarMatrix n n A)) r i j = if i = j then (algebraMap R A) r else 0
                          def CStarMatrix.toOneByOne (n : Type u_2) (R : Type u_3) (A : Type u_5) [Unique n] [Semiring R] [AddCommMonoid A] [Mul A] [Star A] [Module R A] :

                          The ⋆-algebra equivalence between A and 1×1 matrices with its entry in A.

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            noncomputable def CStarMatrix.toCLM {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] :

                            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
                              theorem CStarMatrix.toCLM_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] {M : CStarMatrix m n A} {v : WithCStarModule A (mA)} :
                              (toCLM M) v = (WithCStarModule.equiv A (nA)).symm (Matrix.vecMul v M)
                              theorem CStarMatrix.toCLM_apply_eq_sum {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] {M : CStarMatrix m n A} {v : WithCStarModule A (mA)} :
                              (toCLM M) v = (WithCStarModule.equiv A (nA)).symm fun (j : n) => i : m, v i * M i j

                              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
                              Instances For
                                @[simp]
                                theorem CStarMatrix.toCLM_apply_single {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [DecidableEq m] {M : CStarMatrix m n A} {i : m} (a : A) :
                                (toCLM M) ((WithCStarModule.equiv A (mA)).symm (Pi.single i a)) = (WithCStarModule.equiv A (nA)).symm fun (j : n) => a * M i j
                                theorem CStarMatrix.toCLM_apply_single_apply {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [DecidableEq m] {M : CStarMatrix m n A} {i : m} {j : n} (a : A) :
                                (toCLM M) ((WithCStarModule.equiv A (mA)).symm (Pi.single i a)) j = a * M i j
                                theorem CStarMatrix.mul_entry_mul_eq_inner_toCLM {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] [DecidableEq m] [DecidableEq n] {M : CStarMatrix m n A} {i : m} {j : n} (a b : A) :
                                a * M i j * star b = inner A ((WithCStarModule.equiv A (nA)).symm (Pi.single j b)) ((toCLM M) ((WithCStarModule.equiv A (mA)).symm (Pi.single i a)))
                                theorem CStarMatrix.inner_toCLM_conjTranspose_left {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] {M : CStarMatrix m n A} {v : WithCStarModule A (nA)} {w : WithCStarModule A (mA)} :
                                inner A ((toCLM (Matrix.conjTranspose M)) v) w = inner A v ((toCLM M) w)
                                theorem CStarMatrix.inner_toCLM_conjTranspose_right {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] {M : CStarMatrix m n A} {v : WithCStarModule A (mA)} {w : WithCStarModule A (nA)} :
                                inner A v ((toCLM (Matrix.conjTranspose M)) w) = inner A ((toCLM M) v) w
                                noncomputable instance CStarMatrix.instNorm {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] :

                                The operator norm on CStarMatrix m n A.

                                Equations
                                theorem CStarMatrix.norm_def {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] {M : CStarMatrix m n A} :
                                theorem CStarMatrix.norm_entry_le_norm {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] {M : CStarMatrix m n A} {i : m} {j : n} :
                                theorem CStarMatrix.norm_le_of_forall_inner_le {m : Type u_1} {n : Type u_2} {A : Type u_5} [Fintype m] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype n] {M : CStarMatrix m n A} {C : NNReal} (h : ∀ (v : WithCStarModule A (mA)) (w : WithCStarModule A (nA)), inner A w ((toCLM M) v) C * v * w) :
                                M C
                                instance CStarMatrix.instUniformSpace {A : Type u_1} [NonUnitalCStarAlgebra A] {m : Type u_2} {n : Type u_3} :
                                Equations
                                instance CStarMatrix.instT2Space {A : Type u_1} [NonUnitalCStarAlgebra A] {m : Type u_2} {n : Type u_3} :
                                instance CStarMatrix.instT3Space {A : Type u_1} [NonUnitalCStarAlgebra A] {m : Type u_2} {n : Type u_3} :
                                @[deprecated CStarMatrix.instIsUniformAddGroup (since := "2025-03-31")]

                                Alias of CStarMatrix.instIsUniformAddGroup.

                                instance CStarMatrix.instContinuousSMul {A : Type u_1} [NonUnitalCStarAlgebra A] {m : Type u_2} {n : Type u_3} {R : Type u_4} [SMul R A] [TopologicalSpace R] [ContinuousSMul R A] :
                                Equations
                                • One or more equations did not get rendered due to their size.

                                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
                                • One or more equations did not get rendered due to their size.
                                noncomputable instance CStarMatrix.instNormedRing {A : Type u_1} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] {n : Type u_2} [Fintype n] [DecidableEq n] :
                                Equations
                                noncomputable instance CStarMatrix.instNormedAlgebra {A : Type u_1} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] {n : Type u_2} [Fintype n] [DecidableEq n] :
                                Equations
                                noncomputable instance CStarMatrix.instCStarAlgebra {A : Type u_1} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] {n : Type u_2} [Fintype n] [DecidableEq n] :

                                Matrices with entries in a unital C⋆-algebra form a unital C⋆-algebra.

                                Equations
                                • One or more equations did not get rendered due to their size.
                                def CStarMatrix.ofMatrixL {m : Type u_1} {n : Type u_2} {A : Type u_3} [NonUnitalCStarAlgebra A] :

                                ofMatrix bundled as a continuous linear equivalence.

                                Equations
                                Instances For