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_3} {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_3} {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_3} {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_3} {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_3} {B : Type u_4} (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_3} {B : Type u_4} {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_3} (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_3} (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_3} {B : Type u_4} {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_3} {B : Type u_4} {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_3} (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_3} (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_3} [Star A] (M : CStarMatrix m n A) :

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

          Equations
          Instances For
            instance CStarMatrix.instStar {n : Type u_2} {A : Type u_3} [Star A] :
            Equations
            instance CStarMatrix.instInhabited {m : Type u_1} {n : Type u_2} {A : Type u_3} [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_3} [Add A] :
            Equations
            instance CStarMatrix.instZero {m : Type u_1} {n : Type u_2} {A : Type u_3} [Zero A] :
            Equations
            instance CStarMatrix.instNeg {m : Type u_1} {n : Type u_2} {A : Type u_3} [Neg A] :
            Equations
            instance CStarMatrix.instSub {m : Type u_1} {n : Type u_2} {A : Type u_3} [Sub A] :
            Equations
            instance CStarMatrix.instUnique {m : Type u_1} {n : Type u_2} {A : Type u_3} [Unique A] :
            Equations
            instance CStarMatrix.instSubsingleton {m : Type u_1} {n : Type u_2} {A : Type u_3} [Subsingleton A] :
            instance CStarMatrix.instNontrivial {m : Type u_1} {n : Type u_2} {A : Type u_3} [Nonempty m] [Nonempty n] [Nontrivial A] :
            instance CStarMatrix.instSMul {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} [SMul R A] :
            SMul R (CStarMatrix m n A)
            Equations
            instance CStarMatrix.instSMulCommClass {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} {S : Type u_6} [SMul R A] [SMul S A] [SMulCommClass R S A] :
            instance CStarMatrix.instIsScalarTower {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} {S : Type u_6} [SMul R S] [SMul R A] [SMul S A] [IsScalarTower R S A] :
            instance CStarMatrix.instIsCentralScalar {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} [SMul R A] [SMul Rᵐᵒᵖ A] [IsCentralScalar R A] :
            instance CStarMatrix.instMulAction {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} [Monoid R] [MulAction R A] :
            Equations
            instance CStarMatrix.instModule {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : 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_3} [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_3} [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_3} {B : Type u_4} [SMul B A] (r : B) (M : Matrix 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_3} [Sub A] (M N : Matrix 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_3} [Neg A] (M : Matrix m n A) (i : m) (j : n) :
            (-M) i j = -M i j

            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_3} [Zero A] :
            @[simp]
            theorem CStarMatrix.of_add_of {m : Type u_1} {n : Type u_2} {A : Type u_3} [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_3} [Sub A] (f g : Matrix m n A) :
            @[simp]
            theorem CStarMatrix.neg_of {m : Type u_1} {n : Type u_2} {A : Type u_3} [Neg A] (f : Matrix m n A) :
            @[simp]
            theorem CStarMatrix.smul_of {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} [SMul R A] (r : R) (f : Matrix m n A) :
            instance CStarMatrix.instStarModule {n : Type u_2} {A : Type u_3} {R : Type u_5} [Star R] [Star A] [SMul R A] [StarModule R A] :
            def CStarMatrix.ofMatrixₗ {m : Type u_1} {n : Type u_2} {A : Type u_3} {R : Type u_5} [AddCommMonoid A] [Semiring R] [Module R A] :

            The equivalence to matrices, bundled as a linear equivalence.

            Equations
            Instances For
              instance CStarMatrix.instOne {n : Type u_2} {A : Type u_3} [Zero A] [One A] [DecidableEq n] :
              Equations
              theorem CStarMatrix.one_apply {n : Type u_2} {A : Type u_3} [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_3} [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_3} [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_3} [Zero A] [One A] [DecidableEq n] {i j : n} :
              j i1 i j = 0
              @[defaultInstance 100]
              instance CStarMatrix.instHMulOfFintypeOfMulOfAddCommMonoid {m : Type u_1} {n : Type u_2} {A : Type u_3} {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_3} {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_3} {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} {A : Type u_3} {R : 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} {A : Type u_3} {R : Type u_5} {l : Type u_7} [Fintype n] [Monoid R] [AddCommMonoid A] [Mul A] [DistribMulAction R A] [SMulCommClass R A A] (M : Matrix m n A) (a : R) (N : Matrix n l A) :
              M * a N = a (M * N)
              instance CStarMatrix.instRing {n : Type u_2} {A : Type u_3} [Fintype n] [DecidableEq n] [Ring A] :
              Equations
              def CStarMatrix.ofMatrixRingEquiv {n : Type u_2} {A : Type u_3} [Fintype n] [Semiring A] :

              ofMatrix bundled as a ring equivalence.

              Equations
              Instances For
                instance CStarMatrix.instAlgebra {n : Type u_2} {A : Type u_3} {R : 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

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

                      The operator norm on CStarMatrix m n A.

                      Equations
                      theorem CStarMatrix.norm_def {m : Type u_1} {n : Type u_2} {A : Type u_3} [Fintype n] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype m] {M : CStarMatrix m n A} :
                      theorem CStarMatrix.norm_entry_le_norm {m : Type u_1} {n : Type u_2} {A : Type u_3} [Fintype n] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype m] {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_3} [Fintype n] [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] [Fintype m] {M : CStarMatrix m n A} {C : NNReal} (h : ∀ (v : WithCStarModule (nA)) (w : WithCStarModule (mA)), inner 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} :
                      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] :

                      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
                      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
                      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