diff --git a/thys/Complex_Bounded_Operators/Cblinfun_Code.thy b/thys/Complex_Bounded_Operators/Cblinfun_Code.thy --- a/thys/Complex_Bounded_Operators/Cblinfun_Code.thy +++ b/thys/Complex_Bounded_Operators/Cblinfun_Code.thy @@ -1,659 +1,661 @@ section \\Cblinfun_Code\ -- Support for code generation\ text \This theory provides support for code generation involving on complex vector spaces and bounded operators (e.g., types \cblinfun\ and \ell2\). To fully support code generation, in addition to importing this theory, one need to activate support for code generation (import theory \Jordan_Normal_Form.Matrix_Impl\) and for real and complex numbers (import theory \Real_Impl.Real_Impl\ for support of reals of the form \a + b * sqrt c\ or \Algebraic_Numbers.Real_Factorization\ (much slower) for support of algebraic reals; support of complex numbers comes "for free"). The builtin support for real and complex numbers (in \Complex_Main\) is not sufficient because it does not support the computation of square-roots which are used in the setup below. It is also recommended to import \HOL-Library.Code_Target_Numeral\ for faster support of nats and integers.\ theory Cblinfun_Code imports Cblinfun_Matrix Containers.Set_Impl Jordan_Normal_Form.Matrix_Kernel begin no_notation "Lattice.meet" (infixl "\\" 70) no_notation "Lattice.join" (infixl "\\" 65) hide_const (open) Coset.kernel hide_const (open) Matrix_Kernel.kernel hide_const (open) Order.bottom Order.top +unbundle lattice_syntax unbundle jnf_notation unbundle cblinfun_notation subsection \Code equations for cblinfun operators\ text \In this subsection, we define the code for all operations involving only operators (no combinations of operators/vectors/subspaces)\ text \The following lemma registers cblinfun as an abstract datatype with constructor \<^const>\cblinfun_of_mat\. That means that in generated code, all cblinfun operators will be represented as \<^term>\cblinfun_of_mat X\ where X is a matrix. In code equations for operations involving operators (e.g., +), we can then write the equation directly in terms of matrices by writing, e.g., \<^term>\mat_of_cblinfun (A+B)\ in the lhs, and in the rhs we define the matrix that corresponds to the sum of A,B. In the rhs, we can access the matrices corresponding to A,B by writing \<^term>\mat_of_cblinfun B\. (See, e.g., lemma \cblinfun_of_mat_plusOp\ below). See @{cite "code-generation-tutorial"} for more information on @{theory_text \[code abstype]\}.\ declare mat_of_cblinfun_inverse [code abstype] text \This lemma defines addition. By writing \<^term>\mat_of_cblinfun (M + N)\ on the left hand side, we get access to the\ declare mat_of_cblinfun_plus[code] \ \Code equation for addition of cblinfuns\ declare mat_of_cblinfun_id[code] \ \Code equation for computing the identity operator\ declare mat_of_cblinfun_1[code] \ \Code equation for computing the one-dimensional identity\ declare mat_of_cblinfun_zero[code] \ \Code equation for computing the zero operator\ declare mat_of_cblinfun_uminus[code] \ \Code equation for computing the unary minus on cblinfun's\ declare mat_of_cblinfun_minus[code] \ \Code equation for computing the difference of cblinfun's\ declare mat_of_cblinfun_classical_operator[code] \ \Code equation for computing the "classical operator"\ declare mat_of_cblinfun_compose[code] \ \Code equation for computing the composition/product of cblinfun's\ declare mat_of_cblinfun_scaleC[code] \ \Code equation for multiplication with complex scalar\ declare mat_of_cblinfun_scaleR[code] \ \Code equation for multiplication with real scalar\ declare mat_of_cblinfun_adj[code] \ \Code equation for computing the adj\ text \This instantiation defines a code equation for equality tests for cblinfun.\ instantiation cblinfun :: (onb_enum,onb_enum) equal begin definition [code]: "equal_cblinfun M N \ mat_of_cblinfun M = mat_of_cblinfun N" for M N :: "'a \\<^sub>C\<^sub>L 'b" instance apply intro_classes unfolding equal_cblinfun_def using mat_of_cblinfun_inj injD by fastforce end subsection \Vectors\ text \In this section, we define code for operations on vectors. As with operators above, we do this by using an isomorphism between finite vectors (i.e., types T of sort \complex_vector\) and the type \<^typ>\complex vec\ from \<^session>\Jordan_Normal_Form\. We have developed such an isomorphism in theory \Cblinfun_Matrix\ for any type T of sort \onb_enum\ (i.e., any type with a finite canonical orthonormal basis) as was done above for bounded operators. Unfortunately, we cannot declare code equations for a type class, code equations must be related to a specific type constructor. So we give code definition only for vectors of type \<^typ>\'a ell2\ (where \<^typ>\'a\ must be of sort \enum\ to make make sure that \<^typ>\'a ell2\ is finite dimensional). The isomorphism between \<^typ>\'a ell2\ is given by the constants \ell2_of_vec\ and \vec_of_ell2\ which are copies of the more general \<^const>\basis_enum_of_vec\ and \<^const>\vec_of_basis_enum\ but with a more restricted type to be usable in our code equations. \ definition ell2_of_vec :: "complex vec \ 'a::enum ell2" where "ell2_of_vec = basis_enum_of_vec" definition vec_of_ell2 :: "'a::enum ell2 \ complex vec" where "vec_of_ell2 = vec_of_basis_enum" text \The following theorem registers the isomorphism \ell2_of_vec\/\vec_of_ell2\ for code generation. From now on, code for operations on \<^typ>\_ ell2\ can be expressed by declarations such as \<^term>\vec_of_ell2 (f a b) = g (vec_of_ell2 a) (vec_of_ell2 b)\ if the operation f on \<^typ>\_ ell2\ corresponds to the operation g on \<^typ>\complex vec\.\ lemma vec_of_ell2_inverse [code abstype]: "ell2_of_vec (vec_of_ell2 B) = B" unfolding ell2_of_vec_def vec_of_ell2_def by (rule vec_of_basis_enum_inverse) text \This instantiation defines a code equation for equality tests for ell2.\ instantiation ell2 :: (enum) equal begin definition [code]: "equal_ell2 M N \ vec_of_ell2 M = vec_of_ell2 N" for M N :: "'a::enum ell2" instance apply intro_classes unfolding equal_ell2_def by (metis vec_of_ell2_inverse) end lemma vec_of_ell2_zero[code]: \ \Code equation for computing the zero vector\ "vec_of_ell2 (0::'a::enum ell2) = zero_vec (CARD('a))" by (simp add: vec_of_ell2_def vec_of_basis_enum_zero) lemma vec_of_ell2_ket[code]: \ \Code equation for computing a standard basis vector\ "vec_of_ell2 (ket i) = unit_vec (CARD('a)) (enum_idx i)" for i::"'a::enum" using vec_of_ell2_def vec_of_basis_enum_ket by metis lemma vec_of_ell2_timesScalarVec[code]: \ \Code equation for multiplying a vector with a complex scalar\ "vec_of_ell2 (scaleC a \) = smult_vec a (vec_of_ell2 \)" for \ :: "'a::enum ell2" by (simp add: vec_of_ell2_def vec_of_basis_enum_scaleC) lemma vec_of_ell2_scaleR[code]: \ \Code equation for multiplying a vector with a real scalar\ "vec_of_ell2 (scaleR a \) = smult_vec (complex_of_real a) (vec_of_ell2 \)" for \ :: "'a::enum ell2" by (simp add: vec_of_ell2_def vec_of_basis_enum_scaleR) lemma ell2_of_vec_plus[code]: \ \Code equation for adding vectors\ "vec_of_ell2 (x + y) = (vec_of_ell2 x) + (vec_of_ell2 y)" for x y :: "'a::enum ell2" by (simp add: vec_of_ell2_def vec_of_basis_enum_add) lemma ell2_of_vec_minus[code]: \ \Code equation for subtracting vectors\ "vec_of_ell2 (x - y) = (vec_of_ell2 x) - (vec_of_ell2 y)" for x y :: "'a::enum ell2" by (simp add: vec_of_ell2_def vec_of_basis_enum_minus) lemma ell2_of_vec_uminus[code]: \ \Code equation for negating a vector\ "vec_of_ell2 (- y) = - (vec_of_ell2 y)" for y :: "'a::enum ell2" by (simp add: vec_of_ell2_def vec_of_basis_enum_uminus) lemma cinner_ell2_code' [code]: "cinner \ \ = cscalar_prod (vec_of_ell2 \) (vec_of_ell2 \)" \ \Code equation for the inner product of vectors\ by (simp add: cscalar_prod_vec_of_basis_enum vec_of_ell2_def) lemma norm_ell2_code [code]: \ \Code equation for the norm of a vector\ "norm \ = (let \' = vec_of_ell2 \ in sqrt (\ i \ {0 ..< dim_vec \'}. let z = vec_index \' i in (Re z)\<^sup>2 + (Im z)\<^sup>2))" by (simp add: norm_ell2_vec_of_basis_enum vec_of_ell2_def) lemma times_ell2_code'[code]: \ \Code equation for the product in the algebra of one-dimensional vectors\ fixes \ \ :: "'a::{CARD_1,enum} ell2" shows "vec_of_ell2 (\ * \) = vec_of_list [vec_index (vec_of_ell2 \) 0 * vec_index (vec_of_ell2 \) 0]" by (simp add: vec_of_ell2_def vec_of_basis_enum_times) lemma divide_ell2_code'[code]: \ \Code equation for the product in the algebra of one-dimensional vectors\ fixes \ \ :: "'a::{CARD_1,enum} ell2" shows "vec_of_ell2 (\ / \) = vec_of_list [vec_index (vec_of_ell2 \) 0 / vec_index (vec_of_ell2 \) 0]" by (simp add: vec_of_ell2_def vec_of_basis_enum_divide) lemma inverse_ell2_code'[code]: \ \Code equation for the product in the algebra of one-dimensional vectors\ fixes \ :: "'a::{CARD_1,enum} ell2" shows "vec_of_ell2 (inverse \) = vec_of_list [inverse (vec_index (vec_of_ell2 \) 0)]" by (simp add: vec_of_ell2_def vec_of_basis_enum_to_inverse) lemma one_ell2_code'[code]: \ \Code equation for the unit in the algebra of one-dimensional vectors\ "vec_of_ell2 (1 :: 'a::{CARD_1,enum} ell2) = vec_of_list [1]" by (simp add: vec_of_ell2_def vec_of_basis_enum_1) subsection \Vector/Matrix\ text \We proceed to give code equations for operations involving both operators (cblinfun) and vectors. As explained above, we have to restrict the equations to vectors of type \<^typ>\'a ell2\ even though the theory is available for any type of class \<^class>\onb_enum\. As a consequence, we run into an addition technicality now. For example, to define a code equation for applying an operator to a vector, we might try to give the following lemma: \<^theory_text>\lemma cblinfun_apply_code[code]: "vec_of_ell2 (M *\<^sub>V x) = (mult_mat_vec (mat_of_cblinfun M) (vec_of_ell2 x))" by (simp add: mat_of_cblinfun_cblinfun_apply vec_of_ell2_def)\ Unfortunately, this does not work, Isabelle produces the warning "Projection as head in equation", most likely due to the fact that the type of \<^term>\(*\<^sub>V)\ in the equation is less general than the type of \<^term>\(*\<^sub>V)\ (it is restricted to @{type ell2}). We overcome this problem by defining a constant \cblinfun_apply_code\ which is equal to \<^term>\(*\<^sub>V)\ but has a more restricted type. We then instruct the code generation to replace occurrences of \<^term>\(*\<^sub>V)\ by \cblinfun_apply_code\ (where possible), and we add code generation for \cblinfun_apply_code\ instead of \<^term>\(*\<^sub>V)\. \ definition cblinfun_apply_code :: "'a ell2 \\<^sub>C\<^sub>L 'b ell2 \ 'a ell2 \ 'b ell2" where [code del, code_abbrev]: "cblinfun_apply_code = (*\<^sub>V)" \ \@{attribute code_abbrev} instructs the code generation to replace the rhs \<^term>\(*\<^sub>V)\ by the lhs \<^term>\cblinfun_apply_code\ before starting the actual code generation.\ lemma cblinfun_apply_code[code]: \ \Code equation for \<^term>\cblinfun_apply_code\, i.e., for applying an operator to an \<^type>\ell2\ vector\ "vec_of_ell2 (cblinfun_apply_code M x) = (mult_mat_vec (mat_of_cblinfun M) (vec_of_ell2 x))" by (simp add: cblinfun_apply_code_def mat_of_cblinfun_cblinfun_apply vec_of_ell2_def) text \For the constant \<^term>\vector_to_cblinfun\ (canonical isomorphism from vectors to operators), we have the same problem and define a constant \vector_to_cblinfun_code\ with more restricted type\ definition vector_to_cblinfun_code :: "'a ell2 \ 'b::one_dim \\<^sub>C\<^sub>L 'a ell2" where [code del,code_abbrev]: "vector_to_cblinfun_code = vector_to_cblinfun" \ \@{attribute code_abbrev} instructs the code generation to replace the rhs \<^term>\vector_to_cblinfun\ by the lhs \<^term>\vector_to_cblinfun_code\ before starting the actual code generation.\ lemma vector_to_cblinfun_code[code]: \ \Code equation for translating a vector into an operation (single-column matrix)\ "mat_of_cblinfun (vector_to_cblinfun_code \) = mat_of_cols (CARD('a)) [vec_of_ell2 \]" for \::"'a::enum ell2" by (simp add: mat_of_cblinfun_vector_to_cblinfun vec_of_ell2_def vector_to_cblinfun_code_def) subsection \Subspaces\ text \In this section, we define code equations for handling subspaces, i.e., values of type \<^typ>\'a ccsubspace\. We choose to computationally represent a subspace by a list of vectors that span the subspace. That is, if \<^term>\vecs\ are vectors (type \<^typ>\complex vec\), \SPAN vecs\ is defined to be their span. Then the code generation can simply represent all subspaces in this form, and we need to define the operations on subspaces in terms of list of vectors (e.g., the closed union of two subspaces would be computed as the concatenation of the two lists, to give one of the simplest examples). To support this, \SPAN\ is declared as a "\code_datatype\". (Not as an abstract datatype like \<^term>\cblinfun_of_mat\/\<^term>\mat_of_cblinfun\ because that would require \SPAN\ to be injective.) Then all code equations for different operations need to be formulated as functions of values of the form \SPAN x\. (E.g., \SPAN x + SPAN y = SPAN (\)\.)\ definition [code del]: "SPAN x = (let n = length (canonical_basis :: 'a::onb_enum list) in ccspan (basis_enum_of_vec ` Set.filter (\v. dim_vec v = n) (set x)) :: 'a ccsubspace)" \ \The SPAN of vectors x, as a \<^type>\ccsubspace\. We filter out vectors of the wrong dimension because \SPAN\ needs to have well-defined behavior even in cases that would not actually occur in an execution.\ code_datatype SPAN text \We first declare code equations for \<^term>\Proj\, i.e., for turning a subspace into a projector. This means, we would need a code equation of the form \mat_of_cblinfun (Proj (SPAN S)) = \\. However, this equation is not accepted by the code generation for reasons we do not understand. But if we define an auxiliary constant \mat_of_cblinfun_Proj_code\ that stands for \mat_of_cblinfun (Proj _)\, define a code equation for \mat_of_cblinfun_Proj_code\, and then define a code equation for \mat_of_cblinfun (Proj S)\ in terms of \mat_of_cblinfun_Proj_code\, Isabelle accepts the code equations.\ definition "mat_of_cblinfun_Proj_code S = mat_of_cblinfun (Proj S)" declare mat_of_cblinfun_Proj_code_def[symmetric, code] lemma mat_of_cblinfun_Proj_code_code[code]: \ \Code equation for computing a projector onto a set S of vectors. We first make the vectors S into an orthonormal basis using the Gram-Schmidt procedure and then compute the projector as the sum of the "butterflies" \x * x*\ of the vectors \x\S\ (done by \<^term>\mk_projector_orthog\).\ "mat_of_cblinfun_Proj_code (SPAN S :: 'a::onb_enum ccsubspace) = (let d = length (canonical_basis :: 'a list) in mk_projector_orthog d (gram_schmidt0 d (filter (\v. dim_vec v = d) S)))" proof - have *: "map_option vec_of_basis_enum (if dim_vec x = length (canonical_basis :: 'a list) then Some (basis_enum_of_vec x :: 'a) else None) = (if dim_vec x = length (canonical_basis :: 'a list) then Some x else None)" for x by auto show ?thesis unfolding SPAN_def mat_of_cblinfun_Proj_code_def using mat_of_cblinfun_Proj_ccspan[where S = "map basis_enum_of_vec (filter (\v. dim_vec v = (length (canonical_basis :: 'a list))) S) :: 'a list"] apply (simp only: Let_def map_filter_map_filter filter_set image_set map_map_filter o_def) unfolding * by (simp add: map_filter_map_filter[symmetric]) qed lemma top_ccsubspace_code[code]: \ \Code equation for \<^term>\top\, the subspace containing everything. Top is represented as the span of the standard basis vectors.\ "(top::'a ccsubspace) = (let n = length (canonical_basis :: 'a::onb_enum list) in SPAN (unit_vecs n))" unfolding SPAN_def apply (simp only: index_unit_vec Let_def map_filter_map_filter filter_set image_set map_map_filter map_filter_map o_def unit_vecs_def) apply (simp add: basis_enum_of_vec_unit_vec) apply (subst nth_image) by (auto simp: ) lemma bot_as_span[code]: \ \Code equation for \<^term>\bot\, the subspace containing everything. Top is represented as the span of the standard basis vectors.\ "(bot::'a::onb_enum ccsubspace) = SPAN []" unfolding SPAN_def by (auto simp: Set.filter_def) lemma sup_spans[code]: \ \Code equation for the join (lub) of two subspaces (union of the generating lists)\ "SPAN A \ SPAN B = SPAN (A @ B)" unfolding SPAN_def by (auto simp: ccspan_union image_Un filter_Un Let_def) text \We do not need an equation for \<^term>\(+)\ because \<^term>\(+)\ is defined in terms of \<^term>\(\)\ (for \<^type>\ccsubspace\), thus the code generation automatically computes \<^term>\(+)\ in terms of the code for \<^term>\(\)\\ definition [code del,code_abbrev]: "Span_code (S::'a::enum ell2 set) = (ccspan S)" \ \A copy of \<^term>\ccspan\ with restricted type. For analogous reasons as \<^term>\cblinfun_apply_code\, see there for explanations\ lemma span_Set_Monad[code]: "Span_code (Set_Monad l) = (SPAN (map vec_of_ell2 l))" \ \Code equation for the span of a finite set. (\<^term>\Set_Monad\ is a datatype constructor that represents sets as lists in the computation.)\ apply (simp add: Span_code_def SPAN_def Let_def) apply (subst Set_filter_unchanged) apply (auto simp add: vec_of_ell2_def)[1] by (metis (no_types, lifting) ell2_of_vec_def image_image map_idI set_map vec_of_ell2_inverse) text \This instantiation defines a code equation for equality tests for \<^type>\ccsubspace\. The actual code for equality tests is given below (lemma \equal_ccsubspace_code\).\ instantiation ccsubspace :: (onb_enum) equal begin definition [code del]: "equal_ccsubspace (A::'a ccsubspace) B = (A=B)" instance apply intro_classes unfolding equal_ccsubspace_def by simp end lemma leq_ccsubspace_code[code]: \ \Code equation for deciding inclusion of one space in another. Uses the constant \<^term>\is_subspace_of_vec_list\ which implements the actual computation by checking for each generator of A whether it is in the span of B (by orthogonal projection onto an orthonormal basis of B which is computed using Gram-Schmidt).\ "SPAN A \ (SPAN B :: 'a::onb_enum ccsubspace) \ (let d = length (canonical_basis :: 'a list) in is_subspace_of_vec_list d (filter (\v. dim_vec v = d) A) (filter (\v. dim_vec v = d) B))" proof - define d A' B' where "d = length (canonical_basis :: 'a list)" and "A' = filter (\v. dim_vec v = d) A" and "B' = filter (\v. dim_vec v = d) B" show ?thesis unfolding SPAN_def d_def[symmetric] filter_set Let_def A'_def[symmetric] B'_def[symmetric] image_set apply (subst ccspan_leq_using_vec) unfolding d_def[symmetric] map_map o_def apply (subst map_cong[where xs=A', OF refl]) apply (rule basis_enum_of_vec_inverse) apply (simp add: A'_def d_def) apply (subst map_cong[where xs=B', OF refl]) apply (rule basis_enum_of_vec_inverse) by (simp_all add: B'_def d_def) qed lemma equal_ccsubspace_code[code]: \ \Code equation for equality test. By checking mutual inclusion (for which we have code by the preceding code equation).\ "HOL.equal (A::_ ccsubspace) B = (A\B \ B\A)" unfolding equal_ccsubspace_def by auto lemma apply_cblinfun_code[code]: \ \Code equation for applying an operator \<^term>\A\ to a subspace. Simply by multiplying each generator with \<^term>\A\\ "A *\<^sub>S SPAN S = (let d = length (canonical_basis :: 'a list) in SPAN (map (mult_mat_vec (mat_of_cblinfun A)) (filter (\v. dim_vec v = d) S)))" for A::"'a::onb_enum \\<^sub>C\<^sub>L'b::onb_enum" proof - define dA dB S' where "dA = length (canonical_basis :: 'a list)" and "dB = length (canonical_basis :: 'b list)" and "S' = filter (\v. dim_vec v = dA) S" have "cblinfun_image A (SPAN S) = A *\<^sub>S ccspan (set (map basis_enum_of_vec S'))" unfolding SPAN_def dA_def[symmetric] Let_def S'_def filter_set by simp also have "\ = ccspan ((\x. basis_enum_of_vec (mat_of_cblinfun A *\<^sub>v vec_of_basis_enum (basis_enum_of_vec x :: 'a))) ` set S')" apply (subst cblinfun_apply_ccspan_using_vec) by (simp add: image_image) also have "\ = ccspan ((\x. basis_enum_of_vec (mat_of_cblinfun A *\<^sub>v x)) ` set S')" apply (subst image_cong[OF refl]) apply (subst basis_enum_of_vec_inverse) by (auto simp add: S'_def dA_def) also have "\ = SPAN (map (mult_mat_vec (mat_of_cblinfun A)) S')" unfolding SPAN_def dB_def[symmetric] Let_def filter_set apply (subst filter_True) by (simp_all add: dB_def mat_of_cblinfun_def image_image) finally show ?thesis unfolding dA_def[symmetric] S'_def[symmetric] Let_def by simp qed definition [code del, code_abbrev]: "range_cblinfun_code A = A *\<^sub>S top" \ \A new constant for the special case of applying an operator to the subspace \<^term>\top\ (i.e., for computing the range of the operator). We do this to be able to give more specialized code for this specific situation. (The generic code for \<^term>\(*\<^sub>S)\ would work but is less efficient because it involves repeated matrix multiplications. @{attribute code_abbrev} makes sure occurrences of \<^term>\A *\<^sub>S top\ are replaced before starting the actual code generation.\ lemma range_cblinfun_code[code]: \ \Code equation for computing the range of an operator \<^term>\A\. Returns the columns of the matrix representation of \<^term>\A\.\ fixes A :: "'a::onb_enum \\<^sub>C\<^sub>L 'b::onb_enum" shows "range_cblinfun_code A = SPAN (cols (mat_of_cblinfun A))" proof - define dA dB where "dA = length (canonical_basis :: 'a list)" and "dB = length (canonical_basis :: 'b list)" have carrier_A: "mat_of_cblinfun A \ carrier_mat dB dA" unfolding mat_of_cblinfun_def dA_def dB_def by simp have "range_cblinfun_code A = A *\<^sub>S SPAN (unit_vecs dA)" unfolding range_cblinfun_code_def by (metis dA_def top_ccsubspace_code) also have "\ = SPAN (map (\i. mat_of_cblinfun A *\<^sub>v unit_vec dA i) [0.. = SPAN (map (\x. mat_of_cblinfun A *\<^sub>v col (1\<^sub>m dA) x) [0.. = SPAN (map (col (mat_of_cblinfun A * 1\<^sub>m dA)) [0.. = SPAN (cols (mat_of_cblinfun A))" unfolding cols_def dA_def[symmetric] apply (subst right_mult_one_mat[OF carrier_A]) using carrier_A by blast finally show ?thesis by - qed lemma uminus_Span_code[code]: "- X = range_cblinfun_code (id_cblinfun - Proj X)" \ \Code equation for the orthogonal complement of a subspace \<^term>\X\. Computed as the range of one minus the projector on \<^term>\X\\ unfolding range_cblinfun_code_def by (metis Proj_ortho_compl Proj_range) lemma kernel_code[code]: \ \Computes the kernel of an operator \<^term>\A\. This is implemented using the existing functions for transforming a matrix into row echelon form (\<^term>\gauss_jordan_single\) and for computing a basis of the kernel of such a matrix (\<^term>\find_base_vectors\)\ "kernel A = SPAN (find_base_vectors (gauss_jordan_single (mat_of_cblinfun A)))" for A::"('a::onb_enum,'b::onb_enum) cblinfun" proof - define dA dB Am Ag base where "dA = length (canonical_basis :: 'a list)" and "dB = length (canonical_basis :: 'b list)" and "Am = mat_of_cblinfun A" and "Ag = gauss_jordan_single Am" and "base = find_base_vectors Ag" interpret complex_vec_space dA. have Am_carrier: "Am \ carrier_mat dB dA" unfolding Am_def mat_of_cblinfun_def dA_def dB_def by simp have row_echelon: "row_echelon_form Ag" unfolding Ag_def using Am_carrier refl by (rule gauss_jordan_single) have Ag_carrier: "Ag \ carrier_mat dB dA" unfolding Ag_def using Am_carrier refl by (rule gauss_jordan_single(2)) have base_carrier: "set base \ carrier_vec dA" unfolding base_def using find_base_vectors(1)[OF row_echelon Ag_carrier] using Ag_carrier mat_kernel_def by blast interpret k: kernel dB dA Ag apply standard using Ag_carrier by simp have basis_base: "kernel.basis dA Ag (set base)" using row_echelon Ag_carrier unfolding base_def by (rule find_base_vectors(3)) have "space_as_set (SPAN base) = space_as_set (ccspan (basis_enum_of_vec ` set base :: 'a set))" unfolding SPAN_def dA_def[symmetric] Let_def filter_set apply (subst filter_True) using base_carrier by auto also have "\ = cspan (basis_enum_of_vec ` set base)" apply transfer apply (subst closure_finite_cspan) by simp_all also have "\ = basis_enum_of_vec ` span (set base)" apply (subst basis_enum_of_vec_span) using base_carrier dA_def by auto also have "\ = basis_enum_of_vec ` mat_kernel Ag" using basis_base k.Ker.basis_def k.span_same by auto also have "\ = basis_enum_of_vec ` {v \ carrier_vec dA. Ag *\<^sub>v v = 0\<^sub>v dB}" apply (rule arg_cong[where f="\x. basis_enum_of_vec ` x"]) unfolding mat_kernel_def using Ag_carrier by simp also have "\ = basis_enum_of_vec ` {v \ carrier_vec dA. Am *\<^sub>v v = 0\<^sub>v dB}" using gauss_jordan_single(1)[OF Am_carrier Ag_def[symmetric]] by auto also have "\ = {w. A *\<^sub>V w = 0}" proof - have "basis_enum_of_vec ` {v \ carrier_vec dA. Am *\<^sub>v v = 0\<^sub>v dB} = basis_enum_of_vec ` {v \ carrier_vec dA. A *\<^sub>V basis_enum_of_vec v = 0}" apply (rule arg_cong[where f="\t. basis_enum_of_vec ` t"]) apply (rule Collect_cong) apply (simp add: Am_def) by (metis Am_carrier Am_def carrier_matD(2) carrier_vecD dB_def mat_carrier mat_of_cblinfun_def mat_of_cblinfun_cblinfun_apply vec_of_basis_enum_inverse basis_enum_of_vec_inverse vec_of_basis_enum_zero) also have "\ = {w \ basis_enum_of_vec ` carrier_vec dA. A *\<^sub>V w = 0}" apply (subst Compr_image_eq[symmetric]) by simp also have "\ = {w. A *\<^sub>V w = 0}" apply auto by (metis (no_types, lifting) Am_carrier Am_def carrier_matD(2) carrier_vec_dim_vec dim_vec_of_basis_enum' image_iff mat_carrier mat_of_cblinfun_def vec_of_basis_enum_inverse) finally show ?thesis by - qed also have "\ = space_as_set (kernel A)" apply transfer by auto finally have "SPAN base = kernel A" by (simp add: space_as_set_inject) then show ?thesis by (simp add: base_def Ag_def Am_def) qed lemma inf_ccsubspace_code[code]: \ \Code equation for intersection of subspaces. Reduced to orthogonal complement and sum of subspaces for which we already have code equations.\ "(A::'a::onb_enum ccsubspace) \ B = - (- A \ - B)" by (subst ortho_involution[symmetric], subst compl_inf, simp) lemma Sup_ccsubspace_code[code]: \ \Supremum (sum) of a set of subspaces. Implemented by repeated pairwise sum.\ "Sup (Set_Monad l :: 'a::onb_enum ccsubspace set) = fold sup l bot" unfolding Set_Monad_def by (simp add: Sup_set_fold) lemma Inf_ccsubspace_code[code]: \ \Infimum (intersection) of a set of subspaces. Implemented by the orthogonal complement of the supremum.\ "Inf (Set_Monad l :: 'a::onb_enum ccsubspace set) = - Sup (Set_Monad (map uminus l))" unfolding Set_Monad_def apply (induction l) by auto subsection \Miscellanea\ text \This is a hack to circumvent a bug in the code generation. The automatically generated code for the class \<^class>\uniformity\ has a type that is different from what the generated code later assumes, leading to compilation errors (in ML at least) in any expression involving \<^typ>\_ ell2\ (even if the constant \<^const>\uniformity\ is not actually used). The fragment below circumvents this by forcing Isabelle to use the right type. (The logically useless fragment "\let x = ((=)::'a\_\_)\" achieves this.)\ lemma uniformity_ell2_code[code]: "(uniformity :: ('a ell2 * _) filter) = Filter.abstract_filter (%_. Code.abort STR ''no uniformity'' (%_. let x = ((=)::'a\_\_) in uniformity))" by simp text \Code equation for \<^term>\UNIV\. It is now implemented via type class \<^class>\enum\ (which provides a list of all values).\ declare [[code drop: UNIV]] declare enum_class.UNIV_enum[code] text \Setup for code generation involving sets of \<^type>\ell2\/\<^type>\ccsubspace\. This configures to use lists for representing sets in code.\ derive (eq) ceq ccsubspace derive (no) ccompare ccsubspace derive (monad) set_impl ccsubspace derive (eq) ceq ell2 derive (no) ccompare ell2 derive (monad) set_impl ell2 +unbundle no_lattice_syntax unbundle no_jnf_notation unbundle no_cblinfun_notation end diff --git a/thys/Complex_Bounded_Operators/Cblinfun_Code_Examples.thy b/thys/Complex_Bounded_Operators/Cblinfun_Code_Examples.thy --- a/thys/Complex_Bounded_Operators/Cblinfun_Code_Examples.thy +++ b/thys/Complex_Bounded_Operators/Cblinfun_Code_Examples.thy @@ -1,119 +1,120 @@ section \\Cblinfun_Code_Examples\ -- Examples and test cases for code generation\ theory Cblinfun_Code_Examples imports "Complex_Bounded_Operators.Extra_Pretty_Code_Examples" Jordan_Normal_Form.Matrix_Impl "HOL-Library.Code_Target_Numeral" Cblinfun_Code begin hide_const (open) Order.bottom Order.top no_notation Lattice.join (infixl "\\" 65) no_notation Lattice.meet (infixl "\\" 70) +unbundle lattice_syntax unbundle cblinfun_notation section \Examples\ subsection \Operators\ value "id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "1 :: unit ell2 \\<^sub>C\<^sub>L unit ell2" value "id_cblinfun + id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "0 :: (bool ell2 \\<^sub>C\<^sub>L Enum.finite_3 ell2)" value "- id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "id_cblinfun - id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "classical_operator (\b. Some (\ b))" value "id_cblinfun = (0 :: bool ell2 \\<^sub>C\<^sub>L bool ell2)" value "2 *\<^sub>R id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "imaginary_unit *\<^sub>C id_cblinfun :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "id_cblinfun o\<^sub>C\<^sub>L 0 :: bool ell2 \\<^sub>C\<^sub>L bool ell2" value "id_cblinfun* :: bool ell2 \\<^sub>C\<^sub>L bool ell2" subsection \Vectors\ value "0 :: bool ell2" value "1 :: unit ell2" value "ket False" value "2 *\<^sub>C ket False" value "2 *\<^sub>R ket False" value "ket True + ket False" value "ket True - ket True" value "ket True = ket True" value "- ket True" value "cinner (ket True) (ket True)" value "norm (ket True)" value "ket () * ket ()" value "1 :: unit ell2" value "(1::unit ell2) * (1::unit ell2)" subsection \Vector/Matrix\ value "id_cblinfun *\<^sub>V ket True" value \vector_to_cblinfun (ket True) :: unit ell2 \\<^sub>C\<^sub>L _\ subsection \Subspaces\ value "ccspan {ket False}" value "Proj (ccspan {ket False})" value "top :: bool ell2 ccsubspace" value "bot :: bool ell2 ccsubspace" value "0 :: bool ell2 ccsubspace" value "ccspan {ket False} \ ccspan {ket True}" value "ccspan {ket False} + ccspan {ket True}" value "ccspan {ket False} \ ccspan {ket True}" value "id_cblinfun *\<^sub>S ccspan {ket False}" value "id_cblinfun *\<^sub>S (top :: bool ell2 ccsubspace)" (* Special case, using range_cblinfun_code for efficiency *) value "- ccspan {ket False}" value "ccspan {ket False, ket True} = top" value "ccspan {ket False} \ ccspan {ket True}" value "cblinfun_image id_cblinfun (ccspan {ket True})" value "kernel id_cblinfun :: bool ell2 ccsubspace" value "eigenspace 1 id_cblinfun :: bool ell2 ccsubspace" value "Inf {ccspan {ket False}, top}" value "Sup {ccspan {ket False}, top}" end diff --git a/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy b/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy --- a/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy +++ b/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy @@ -1,2981 +1,2984 @@ section \\Complex_Bounded_Linear_Function\ -- Complex bounded linear functions (bounded operators)\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Bounded_Linear_Function imports Complex_Inner_Product One_Dimensional_Spaces Banach_Steinhaus.Banach_Steinhaus "HOL-Types_To_Sets.Types_To_Sets" Complex_Bounded_Linear_Function0 begin +unbundle lattice_syntax + subsection \Misc basic facts and declarations\ notation cblinfun_apply (infixr "*\<^sub>V" 70) lemma id_cblinfun_apply[simp]: "id_cblinfun *\<^sub>V \ = \" apply transfer by simp lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" apply transfer by (simp add: clinear_continuous_at) declare cblinfun.scaleC_left[simp] lemma cblinfun_apply_clinear[simp]: \clinear (cblinfun_apply A)\ using bounded_clinear.axioms(1) cblinfun_apply by blast lemma cblinfun_cinner_eqI: fixes A B :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ assumes \\\. cinner \ (A *\<^sub>V \) = cinner \ (B *\<^sub>V \)\ shows \A = B\ proof - define C where \C = A - B\ have C0[simp]: \cinner \ (C \) = 0\ for \ by (simp add: C_def assms cblinfun.diff_left cinner_diff_right) { fix f g \ have \0 = cinner (f + \ *\<^sub>C g) (C *\<^sub>V (f + \ *\<^sub>C g))\ by (simp add: cinner_diff_right minus_cblinfun.rep_eq) also have \\ = \ *\<^sub>C cinner f (C g) + cnj \ *\<^sub>C cinner g (C f)\ by (smt (z3) C0 add.commute add.right_neutral cblinfun.add_right cblinfun.scaleC_right cblinfun_cinner_right.rep_eq cinner_add_left cinner_scaleC_left complex_scaleC_def) finally have \\ *\<^sub>C cinner f (C g) = - cnj \ *\<^sub>C cinner g (C f)\ by (simp add: eq_neg_iff_add_eq_0) } then have \cinner f (C g) = 0\ for f g by (metis complex_cnj_i complex_cnj_one complex_vector.scale_cancel_right complex_vector.scale_left_imp_eq equation_minus_iff i_squared mult_eq_0_iff one_neq_neg_one) then have \C g = 0\ for g using cinner_eq_zero_iff by blast then have \C = 0\ by (simp add: cblinfun_eqI) then show \A = B\ using C_def by auto qed lemma id_cblinfun_not_0[simp]: \(id_cblinfun :: 'a::{complex_normed_vector, not_singleton} \\<^sub>C\<^sub>L _) \ 0\ by (metis (full_types) Extra_General.UNIV_not_singleton cblinfun.zero_left cblinfun_id_cblinfun_apply ex_norm1 norm_zero one_neq_zero) lemma cblinfun_norm_geqI: assumes \norm (f *\<^sub>V x) / norm x \ K\ shows \norm f \ K\ using assms apply transfer by (smt (z3) bounded_clinear.bounded_linear le_onorm) (* This lemma is proven in Complex_Bounded_Linear_Function0 but we add the [simp] only here because we try to keep Complex_Bounded_Linear_Function0 as close to Bounded_Linear_Function as possible. *) declare scaleC_conv_of_complex[simp] lemma cblinfun_eq_0_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = 0" shows \F *\<^sub>V x = 0\ apply (rule complex_vector.linear_eq_0_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = G *\<^sub>V s" shows \F *\<^sub>V x = G *\<^sub>V x\ apply (rule complex_vector.linear_eq_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_0_on_UNIV_span: fixes basis::\'a::complex_normed_vector set\ assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = 0" shows \F = 0\ by (metis cblinfun_eq_0_on_span UNIV_I assms cblinfun.zero_left cblinfun_eqI) lemma cblinfun_eq_on_UNIV_span: fixes basis::"'a::complex_normed_vector set" and \::"'a \ 'b::complex_normed_vector" assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = G *\<^sub>V s" shows \F = G\ proof- have "F - G = 0" apply (rule cblinfun_eq_0_on_UNIV_span[where basis=basis]) using assms by (auto simp add: cblinfun.diff_left) thus ?thesis by simp qed lemma cblinfun_eq_on_canonical_basis: fixes f g::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = g *\<^sub>V u" shows "f = g" apply (rule cblinfun_eq_on_UNIV_span[where basis=basis]) using assms is_generator_set is_cindependent_set by auto lemma cblinfun_eq_0_on_canonical_basis: fixes f ::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = 0" shows "f = 0" by (simp add: assms cblinfun_eq_on_canonical_basis) lemma cinner_canonical_basis_eq_0: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = 0" shows "F = 0" proof- have "F *\<^sub>V u = 0" if "u\basisA" for u proof- have "\v. v\basisB \ \v, F *\<^sub>V u\ = 0" by (simp add: assms(3) that) moreover have "(\v. v\basisB \ \v, x\ = 0) \ x = 0" for x proof- assume r1: "\v. v\basisB \ \v, x\ = 0" have "\v, x\ = 0" for v proof- have "cspan basisB = UNIV" using basisB_def is_generator_set by auto hence "v \ cspan basisB" by (smt iso_tuple_UNIV_I) hence "\t s. v = (\a\t. s a *\<^sub>C a) \ finite t \ t \ basisB" using complex_vector.span_explicit by (smt mem_Collect_eq) then obtain t s where b1: "v = (\a\t. s a *\<^sub>C a)" and b2: "finite t" and b3: "t \ basisB" by blast have "\v, x\ = \(\a\t. s a *\<^sub>C a), x\" by (simp add: b1) also have "\ = (\a\t. \s a *\<^sub>C a, x\)" using cinner_sum_left by blast also have "\ = (\a\t. cnj (s a) * \a, x\)" by auto also have "\ = 0" using b3 r1 subsetD by force finally show ?thesis by simp qed thus ?thesis by (simp add: \\v. \v, x\ = 0\ cinner_extensionality) qed ultimately show ?thesis by simp qed thus ?thesis using basisA_def cblinfun_eq_0_on_canonical_basis by auto qed lemma cinner_canonical_basis_eq: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = \v, G *\<^sub>V u\" shows "F = G" proof- define H where "H = F - G" have "\u v. u\basisA \ v\basisB \ \v, H *\<^sub>V u\ = 0" unfolding H_def by (simp add: assms(3) cinner_diff_right minus_cblinfun.rep_eq) hence "H = 0" by (simp add: basisA_def basisB_def cinner_canonical_basis_eq_0) thus ?thesis unfolding H_def by simp qed lemma cinner_canonical_basis_eq': defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \F *\<^sub>V u, v\ = \G *\<^sub>V u, v\" shows "F = G" using cinner_canonical_basis_eq assms by (metis cinner_commute') lemma cblinfun_norm_approx_witness: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ > 0\ shows \\\. norm (A *\<^sub>V \) \ norm A - \ \ norm \ = 1\ proof (transfer fixing: \) fix A :: \'a \ 'b\ assume [simp]: \bounded_clinear A\ have \\y\{norm (A x) |x. norm x = 1}. y > \ {norm (A x) |x. norm x = 1} - \\ apply (rule Sup_real_close) using assms by (auto simp: ex_norm1 bounded_clinear.bounded_linear bdd_above_norm_f) also have \\ {norm (A x) |x. norm x = 1} = onorm A\ by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear onorm_sphere) finally show \\\. onorm A - \ \ norm (A \) \ norm \ = 1\ by force qed lemma cblinfun_norm_approx_witness_mult: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ < 1\ shows \\\. norm (A *\<^sub>V \) \ norm A * \ \ norm \ = 1\ proof (cases \norm A = 0\) case True then show ?thesis apply auto by (simp add: ex_norm1) next case False then have \(1 - \) * norm A > 0\ using assms by fastforce then obtain \ where geq: \norm (A *\<^sub>V \) \ norm A - ((1 - \) * norm A)\ and \norm \ = 1\ using cblinfun_norm_approx_witness by blast have \norm A * \ = norm A - (1 - \) * norm A\ by (simp add: mult.commute right_diff_distrib') also have \\ \ norm (A *\<^sub>V \)\ by (rule geq) finally show ?thesis using \norm \ = 1\ by auto qed lemma cblinfun_to_CARD_1_0[simp]: \(A :: _ \\<^sub>C\<^sub>L _::CARD_1) = 0\ apply (rule cblinfun_eqI) by auto lemma cblinfun_from_CARD_1_0[simp]: \(A :: _::CARD_1 \\<^sub>C\<^sub>L _) = 0\ apply (rule cblinfun_eqI) apply (subst CARD_1_vec_0) by auto lemma cblinfun_cspan_UNIV: fixes basis :: \('a::{complex_normed_vector,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_normed_vector) set\ and basisA :: \'a set\ and basisB :: \'b set\ assumes \cspan basisA = UNIV\ and \cspan basisB = UNIV\ assumes basis: \\a b. a\basisA \ b\basisB \ \F\basis. \a'\basisA. F *\<^sub>V a' = (if a'=a then b else 0)\ shows \cspan basis = UNIV\ proof - obtain basisA' where \basisA' \ basisA\ and \cindependent basisA'\ and \cspan basisA' = UNIV\ by (metis assms(1) complex_vector.maximal_independent_subset complex_vector.span_eq top_greatest) then have [simp]: \finite basisA'\ by (simp add: cindependent_cfinite_dim_finite) have basis': \\a b. a\basisA' \ b\basisB \ \F\basis. \a'\basisA'. F *\<^sub>V a' = (if a'=a then b else 0)\ using basis \basisA' \ basisA\ by fastforce obtain F where F: \F a b \ basis \ F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using basis' by metis then have F_apply: \F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' using that by auto have F_basis: \F a b \ basis\ if \a\basisA'\ \b\basisB\ for a b using that F by auto have b_span: \\G\cspan {F a b|b. b\basisB}. \a'\basisA'. G *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ for a b proof - from \cspan basisB = UNIV\ obtain r t where \finite t\ and \t \ basisB\ and b_lincom: \b = (\a\t. r a *\<^sub>C a)\ unfolding complex_vector.span_alt apply atomize_elim by blast define G where \G = (\i\t. r i *\<^sub>C F a i)\ have \G \ cspan {F a b|b. b\basisB}\ using \finite t\ \t \ basisB\ unfolding G_def by (smt (verit, ccfv_threshold) complex_vector.span_base complex_vector.span_scale complex_vector.span_sum mem_Collect_eq subset_eq) moreover have \G *\<^sub>V a' = (if a'=a then b else 0)\ if \a'\basisA'\ for a' apply (cases \a'=a\) using \t \ basisB\ \a\basisA'\ \a'\basisA'\ by (auto simp: b_lincom G_def cblinfun.sum_left F_apply intro!: sum.neutral sum.cong) ultimately show ?thesis by blast qed have a_span: \cspan (\a\basisA'. cspan {F a b|b. b\basisB}) = UNIV\ proof (intro equalityI subset_UNIV subsetI, rename_tac H) fix H obtain G where G: \G a b \ cspan {F a b|b. b\basisB} \ G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using b_span by blast then have G_cspan: \G a b \ cspan {F a b|b. b\basisB}\ if \a\basisA'\ for a b using that by auto from G have G: \G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' using that by auto define H' where \H' = (\a\basisA'. G a (H *\<^sub>V a))\ have \H' \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ unfolding H'_def using G_cspan by (smt (verit, del_insts) UN_iff complex_vector.span_clauses(1) complex_vector.span_sum) moreover have \H' = H\ using \cspan basisA' = UNIV\ apply (rule cblinfun_eq_on_UNIV_span) apply (auto simp: H'_def cblinfun.sum_left) apply (subst sum_single) by (auto simp: G) ultimately show \H \ cspan (\a\basisA'. cspan {F a b |b. b \ basisB})\ by simp qed moreover have \cspan basis \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ using F_basis by (smt (z3) UN_subset_iff complex_vector.span_alt complex_vector.span_minimal complex_vector.subspace_span mem_Collect_eq subset_iff) ultimately show \cspan basis = UNIV\ by auto qed instance cblinfun :: (\{cfinite_dim,complex_normed_vector}\, \{cfinite_dim,complex_normed_vector}\) cfinite_dim proof intro_classes obtain basisA :: \'a set\ where [simp]: \cspan basisA = UNIV\ \cindependent basisA\ \finite basisA\ using finite_basis by blast obtain basisB :: \'b set\ where [simp]: \cspan basisB = UNIV\ \cindependent basisB\ \finite basisB\ using finite_basis by blast define f where \f a b = cconstruct basisA (\x. if x=a then b else 0)\ for a :: 'a and b :: 'b have f_a: \f a b a = b\ if \a : basisA\ for a b by (simp add: complex_vector.construct_basis f_def that) have f_not_a: \f a b c = 0\ if \a : basisA\ and \c : basisA\ and \a \ c\for a b c using that by (simp add: complex_vector.construct_basis f_def) define F where \F a b = CBlinfun (f a b)\ for a b have \clinear (f a b)\ for a b by (auto intro: complex_vector.linear_construct simp: f_def) then have \bounded_clinear (f a b)\ for a b by auto then have F_apply: \cblinfun_apply (F a b) = f a b\ for a b by (simp add: F_def bounded_clinear_CBlinfun_apply) define basis where \basis = {F a b| a b. a\basisA \ b\basisB}\ have \cspan basis = UNIV\ apply (rule cblinfun_cspan_UNIV[where basisA=basisA and basisB=basisB]) apply (auto simp: basis_def) by (metis F_apply f_a f_not_a) moreover have \finite basis\ unfolding basis_def apply (rule finite_image_set2) by auto ultimately show \\S :: ('a \\<^sub>C\<^sub>L 'b) set. finite S \ cspan S = UNIV\ by auto qed subsection \Relationship to real bounded operators (\<^typ>\_ \\<^sub>L _\)\ instantiation blinfun :: (real_normed_vector, complex_normed_vector) "complex_normed_vector" begin lift_definition scaleC_blinfun :: \complex \ ('a::real_normed_vector, 'b::complex_normed_vector) blinfun \ ('a, 'b) blinfun\ is \\ c::complex. \ f::'a\'b. (\ x. c *\<^sub>C (f x) )\ proof fix c::complex and f :: \'a\'b\ and b1::'a and b2::'a assume \bounded_linear f\ show \c *\<^sub>C f (b1 + b2) = c *\<^sub>C f b1 + c *\<^sub>C f b2\ by (simp add: \bounded_linear f\ linear_simps scaleC_add_right) fix c::complex and f :: \'a\'b\ and b::'a and r::real assume \bounded_linear f\ show \c *\<^sub>C f (r *\<^sub>R b) = r *\<^sub>R (c *\<^sub>C f b)\ by (simp add: \bounded_linear f\ linear_simps(5) scaleR_scaleC) fix c::complex and f :: \'a\'b\ assume \bounded_linear f\ have \\ K. \ x. norm (f x) \ norm x * K\ using \bounded_linear f\ by (simp add: bounded_linear.bounded) then obtain K where \\ x. norm (f x) \ norm x * K\ by blast have \cmod c \ 0\ by simp hence \\ x. (cmod c) * norm (f x) \ (cmod c) * norm x * K\ using \\ x. norm (f x) \ norm x * K\ by (metis ordered_comm_semiring_class.comm_mult_left_mono vector_space_over_itself.scale_scale) moreover have \norm (c *\<^sub>C f x) = (cmod c) * norm (f x)\ for x by simp ultimately show \\K. \x. norm (c *\<^sub>C f x) \ norm x * K\ by (metis ab_semigroup_mult_class.mult_ac(1) mult.commute) qed instance proof have "r *\<^sub>R x = complex_of_real r *\<^sub>C x" for x :: "('a, 'b) blinfun" and r apply transfer by (simp add: scaleR_scaleC) thus "((*\<^sub>R) r::'a \\<^sub>L 'b \ _) = (*\<^sub>C) (complex_of_real r)" for r by auto show "a *\<^sub>C (x + y) = a *\<^sub>C x + a *\<^sub>C y" for a :: complex and x y :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_right) show "(a + b) *\<^sub>C x = a *\<^sub>C x + b *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_left) show "a *\<^sub>C b *\<^sub>C x = (a * b) *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by simp have \1 *\<^sub>C f x = f x\ for f :: \'a\'b\ and x by auto thus "1 *\<^sub>C x = x" for x :: "'a \\<^sub>L 'b" by (simp add: scaleC_blinfun.rep_eq blinfun_eqI) have \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ if \bounded_linear f\ for f :: \'a \ 'b\ and a :: complex proof- have \cmod a \ 0\ by simp have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \\ x. (cmod a) *(\ ereal ((norm (f x)) / (norm x)) \) \ (cmod a) * K\ using \cmod a \ 0\ by (metis abs_ereal.simps(1) abs_ereal_pos abs_pos ereal_mult_left_mono times_ereal.simps(1)) hence \\ x. (\ ereal ((cmod a) * (norm (f x)) / (norm x)) \) \ (cmod a) * K\ by simp hence \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have p1: \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ using \\ x. \ ereal (cmod a * (norm (f x)) / (norm x)) \ \ cmod a * K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) have p2: \\i. i \ UNIV \ 0 \ ereal (cmod a * norm (f i) / norm i)\ by simp hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ (SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\)\ using \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I p2 abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ cmod a * K\ using \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ by simp hence \\ ( SUP i\UNIV::'a set. ereal ((\ x. (cmod a) * (norm (f x)) / norm x) i)) \ \ \\ by auto hence w2: \( SUP i\UNIV::'a set. ereal ((\ x. cmod a * (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. cmod a * (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) have \(UNIV::('a set)) \ {}\ by simp moreover have \\ i. i \ (UNIV::('a set)) \ (\ x. (norm (f x)) / norm x :: ereal) i \ 0\ by simp moreover have \cmod a \ 0\ by simp ultimately have \(SUP i\(UNIV::('a set)). ((cmod a)::ereal) * (\ x. (norm (f x)) / norm x :: ereal) i ) = ((cmod a)::ereal) * ( SUP i\(UNIV::('a set)). (\ x. (norm (f x)) / norm x :: ereal) i )\ by (simp add: Sup_ereal_mult_left') hence \(SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp hence z1: \real_of_ereal ( (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) ) = real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) )\ by simp have z2: \real_of_ereal (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = (SUP x. cmod a * (norm (f x) / norm x))\ using w2 by auto have \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = (cmod a) * real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp moreover have \real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) ) = ( SUP x. ((norm (f x)) / norm x) )\ proof- have \\ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) \ \ \\ proof- have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ using \\ x. \ ereal ((norm (f x)) / (norm x)) \ \ K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ (SUP x. \ereal ((norm (f x)) / (norm x))\)\ using \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I \\i. i \ UNIV \ 0 \ ereal (norm (f i) / norm i)\ abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ K\ using \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ by simp thus ?thesis by auto qed hence \ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) thus ?thesis by simp qed have z3: \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = cmod a * (SUP x. norm (f x) / norm x)\ by (simp add: \real_of_ereal (SUP x. ereal (norm (f x) / norm x)) = (SUP x. norm (f x) / norm x)\) hence w1: \(SUP x. cmod a * (norm (f x) / norm x)) = cmod a * (SUP x. norm (f x) / norm x)\ using z1 z2 by linarith have v1: \onorm (\x. a *\<^sub>C f x) = (SUP x. norm (a *\<^sub>C f x) / norm x)\ by (simp add: onorm_def) have v2: \(SUP x. norm (a *\<^sub>C f x) / norm x) = (SUP x. ((cmod a) * norm (f x)) / norm x)\ by simp have v3: \(SUP x. ((cmod a) * norm (f x)) / norm x) = (SUP x. (cmod a) * ((norm (f x)) / norm x))\ by simp have v4: \(SUP x. (cmod a) * ((norm (f x)) / norm x)) = (cmod a) * (SUP x. ((norm (f x)) / norm x))\ using w1 by blast show \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ using v1 v2 v3 v4 by (metis (mono_tags, lifting) onorm_def) qed thus \norm (a *\<^sub>C x) = cmod a * norm x\ for a::complex and x::\('a, 'b) blinfun\ apply transfer by blast qed end (* We do not have clinear_blinfun_compose_right *) lemma clinear_blinfun_compose_left: \clinear (\x. blinfun_compose x y)\ by (auto intro!: clinearI simp: blinfun_eqI scaleC_blinfun.rep_eq bounded_bilinear.add_left bounded_bilinear_blinfun_compose) instantiation blinfun :: (real_normed_vector, cbanach) "cbanach" begin instance.. end lemma blinfun_compose_assoc: "(A o\<^sub>L B) o\<^sub>L C = A o\<^sub>L (B o\<^sub>L C)" by (simp add: blinfun_eqI) lift_definition blinfun_of_cblinfun::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a \\<^sub>L 'b\ is "id" apply transfer by (simp add: bounded_clinear.bounded_linear) lift_definition blinfun_cblinfun_eq :: \'a \\<^sub>L 'b \ 'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ bool\ is "(=)" . lemma blinfun_cblinfun_eq_bi_unique[transfer_rule]: \bi_unique blinfun_cblinfun_eq\ unfolding bi_unique_def apply transfer by auto lemma blinfun_cblinfun_eq_right_total[transfer_rule]: \right_total blinfun_cblinfun_eq\ unfolding right_total_def apply transfer by (simp add: bounded_clinear.bounded_linear) named_theorems cblinfun_blinfun_transfer lemma cblinfun_blinfun_transfer_0[cblinfun_blinfun_transfer]: "blinfun_cblinfun_eq (0::(_,_) blinfun) (0::(_,_) cblinfun)" apply transfer by simp lemma cblinfun_blinfun_transfer_plus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (+) (+)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_minus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (-) (-)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_uminus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (uminus) (uminus)" unfolding rel_fun_def apply transfer by auto definition "real_complex_eq r c \ complex_of_real r = c" lemma bi_unique_real_complex_eq[transfer_rule]: \bi_unique real_complex_eq\ unfolding real_complex_eq_def bi_unique_def by auto lemma left_total_real_complex_eq[transfer_rule]: \left_total real_complex_eq\ unfolding real_complex_eq_def left_total_def by auto lemma cblinfun_blinfun_transfer_scaleC[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(real_complex_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (scaleR) (scaleC)" unfolding rel_fun_def apply transfer by (simp add: real_complex_eq_def scaleR_scaleC) lemma cblinfun_blinfun_transfer_CBlinfun[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(eq_onp bounded_clinear ===> blinfun_cblinfun_eq) Blinfun CBlinfun" unfolding rel_fun_def blinfun_cblinfun_eq.rep_eq eq_onp_def by (auto simp: CBlinfun_inverse Blinfun_inverse bounded_clinear.bounded_linear) lemma cblinfun_blinfun_transfer_norm[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=)) norm norm" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_dist[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> (=)) dist dist" unfolding rel_fun_def dist_norm apply transfer by auto lemma cblinfun_blinfun_transfer_sgn[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) sgn sgn" unfolding rel_fun_def sgn_blinfun_def sgn_cblinfun_def apply transfer by (auto simp: scaleR_scaleC) lemma cblinfun_blinfun_transfer_Cauchy[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> (=)) Cauchy Cauchy" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding Cauchy_def by transfer_prover qed lemma cblinfun_blinfun_transfer_tendsto[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> blinfun_cblinfun_eq ===> (=) ===> (=)) tendsto tendsto" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding tendsto_iff by transfer_prover qed lemma cblinfun_blinfun_transfer_compose[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (o\<^sub>L) (o\<^sub>C\<^sub>L)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_apply[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=) ===> (=)) blinfun_apply cblinfun_apply" unfolding rel_fun_def apply transfer by auto lemma blinfun_of_cblinfun_inj: \blinfun_of_cblinfun f = blinfun_of_cblinfun g \ f = g\ by (metis cblinfun_apply_inject blinfun_of_cblinfun.rep_eq) lemma blinfun_of_cblinfun_inv: assumes "\c. \x. f *\<^sub>v (c *\<^sub>C x) = c *\<^sub>C (f *\<^sub>v x)" shows "\g. blinfun_of_cblinfun g = f" using assms proof transfer show "\g\Collect bounded_clinear. id g = f" if "bounded_linear f" and "\c x. f (c *\<^sub>C x) = c *\<^sub>C f x" for f :: "'a \ 'b" using that bounded_linear_bounded_clinear by auto qed lemma blinfun_of_cblinfun_zero: \blinfun_of_cblinfun 0 = 0\ apply transfer by simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ apply transfer by auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_norm: fixes f::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ shows \norm f = norm (blinfun_of_cblinfun f)\ apply transfer by auto subsection \Composition\ lemma blinfun_of_cblinfun_cblinfun_compose: fixes f::\'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector\ and g::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b\ shows \blinfun_of_cblinfun (f o\<^sub>C\<^sub>L g) = (blinfun_of_cblinfun f) o\<^sub>L (blinfun_of_cblinfun g)\ apply transfer by auto lemma cblinfun_compose_assoc: shows "(A o\<^sub>C\<^sub>L B) o\<^sub>C\<^sub>L C = A o\<^sub>C\<^sub>L (B o\<^sub>C\<^sub>L C)" by (metis (no_types, lifting) cblinfun_apply_inject fun.map_comp cblinfun_compose.rep_eq) lemma cblinfun_compose_zero_right[simp]: "U o\<^sub>C\<^sub>L 0 = 0" using bounded_cbilinear.zero_right bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_zero_left[simp]: "0 o\<^sub>C\<^sub>L U = 0" using bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_scaleC_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>C A) o\<^sub>C\<^sub>L B = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ by (simp add: bounded_cbilinear.scaleC_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_scaleR_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>R A) o\<^sub>C\<^sub>L B = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_scaleC_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>C B) = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ apply transfer by (auto intro!: ext bounded_clinear.clinear complex_vector.linear_scale) lemma cblinfun_compose_scaleR_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>R B) = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_id_right[simp]: shows "U o\<^sub>C\<^sub>L id_cblinfun = U" apply transfer by auto lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" apply transfer by auto lemma cblinfun_eq_on: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \t \ closure (cspan G)\ shows "A *\<^sub>V t = B *\<^sub>V t" using assms apply transfer using bounded_clinear_eq_on by blast lemma cblinfun_eq_gen_eqI: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \ccspan G = \\ shows "A = B" apply (rule cblinfun_eqI) apply (rule cblinfun_eq_on[where G=G]) using assms apply auto by (metis ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) lemma cblinfun_compose_add_left: \(a + b) o\<^sub>C\<^sub>L c = (a o\<^sub>C\<^sub>L c) + (b o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_add_right: \a o\<^sub>C\<^sub>L (b + c) = (a o\<^sub>C\<^sub>L b) + (a o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) lemma cbilinear_cblinfun_compose[simp]: "cbilinear cblinfun_compose" by (auto intro!: clinearI simp add: cbilinear_def bounded_cbilinear.add_left bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) subsection \Adjoint\ lift_definition adj :: "'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ 'b \\<^sub>C\<^sub>L 'a" ("_*" [99] 100) is cadjoint by (fact cadjoint_bounded_clinear) lemma id_cblinfun_adjoint[simp]: "id_cblinfun* = id_cblinfun" apply transfer using cadjoint_id by (metis eq_id_iff) lemma double_adj[simp]: "(A*)* = A" apply transfer using double_cadjoint by blast lemma adj_cblinfun_compose[simp]: fixes B::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and A::\'b \\<^sub>C\<^sub>L 'c::complex_inner\ shows "(A o\<^sub>C\<^sub>L B)* = (B*) o\<^sub>C\<^sub>L (A*)" proof transfer fix A :: \'b \ 'c\ and B :: \'a \ 'b\ assume \bounded_clinear A\ and \bounded_clinear B\ hence \bounded_clinear (A \ B)\ by (simp add: comp_bounded_clinear) have \\ (A \ B) u, v \ = \ u, (B\<^sup>\ \ A\<^sup>\) v \\ for u v by (metis (no_types, lifting) cadjoint_univ_prop \bounded_clinear A\ \bounded_clinear B\ cinner_commute' comp_def) thus \(A \ B)\<^sup>\ = B\<^sup>\ \ A\<^sup>\\ using \bounded_clinear (A \ B)\ by (metis cadjoint_eqI cinner_commute') qed lemma scaleC_adj[simp]: "(a *\<^sub>C A)* = (cnj a) *\<^sub>C (A*)" apply transfer by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear bounded_clinear_def complex_vector.linear_scale scaleC_cadjoint) lemma scaleR_adj[simp]: "(a *\<^sub>R A)* = a *\<^sub>R (A*)" by (simp add: scaleR_scaleC) lemma adj_plus: \(A + B)* = (A*) + (B*)\ proof transfer fix A B::\'b \ 'a\ assume a1: \bounded_clinear A\ and a2: \bounded_clinear B\ define F where \F = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ define G where \G = (\x. A x + B x)\ have \bounded_clinear G\ unfolding G_def by (simp add: a1 a2 bounded_clinear_add) moreover have \\F u, v\ = \u, G v\\ for u v unfolding F_def G_def using cadjoint_univ_prop a1 a2 cinner_add_left by (simp add: cadjoint_univ_prop cinner_add_left cinner_add_right) ultimately have \F = G\<^sup>\ \ using cadjoint_eqI by blast thus \(\x. A x + B x)\<^sup>\ = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ unfolding F_def G_def by auto qed lemma cinner_sup_norm_cblinfun: fixes A :: \'a::{complex_normed_vector,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ shows \norm A = (SUP (\,\). cmod (cinner \ (A *\<^sub>V \)) / (norm \ * norm \))\ apply transfer apply (rule cinner_sup_onorm) by (simp add: bounded_clinear.bounded_linear) lemma cinner_adj_left: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\G* *\<^sub>V x, y\ = \x, G *\<^sub>V y\\ apply transfer using cadjoint_univ_prop by blast lemma cinner_adj_right: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\x, G* *\<^sub>V y\ = \G *\<^sub>V x, y\\ apply transfer using cadjoint_univ_prop' by blast lemma adj_0[simp]: \0* = 0\ by (metis add_cancel_right_left adj_plus) lemma norm_adj[simp]: \norm (A*) = norm A\ for A :: \'b::chilbert_space \\<^sub>C\<^sub>L 'c::complex_inner\ proof (cases \(\x y :: 'b. x \ y) \ (\x y :: 'c. x \ y)\) case True then have c1: \class.not_singleton TYPE('b)\ apply intro_classes by simp from True have c2: \class.not_singleton TYPE('c)\ apply intro_classes by simp have normA: \norm A = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c1) have normAadj: \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A* *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c2) have \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ unfolding normAadj apply (subst cinner_adj_right) apply (subst cinner_commute) apply (subst complex_mod_cnj) by rule also have \\ = Sup ((\(\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \)) ` prod.swap ` UNIV)\ by auto also have \\ = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (subst image_image) by auto also have \\ = norm A\ unfolding normA by (simp add: mult.commute) finally show ?thesis by - next case False then consider (b) \\x::'b. x = 0\ | (c) \\x::'c. x = 0\ by auto then have \A = 0\ apply (cases; transfer) apply (metis (full_types) bounded_clinear_def complex_vector.linear_0) by auto then show \norm (A*) = norm A\ by simp qed lemma antilinear_adj[simp]: \antilinear adj\ apply (rule antilinearI) by (auto simp add: adj_plus) lemma bounded_antilinear_adj[bounded_antilinear, simp]: \bounded_antilinear adj\ by (auto intro!: antilinearI exI[of _ 1] simp: bounded_antilinear_def bounded_antilinear_axioms_def adj_plus) lemma adjoint_eqI: fixes G:: \'b::chilbert_space \\<^sub>C\<^sub>L 'a::chilbert_space\ and F:: \'a \\<^sub>C\<^sub>L 'b\ assumes \\x y. \(cblinfun_apply F) x, y\ = \x, (cblinfun_apply G) y\\ shows \F = G*\ using assms apply transfer using cadjoint_eqI by auto lemma cinner_real_hermiteanI: \ \Prop. II.2.12 in @{cite conway2013course}\ assumes \\\. cinner \ (A *\<^sub>V \) \ \\ shows \A = A*\ proof - { fix g h :: 'a { fix \ :: complex have \cinner h (A h) + cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) + (abs \)\<^sup>2 * cinner g (A g) = cinner (h + \ *\<^sub>C g) (A *\<^sub>V (h + \ *\<^sub>C g))\ (is \?sum4 = _\) apply (auto simp: cinner_add_right cinner_add_left cblinfun.add_right cblinfun.scaleC_right ring_class.ring_distribs) by (metis cnj_x_x mult.commute) also have \\ \ \\ using assms by auto finally have \?sum4 = cnj ?sum4\ using Reals_cnj_iff by fastforce then have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner (A h) g + cnj \ *\<^sub>C cinner (A g) h\ using Reals_cnj_iff abs_complex_real assms by force also have \\ = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by (simp add: cinner_adj_right) finally have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by - } from this[where \2=1] this[where \2=\] have 1: \cinner g (A h) + cinner h (A g) = cinner h (A* *\<^sub>V g) + cinner g (A* *\<^sub>V h)\ and i: \- \ * cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) - \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by auto from arg_cong2[OF 1 arg_cong[OF i, where f=\(*) (-\)\], where f=plus] have \cinner h (A g) = cinner h (A* *\<^sub>V g)\ by (auto simp: ring_class.ring_distribs) } then show "A = A*" by (simp add: adjoint_eqI cinner_adj_right) qed lemma norm_AAadj[simp]: \norm (A o\<^sub>C\<^sub>L A*) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof (cases \class.not_singleton TYPE('b)\) case True then have [simp]: \class.not_singleton TYPE('b)\ by - have 1: \(norm A)\<^sup>2 * \ \ norm (A o\<^sub>C\<^sub>L A*)\ if \\ < 1\ and \\ \ 0\ for \ proof - obtain \ where \: \norm ((A*) *\<^sub>V \) \ norm (A*) * sqrt \\ and [simp]: \norm \ = 1\ apply atomize_elim apply (rule cblinfun_norm_approx_witness_mult[internalize_sort' 'a]) using \\ < 1\ by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) have \complex_of_real ((norm A)\<^sup>2 * \) = (norm (A*) * sqrt \)\<^sup>2\ by (simp add: ordered_field_class.sign_simps(23) that(2)) also have \\ \ (norm ((A* *\<^sub>V \)))\<^sup>2\ apply (rule complex_of_real_mono) using \ apply (rule power_mono) using \\ \ 0\ by auto also have \\ \ cinner (A* *\<^sub>V \) (A* *\<^sub>V \)\ by (auto simp flip: power2_norm_eq_cinner) also have \\ = cinner \ (A *\<^sub>V A* *\<^sub>V \)\ by (simp add: cinner_adj_left) also have \\ = cinner \ ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ by auto also have \\ \ norm (A o\<^sub>C\<^sub>L A*)\ using \norm \ = 1\ by (smt (verit, best) Im_complex_of_real Re_complex_of_real \(A* *\<^sub>V \) \\<^sub>C (A* *\<^sub>V \) = \ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \)\ \\ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \) = \ \\<^sub>C ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ cdot_square_norm cinner_ge_zero cmod_Re complex_inner_class.Cauchy_Schwarz_ineq2 less_eq_complex_def mult_cancel_left1 mult_cancel_right1 norm_cblinfun) finally show ?thesis by (auto simp: less_eq_complex_def) qed then have 1: \(norm A)\<^sup>2 \ norm (A o\<^sub>C\<^sub>L A*)\ by (metis field_le_mult_one_interval less_eq_real_def ordered_field_class.sign_simps(5)) have 2: \norm (A o\<^sub>C\<^sub>L A*) \ (norm A)\<^sup>2\ proof (rule norm_cblinfun_bound) show \0 \ (norm A)\<^sup>2\ by simp fix \ have \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) = norm (A *\<^sub>V A* *\<^sub>V \)\ by auto also have \\ \ norm A * norm (A* *\<^sub>V \)\ by (simp add: norm_cblinfun) also have \\ \ norm A * norm (A*) * norm \\ by (metis mult.assoc norm_cblinfun norm_imp_pos_and_ge ordered_comm_semiring_class.comm_mult_left_mono) also have \\ = (norm A)\<^sup>2 * norm \\ by (simp add: power2_eq_square) finally show \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) \ (norm A)\<^sup>2 * norm \\ by - qed from 1 2 show ?thesis by simp next case False then have [simp]: \class.CARD_1 TYPE('b)\ by (rule not_singleton_vs_CARD_1) have \A = 0\ apply (rule cblinfun_to_CARD_1_0[internalize_sort' 'b]) by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) then show ?thesis by auto qed subsection \Unitaries / isometries\ definition isometry::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun\ definition unitary::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \unitary U \ (U* o\<^sub>C\<^sub>L U = id_cblinfun) \ (U o\<^sub>C\<^sub>L U* = id_cblinfun)\ lemma unitary_twosided_isometry: "unitary U \ isometry U \ isometry (U*)" unfolding unitary_def isometry_def by simp lemma isometryD[simp]: "isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding isometry_def by simp (* Not [simp] because isometryD[simp] + unitary_isometry[simp] already have the same effect *) lemma unitaryD1: "unitary U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding unitary_def by simp lemma unitaryD2[simp]: "unitary U \ U o\<^sub>C\<^sub>L U* = id_cblinfun" unfolding unitary_def by simp lemma unitary_isometry[simp]: "unitary U \ isometry U" unfolding unitary_def isometry_def by simp lemma unitary_adj[simp]: "unitary (U*) = unitary U" unfolding unitary_def by auto lemma isometry_cblinfun_compose[simp]: assumes "isometry A" and "isometry B" shows "isometry (A o\<^sub>C\<^sub>L B)" proof- have "B* o\<^sub>C\<^sub>L A* o\<^sub>C\<^sub>L (A o\<^sub>C\<^sub>L B) = id_cblinfun" if "A* o\<^sub>C\<^sub>L A = id_cblinfun" and "B* o\<^sub>C\<^sub>L B = id_cblinfun" using that by (smt (verit, del_insts) adjoint_eqI cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply) thus ?thesis using assms unfolding isometry_def by simp qed lemma unitary_cblinfun_compose[simp]: "unitary (A o\<^sub>C\<^sub>L B)" if "unitary A" and "unitary B" using that by (smt (z3) adj_cblinfun_compose cblinfun_compose_assoc cblinfun_compose_id_right double_adj isometryD isometry_cblinfun_compose unitary_def unitary_isometry) lemma unitary_surj: assumes "unitary U" shows "surj (cblinfun_apply U)" apply (rule surjI[where f=\cblinfun_apply (U*)\]) using assms unfolding unitary_def apply transfer using comp_eq_dest_lhs by force lemma unitary_id[simp]: "unitary id_cblinfun" by (simp add: unitary_def) lemma orthogonal_on_basis_is_isometry: assumes spanB: \ccspan B = \\ assumes orthoU: \\b c. b\B \ c\B \ cinner (U *\<^sub>V b) (U *\<^sub>V c) = cinner b c\ shows \isometry U\ proof - have [simp]: \b \ closure (cspan B)\ for b using spanB apply transfer by simp have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ and \\\B\ for \ \ by (simp add: cinner_adj_left orthoU that(1) that(2)) have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ for \ \ apply (rule bounded_clinear_eq_on[where t=\ and G=B]) using bounded_clinear_cinner_right *[OF that] by auto have \U* *\<^sub>V U *\<^sub>V \ = \\ if \\\B\ for \ apply (rule cinner_extensionality) apply (subst cinner_eq_flip) by (simp add: * that) then have \U* o\<^sub>C\<^sub>L U = id_cblinfun\ by (metis cblinfun_apply_cblinfun_compose cblinfun_eq_gen_eqI cblinfun_id_cblinfun_apply spanB) then show \isometry U\ using isometry_def by blast qed subsection \Images\ (* Closure is necessary. See email 47a3bb3d-3cc3-0934-36eb-3ef0f7b70a85@ut.ee *) lift_definition cblinfun_image :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a ccsubspace \ 'b ccsubspace\ (infixr "*\<^sub>S" 70) is "\A S. closure (A ` S)" using bounded_clinear_def closed_closure closed_csubspace.intro by (simp add: bounded_clinear_def complex_vector.linear_subspace_image closure_is_closed_csubspace) lemma cblinfun_image_mono: assumes a1: "S \ T" shows "A *\<^sub>S S \ A *\<^sub>S T" using a1 by (simp add: cblinfun_image.rep_eq closure_mono image_mono less_eq_ccsubspace.rep_eq) lemma cblinfun_image_0[simp]: shows "U *\<^sub>S 0 = 0" thm zero_ccsubspace_def apply transfer by (simp add: bounded_clinear_def complex_vector.linear_0) lemma cblinfun_image_bot[simp]: "U *\<^sub>S bot = bot" using cblinfun_image_0 by auto lemma cblinfun_image_sup[simp]: fixes A B :: \'a::chilbert_space ccsubspace\ and U :: "'a \\<^sub>C\<^sub>L'b::chilbert_space" shows \U *\<^sub>S (sup A B) = sup (U *\<^sub>S A) (U *\<^sub>S B)\ apply transfer using bounded_clinear.bounded_linear closure_image_closed_sum by blast lemma scaleC_cblinfun_image[simp]: fixes A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b :: chilbert_space\ and S :: \'a ccsubspace\ and \ :: complex shows \(\ *\<^sub>C A) *\<^sub>S S = \ *\<^sub>C (A *\<^sub>S S)\ proof- have \closure ( ( ((*\<^sub>C) \) \ (cblinfun_apply A) ) ` space_as_set S) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis closure_scaleC image_comp) hence \(closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis (mono_tags, lifting) comp_apply image_cong scaleC_cblinfun.rep_eq) hence \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by (metis space_as_set_inverse cblinfun_image.rep_eq scaleC_ccsubspace.rep_eq) have x1: "Abs_clinear_space (closure ((*\<^sub>V) (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure ((*\<^sub>V) A ` space_as_set S))" using \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by blast show ?thesis unfolding cblinfun_image_def using x1 by force qed lemma cblinfun_image_id[simp]: "id_cblinfun *\<^sub>S \ = \" apply transfer by (simp add: closed_csubspace.closed) lemma cblinfun_compose_image: \(A o\<^sub>C\<^sub>L B) *\<^sub>S S = A *\<^sub>S (B *\<^sub>S S)\ apply transfer unfolding image_comp[symmetric] apply (rule closure_bounded_linear_image_subset_eq[symmetric]) by (simp add: bounded_clinear.bounded_linear) lemmas cblinfun_assoc_left = cblinfun_compose_assoc[symmetric] cblinfun_compose_image[symmetric] add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space", symmetric] lemmas cblinfun_assoc_right = cblinfun_compose_assoc cblinfun_compose_image add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space"] lemma cblinfun_image_INF_leq[simp]: fixes U :: "'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::cbanach" and V :: "'a \ 'b ccsubspace" shows \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ apply transfer by (simp add: INT_greatest Inter_lower closure_mono image_mono) lemma isometry_cblinfun_image_inf_distrib': fixes U::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::cbanach\ and B C::"'a ccsubspace" shows "U *\<^sub>S (inf B C) \ inf (U *\<^sub>S B) (U *\<^sub>S C)" proof - define V where \V b = (if b then B else C)\ for b have \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ by auto then show ?thesis unfolding V_def by (metis (mono_tags, lifting) INF_UNIV_bool_expand) qed lemma cblinfun_image_eq: fixes S :: "'a::cbanach ccsubspace" and A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::cbanach" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and "ccspan G \ S" shows "A *\<^sub>S S = B *\<^sub>S S" proof (use assms in transfer) fix G :: "'a set" and A :: "'a \ 'b" and B :: "'a \ 'b" and S :: "'a set" assume a1: "bounded_clinear A" assume a2: "bounded_clinear B" assume a3: "\x. x \ G \ A x = B x" assume a4: "S \ closure (cspan G)" have "A ` closure S = B ` closure S" by (smt (verit, best) UnCI a1 a2 a3 a4 bounded_clinear_eq_on closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis Complex_Vector_Spaces0.bounded_clinear.bounded_linear a1 a2 closure_bounded_linear_image_subset_eq) qed lemma cblinfun_fixes_range: assumes "A o\<^sub>C\<^sub>L B = B" and "\ \ space_as_set (B *\<^sub>S top)" shows "A *\<^sub>V \ = \" proof- define rangeB rangeB' where "rangeB = space_as_set (B *\<^sub>S top)" and "rangeB' = range (cblinfun_apply B)" from assms have "\ \ closure rangeB'" by (simp add: cblinfun_image.rep_eq rangeB'_def top_ccsubspace.rep_eq) then obtain \i where \i_lim: "\i \ \" and \i_B: "\i i \ rangeB'" for i using closure_sequential by blast have A_invariant: "A *\<^sub>V \i i = \i i" for i proof- from \i_B obtain \ where \: "\i i = B *\<^sub>V \" using rangeB'_def by blast hence "A *\<^sub>V \i i = (A o\<^sub>C\<^sub>L B) *\<^sub>V \" by (simp add: cblinfun_compose.rep_eq) also have "\ = B *\<^sub>V \" by (simp add: assms) also have "\ = \i i" by (simp add: \) finally show ?thesis. qed from \i_lim have "(\i. A *\<^sub>V (\i i)) \ A *\<^sub>V \" by (rule isCont_tendsto_compose[rotated], simp) with A_invariant have "(\i. \i i) \ A *\<^sub>V \" by auto with \i_lim show "A *\<^sub>V \ = \" using LIMSEQ_unique by blast qed lemma zero_cblinfun_image[simp]: "0 *\<^sub>S S = (0::_ ccsubspace)" apply transfer by (simp add: complex_vector.subspace_0 image_constant[where x=0]) lemma cblinfun_image_INF_eq_general: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L'c::chilbert_space" and Uinv :: "'c \\<^sub>C\<^sub>L'b" assumes UinvUUinv: "Uinv o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L Uinv = Uinv" and UUinvU: "U o\<^sub>C\<^sub>L Uinv o\<^sub>C\<^sub>L U = U" \ \Meaning: \<^term>\Uinv\ is a Pseudoinverse of \<^term>\U\\ and V: "\i. V i \ Uinv *\<^sub>S top" shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof (rule antisym) show "U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S V i)" by (rule cblinfun_image_INF_leq) next define rangeU rangeUinv where "rangeU = U *\<^sub>S top" and "rangeUinv = Uinv *\<^sub>S top" define INFUV INFV where INFUV_def: "INFUV = (INF i. U *\<^sub>S V i)" and INFV_def: "INFV = (INF i. V i)" from assms have "V i \ rangeUinv" for i unfolding rangeUinv_def by simp moreover have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set rangeUinv" for \ using UinvUUinv cblinfun_fixes_range rangeUinv_def that by fastforce ultimately have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set (V i)" for \ i using less_eq_ccsubspace.rep_eq that by blast hence d1: "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>S (V i) = (V i)" for i proof transfer show "closure ((Uinv \ U) ` V i) = V i" if "pred_fun \ closed_csubspace V" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ (Uinv \ U) \ = \" for V :: "'a \ 'b set" and Uinv :: "'c \ 'b" and U :: "'b \ 'c" and i :: 'a using that proof auto show "x \ V i" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ closure (V i)" for x :: 'b using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure (V i)" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ V i" for x :: 'b using that using setdist_eq_0_sing_1 setdist_sing_in_set by blast qed qed have "U *\<^sub>S V i \ rangeU" for i by (simp add: cblinfun_image_mono rangeU_def) hence "INFUV \ rangeU" unfolding INFUV_def by (meson INF_lower UNIV_I order_trans) moreover have "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set rangeU" for \ using UUinvU cblinfun_fixes_range rangeU_def that by fastforce ultimately have x: "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set INFUV" for \ by (simp add: in_mono less_eq_ccsubspace.rep_eq that) have "closure ((U \ Uinv) ` INFUV) = INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ (U \ Uinv) \ = \" for INFUV :: "'c set" and U :: "'b \ 'c" and Uinv :: "'c \ 'b" using that proof auto show "x \ INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ closure INFUV" for x :: 'c using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ INFUV" for x :: 'c using that using setdist_eq_0_sing_1 setdist_sing_in_set by (simp add: closed_csubspace.closed) qed hence "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>S INFUV = INFUV" by (metis (mono_tags, opaque_lifting) x cblinfun_image.rep_eq cblinfun_image_id id_cblinfun_apply image_cong space_as_set_inject) hence "INFUV = U *\<^sub>S Uinv *\<^sub>S INFUV" by (simp add: cblinfun_compose_image) also have "\ \ U *\<^sub>S (INF i. Uinv *\<^sub>S U *\<^sub>S V i)" unfolding INFUV_def by (metis cblinfun_image_mono cblinfun_image_INF_leq) also have "\ = U *\<^sub>S INFV" using d1 by (metis (no_types, lifting) INFV_def cblinfun_assoc_left(2) image_cong) finally show "INFUV \ U *\<^sub>S INFV". qed lemma unitary_range[simp]: assumes "unitary U" shows "U *\<^sub>S top = top" using assms unfolding unitary_def apply transfer by (metis closure_UNIV comp_apply surj_def) lemma range_adjoint_isometry: assumes "isometry U" shows "U* *\<^sub>S top = top" proof- from assms have "top = U* *\<^sub>S U *\<^sub>S top" by (simp add: cblinfun_assoc_left(2)) also have "\ \ U* *\<^sub>S top" by (simp add: cblinfun_image_mono) finally show ?thesis using top.extremum_unique by blast qed lemma cblinfun_image_INF_eq[simp]: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L 'c::chilbert_space" assumes \isometry U\ shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof - from \isometry U\ have "U* o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L U* = U*" unfolding isometry_def by simp moreover from \isometry U\ have "U o\<^sub>C\<^sub>L U* o\<^sub>C\<^sub>L U = U" unfolding isometry_def by (simp add: cblinfun_compose_assoc) moreover have "V i \ U* *\<^sub>S top" for i by (simp add: range_adjoint_isometry assms) ultimately show ?thesis by (rule cblinfun_image_INF_eq_general) qed lemma isometry_cblinfun_image_inf_distrib[simp]: fixes U::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and X Y::"'a ccsubspace" assumes "isometry U" shows "U *\<^sub>S (inf X Y) = inf (U *\<^sub>S X) (U *\<^sub>S Y)" using cblinfun_image_INF_eq[where V="\b. if b then X else Y" and U=U] unfolding INF_UNIV_bool_expand using assms by auto lemma cblinfun_image_ccspan: shows "A *\<^sub>S ccspan G = ccspan ((*\<^sub>V) A ` G)" apply transfer by (simp add: bounded_clinear.bounded_linear bounded_clinear_def closure_bounded_linear_image_subset_eq complex_vector.linear_span_image) lemma cblinfun_apply_in_image[simp]: "A *\<^sub>V \ \ space_as_set (A *\<^sub>S \)" by (metis cblinfun_image.rep_eq closure_subset in_mono range_eqI top_ccsubspace.rep_eq) lemma cblinfun_plus_image_distr: \(A + B) *\<^sub>S S \ A *\<^sub>S S \ B *\<^sub>S S\ apply transfer by (smt (verit, ccfv_threshold) closed_closure closed_sum_def closure_minimal closure_subset image_subset_iff set_plus_intro subset_eq) lemma cblinfun_sum_image_distr: \(\i\I. A i) *\<^sub>S S \ (SUP i\I. A i *\<^sub>S S)\ proof (cases \finite I\) case True then show ?thesis proof induction case empty then show ?case by auto next case (insert x F) then show ?case apply auto by (smt (z3) cblinfun_plus_image_distr inf_sup_aci(6) le_iff_sup) qed next case False then show ?thesis by auto qed subsection \Sandwiches\ lift_definition sandwich :: \('a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner) \ (('a \\<^sub>C\<^sub>L 'a) \\<^sub>C\<^sub>L ('b \\<^sub>C\<^sub>L 'b))\ is \\(A::'a\\<^sub>C\<^sub>L'b) B. A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ proof fix A :: \'a \\<^sub>C\<^sub>L 'b\ and B B1 B2 :: \'a \\<^sub>C\<^sub>L 'a\ and c :: complex show \A o\<^sub>C\<^sub>L (B1 + B2) o\<^sub>C\<^sub>L A* = (A o\<^sub>C\<^sub>L B1 o\<^sub>C\<^sub>L A*) + (A o\<^sub>C\<^sub>L B2 o\<^sub>C\<^sub>L A*)\ by (simp add: cblinfun_compose_add_left cblinfun_compose_add_right) show \A o\<^sub>C\<^sub>L (c *\<^sub>C B) o\<^sub>C\<^sub>L A* = c *\<^sub>C (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*)\ by auto show \\K. \B. norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * K\ proof (rule exI[of _ \norm A * norm (A*)\], rule allI) fix B have \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by blast also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) finally show \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * (norm A * norm (A*))\ by (simp add: mult.assoc vector_space_over_itself.scale_left_commute) qed qed lemma sandwich_0[simp]: \sandwich 0 = 0\ by (simp add: cblinfun_eqI sandwich.rep_eq) lemma sandwich_apply: \sandwich A *\<^sub>V B = A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ apply (transfer fixing: A B) by auto lemma norm_sandwich: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'a::{chilbert_space} \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof - have main: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'c::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'd::{complex_inner}\ proof (rule norm_cblinfun_eqI) show \(norm A)\<^sup>2 \ norm (sandwich A *\<^sub>V id_cblinfun) / norm (id_cblinfun :: 'c \\<^sub>C\<^sub>L _)\ apply (auto simp: sandwich_apply) by - fix B have \norm (sandwich A *\<^sub>V B) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by (auto simp: sandwich_apply simp del: norm_adj) also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) also have \\ \ (norm A)\<^sup>2 * norm B\ by (simp add: power2_eq_square mult.assoc vector_space_over_itself.scale_left_commute) finally show \norm (sandwich A *\<^sub>V B) \ (norm A)\<^sup>2 * norm B\ by - show \0 \ (norm A)\<^sup>2\ by auto qed show ?thesis proof (cases \class.not_singleton TYPE('a)\) case True show ?thesis apply (rule main[internalize_sort' 'c2]) apply standard[1] using True by simp next case False have \A = 0\ apply (rule cblinfun_from_CARD_1_0[internalize_sort' 'a]) apply (rule not_singleton_vs_CARD_1) apply (rule False) by standard then show ?thesis by simp qed qed lemma sandwich_apply_adj: \sandwich A (B*) = (sandwich A B)*\ by (simp add: cblinfun_assoc_left(1) sandwich_apply) lemma sandwich_id[simp]: "sandwich id_cblinfun = id_cblinfun" apply (rule cblinfun_eqI) by (auto simp: sandwich_apply) subsection \Projectors\ lift_definition Proj :: "('a::chilbert_space) ccsubspace \ 'a \\<^sub>C\<^sub>L'a" is \projection\ by (rule projection_bounded_clinear) lemma Proj_range[simp]: "Proj S *\<^sub>S top = S" proof transfer fix S :: \'a set\ assume \closed_csubspace S\ then have "closure (range (projection S)) \ S" by (metis closed_csubspace.closed closed_csubspace.subspace closure_closed complex_vector.subspace_0 csubspace_is_convex dual_order.eq_iff insert_absorb insert_not_empty projection_image) moreover have "S \ closure (range (projection S))" using \closed_csubspace S\ by (metis closed_csubspace_def closure_subset csubspace_is_convex equals0D projection_image subset_iff) ultimately show \closure (range (projection S)) = S\ by auto qed lemma adj_Proj: \(Proj M)* = Proj M\ apply transfer by (simp add: projection_cadjoint) lemma Proj_idempotent[simp]: \Proj M o\<^sub>C\<^sub>L Proj M = Proj M\ proof - have u1: \(cblinfun_apply (Proj M)) = projection (space_as_set M)\ apply transfer by blast have \closed_csubspace (space_as_set M)\ using space_as_set by auto hence u2: \(projection (space_as_set M))\(projection (space_as_set M)) = (projection (space_as_set M))\ using projection_idem by fastforce have \(cblinfun_apply (Proj M)) \ (cblinfun_apply (Proj M)) = cblinfun_apply (Proj M)\ using u1 u2 by simp hence \cblinfun_apply ((Proj M) o\<^sub>C\<^sub>L (Proj M)) = cblinfun_apply (Proj M)\ by (simp add: cblinfun_compose.rep_eq) thus ?thesis using cblinfun_apply_inject by auto qed lift_definition is_Proj::\'a::chilbert_space \\<^sub>C\<^sub>L 'a \ bool\ is \\P. \M. closed_csubspace M \ is_projection_on P M\ . lemma Proj_on_own_range': fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \P o\<^sub>C\<^sub>L P = P\ and \P = P*\ shows \Proj (P *\<^sub>S top) = P\ proof- define M where "M = P *\<^sub>S top" have v3: "x \ (\x. x - P *\<^sub>V x) -` {0}" if "x \ range (cblinfun_apply P)" for x :: 'a proof- have v3_1: \cblinfun_apply P \ cblinfun_apply P = cblinfun_apply P\ by (metis \P o\<^sub>C\<^sub>L P = P\ cblinfun_compose.rep_eq) have \\t. P *\<^sub>V t = x\ using that by blast then obtain t where t_def: \P *\<^sub>V t = x\ by blast hence \x - P *\<^sub>V x = x - P *\<^sub>V (P *\<^sub>V t)\ by simp also have \\ = x - (P *\<^sub>V t)\ using v3_1 by (metis comp_apply) also have \\ = 0\ by (simp add: t_def) finally have \x - P *\<^sub>V x = 0\ by blast thus ?thesis by simp qed have v1: "range (cblinfun_apply P) \ (\x. x - cblinfun_apply P x) -` {0}" using v3 by blast have "x \ range (cblinfun_apply P)" if "x \ (\x. x - P *\<^sub>V x) -` {0}" for x :: 'a proof- have x1:\x - P *\<^sub>V x = 0\ using that by blast have \x = P *\<^sub>V x\ by (simp add: x1 eq_iff_diff_eq_0) thus ?thesis by blast qed hence v2: "(\x. x - cblinfun_apply P x) -` {0} \ range (cblinfun_apply P)" by blast have i1: \range (cblinfun_apply P) = (\ x. x - cblinfun_apply P x) -` {0}\ using v1 v2 by (simp add: v1 dual_order.antisym) have p1: \closed {(0::'a)}\ by simp have p2: \continuous (at x) (\ x. x - P *\<^sub>V x)\ for x proof- have \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by (simp add: id_cblinfun.rep_eq minus_cblinfun.rep_eq) hence \bounded_clinear (cblinfun_apply (id_cblinfun - P))\ using cblinfun_apply by blast hence \continuous (at x) (cblinfun_apply (id_cblinfun - P))\ by (simp add: clinear_continuous_at) thus ?thesis using \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by simp qed have i2: \closed ( (\ x. x - P *\<^sub>V x) -` {0} )\ using p1 p2 by (rule Abstract_Topology.continuous_closed_vimage) have \closed (range (cblinfun_apply P))\ using i1 i2 by simp have u2: \cblinfun_apply P x \ space_as_set M\ for x by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq top_ccsubspace.rep_eq) have xy: \\ x - P *\<^sub>V x, y \ = 0\ if y1: \y \ space_as_set M\ for x y proof- have \\t. y = P *\<^sub>V t\ using y1 by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq image_iff top_ccsubspace.rep_eq) then obtain t where t_def: \y = P *\<^sub>V t\ by blast have \\ x - P *\<^sub>V x, y \ = \ x - P *\<^sub>V x, P *\<^sub>V t \\ by (simp add: t_def) also have \\ = \ P *\<^sub>V (x - P *\<^sub>V x), t \\ by (metis \P = P*\ cinner_adj_left) also have \\ = \ P *\<^sub>V x - P *\<^sub>V (P *\<^sub>V x), t \\ by (simp add: cblinfun.diff_right) also have \\ = \ P *\<^sub>V x - P *\<^sub>V x, t \\ by (metis assms(1) comp_apply cblinfun_compose.rep_eq) also have \\ = \ 0, t \\ by simp also have \\ = 0\ by simp finally show ?thesis by blast qed hence u1: \x - P *\<^sub>V x \ orthogonal_complement (space_as_set M)\ for x by (simp add: orthogonal_complementI) have "closed_csubspace (space_as_set M)" using space_as_set by auto hence f1: "(Proj M) *\<^sub>V a = P *\<^sub>V a" for a by (simp add: Proj.rep_eq projection_eqI u1 u2) have "(+) ((P - Proj M) *\<^sub>V a) = id" for a using f1 by (auto intro!: ext simp add: minus_cblinfun.rep_eq) hence "b - b = cblinfun_apply (P - Proj M) a" for a b by (metis (no_types) add_diff_cancel_right' id_apply) hence "cblinfun_apply (id_cblinfun - (P - Proj M)) a = a" for a by (simp add: minus_cblinfun.rep_eq) thus ?thesis using u1 u2 cblinfun_apply_inject diff_diff_eq2 diff_eq_diff_eq eq_id_iff id_cblinfun.rep_eq by (metis (no_types, opaque_lifting) M_def) qed lemma Proj_range_closed: assumes "is_Proj P" shows "closed (range (cblinfun_apply P))" using assms apply transfer using closed_csubspace.closed is_projection_on_image by blast lemma Proj_is_Proj[simp]: fixes M::\'a::chilbert_space ccsubspace\ shows \is_Proj (Proj M)\ proof- have u1: "closed_csubspace (space_as_set M)" using space_as_set by blast have v1: "h - Proj M *\<^sub>V h \ orthogonal_complement (space_as_set M)" for h by (simp add: Proj.rep_eq orthogonal_complementI projection_orthogonal u1) have v2: "Proj M *\<^sub>V h \ space_as_set M" for h by (metis Proj.rep_eq mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) have u2: "is_projection_on ((*\<^sub>V) (Proj M)) (space_as_set M)" unfolding is_projection_on_def by (simp add: smallest_dist_is_ortho u1 v1 v2) show ?thesis using u1 u2 is_Proj.rep_eq by blast qed lemma is_Proj_algebraic: fixes P::\'a::chilbert_space \\<^sub>C\<^sub>L 'a\ shows \is_Proj P \ P o\<^sub>C\<^sub>L P = P \ P = P*\ proof have "P o\<^sub>C\<^sub>L P = P" if "is_Proj P" using that apply transfer using is_projection_on_idem by fastforce moreover have "P = P*" if "is_Proj P" using that apply transfer by (metis is_projection_on_cadjoint) ultimately show "P o\<^sub>C\<^sub>L P = P \ P = P*" if "is_Proj P" using that by blast show "is_Proj P" if "P o\<^sub>C\<^sub>L P = P \ P = P*" using that Proj_on_own_range' Proj_is_Proj by metis qed lemma Proj_on_own_range: fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \is_Proj P\ shows \Proj (P *\<^sub>S top) = P\ using Proj_on_own_range' assms is_Proj_algebraic by blast lemma Proj_image_leq: "(Proj S) *\<^sub>S A \ S" by (metis Proj_range inf_top_left le_inf_iff isometry_cblinfun_image_inf_distrib') lemma Proj_sandwich: fixes A::"'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space" assumes "isometry A" shows "sandwich A *\<^sub>V Proj S = Proj (A *\<^sub>S S)" proof- define P where \P = A o\<^sub>C\<^sub>L Proj S o\<^sub>C\<^sub>L (A*)\ have \P o\<^sub>C\<^sub>L P = P\ using assms unfolding P_def isometry_def by (metis (no_types, lifting) Proj_idempotent cblinfun_assoc_left(1) cblinfun_compose_id_left) moreover have \P = P*\ unfolding P_def by (metis adj_Proj adj_cblinfun_compose cblinfun_assoc_left(1) double_adj) ultimately have \\M. P = Proj M \ space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ using P_def Proj_on_own_range' by (metis Proj_is_Proj Proj_range_closed cblinfun_image.rep_eq closure_closed top_ccsubspace.rep_eq) then obtain M where \P = Proj M\ and \space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ by blast have f1: "A o\<^sub>C\<^sub>L Proj S = P o\<^sub>C\<^sub>L A" by (simp add: P_def assms cblinfun_compose_assoc) hence "P o\<^sub>C\<^sub>L A o\<^sub>C\<^sub>L A* = P" using P_def by presburger hence "(P o\<^sub>C\<^sub>L A) *\<^sub>S (c \ A* *\<^sub>S d) = P *\<^sub>S (A *\<^sub>S c \ d)" for c d by (simp add: cblinfun_assoc_left(2)) hence "P *\<^sub>S (A *\<^sub>S \ \ c) = (P o\<^sub>C\<^sub>L A) *\<^sub>S \" for c by (metis sup_top_left) hence \M = A *\<^sub>S S\ using f1 by (metis \P = Proj M\ cblinfun_assoc_left(2) Proj_range sup_top_right) thus ?thesis using \P = Proj M\ unfolding P_def sandwich_apply by blast qed lemma Proj_orthog_ccspan_union: assumes "\x y. x \ X \ y \ Y \ is_orthogonal x y" shows \Proj (ccspan (X \ Y)) = Proj (ccspan X) + Proj (ccspan Y)\ proof - have \x \ cspan X \ y \ cspan Y \ is_orthogonal x y\ for x y apply (rule is_orthogonal_closure_cspan[where X=X and Y=Y]) using closure_subset assms by auto then have \x \ closure (cspan X) \ y \ closure (cspan Y) \ is_orthogonal x y\ for x y by (metis orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') then show ?thesis apply (transfer fixing: X Y) apply (subst projection_plus[symmetric]) by auto qed abbreviation proj :: "'a::chilbert_space \ 'a \\<^sub>C\<^sub>L 'a" where "proj \ \ Proj (ccspan {\})" lemma proj_0[simp]: \proj 0 = 0\ apply transfer by auto lemma surj_isometry_is_unitary: fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes \isometry U\ assumes \U *\<^sub>S \ = \\ shows \unitary U\ by (metis Proj_sandwich sandwich_apply Proj_on_own_range' assms(1) assms(2) cblinfun_compose_id_right isometry_def unitary_def unitary_id unitary_range) lemma ccsubspace_supI_via_Proj: fixes A B C::"'a::chilbert_space ccsubspace" assumes a1: \Proj (- C) *\<^sub>S A \ B\ shows "A \ sup B C" proof- have x2: \x \ space_as_set B\ if "x \ closure ( (projection (orthogonal_complement (space_as_set C))) ` space_as_set A)" for x using that by (metis Proj.rep_eq cblinfun_image.rep_eq assms less_eq_ccsubspace.rep_eq subsetD uminus_ccsubspace.rep_eq) have q1: \x \ closure {\ + \ |\ \. \ \ space_as_set B \ \ \ space_as_set C}\ if \x \ space_as_set A\ for x proof- have p1: \closed_csubspace (space_as_set C)\ using space_as_set by auto hence \x = (projection (space_as_set C)) x + (projection (orthogonal_complement (space_as_set C))) x\ by simp hence \x = (projection (orthogonal_complement (space_as_set C))) x + (projection (space_as_set C)) x\ by (metis ordered_field_class.sign_simps(2)) moreover have \(projection (orthogonal_complement (space_as_set C))) x \ space_as_set B\ using x2 by (meson closure_subset image_subset_iff that) moreover have \(projection (space_as_set C)) x \ space_as_set C\ by (metis mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) ultimately show ?thesis using closure_subset by fastforce qed have x1: \x \ (space_as_set B +\<^sub>M space_as_set C)\ if "x \ space_as_set A" for x proof - have f1: "x \ closure {a + b |a b. a \ space_as_set B \ b \ space_as_set C}" by (simp add: q1 that) have "{a + b |a b. a \ space_as_set B \ b \ space_as_set C} = {a. \p. p \ space_as_set B \ (\q. q \ space_as_set C \ a = p + q)}" by blast hence "x \ closure {a. \b\space_as_set B. \c\space_as_set C. a = b + c}" using f1 by (simp add: Bex_def_raw) thus ?thesis using that unfolding closed_sum_def set_plus_def by blast qed hence \x \ space_as_set (Abs_clinear_space (space_as_set B +\<^sub>M space_as_set C))\ if "x \ space_as_set A" for x using that by (metis space_as_set_inverse sup_ccsubspace.rep_eq) thus ?thesis by (simp add: x1 less_eq_ccsubspace.rep_eq subset_eq sup_ccsubspace.rep_eq) qed lemma is_Proj_idempotent: assumes "is_Proj P" shows "P o\<^sub>C\<^sub>L P = P" using assms unfolding is_Proj_def using assms is_Proj_algebraic by auto lemma is_proj_selfadj: assumes "is_Proj P" shows "P* = P" using assms unfolding is_Proj_def by (metis is_Proj_algebraic is_Proj_def) lemma is_Proj_I: assumes "P o\<^sub>C\<^sub>L P = P" and "P* = P" shows "is_Proj P" using assms is_Proj_algebraic by metis lemma is_Proj_0[simp]: "is_Proj 0" by (metis add_left_cancel adj_plus bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose group_cancel.rule0 is_Proj_I) lemma is_Proj_complement[simp]: assumes a1: "is_Proj P" shows "is_Proj (id_cblinfun-P)" by (smt (z3) add_diff_cancel_left add_diff_cancel_left' adj_cblinfun_compose adj_plus assms bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel id_cblinfun_adjoint is_Proj_algebraic cblinfun_compose_id_left) lemma Proj_bot[simp]: "Proj bot = 0" by (metis zero_cblinfun_image Proj_on_own_range' is_Proj_0 is_Proj_algebraic zero_ccsubspace_def) lemma Proj_ortho_compl: "Proj (- X) = id_cblinfun - Proj X" by (transfer , auto) lemma Proj_inj: assumes "Proj X = Proj Y" shows "X = Y" by (metis assms Proj_range) subsection \Kernel\ lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector \ 'a ccsubspace" is "\ f. f -` {0}" by (metis kernel_is_closed_csubspace) definition eigenspace :: "complex \ 'a::complex_normed_vector \\<^sub>C\<^sub>L'a \ 'a ccsubspace" where "eigenspace a A = kernel (A - a *\<^sub>C id_cblinfun)" lemma kernel_scaleC[simp]: "a\0 \ kernel (a *\<^sub>C A) = kernel A" for a :: complex and A :: "(_,_) cblinfun" apply transfer using complex_vector.scale_eq_0_iff by blast lemma kernel_0[simp]: "kernel 0 = top" apply transfer by auto lemma kernel_id[simp]: "kernel id_cblinfun = 0" apply transfer by simp lemma eigenspace_scaleC[simp]: assumes a1: "a \ 0" shows "eigenspace b (a *\<^sub>C A) = eigenspace (b/a) A" proof - have "b *\<^sub>C (id_cblinfun::('a, _) cblinfun) = a *\<^sub>C (b / a) *\<^sub>C id_cblinfun" using a1 by (metis ceq_vector_fraction_iff) hence "kernel (a *\<^sub>C A - b *\<^sub>C id_cblinfun) = kernel (A - (b / a) *\<^sub>C id_cblinfun)" using a1 by (metis (no_types) complex_vector.scale_right_diff_distrib kernel_scaleC) thus ?thesis unfolding eigenspace_def by blast qed lemma eigenspace_memberD: assumes "x \ space_as_set (eigenspace e A)" shows "A *\<^sub>V x = e *\<^sub>C x" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms apply transfer by auto lemma eigenspace_memberI: assumes "A *\<^sub>V x = e *\<^sub>C x" shows "x \ space_as_set (eigenspace e A)" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms apply transfer by auto subsection \Isomorphisms and inverses\ definition iso_cblinfun :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ bool\ where \iso_cblinfun A = (\ B. A o\<^sub>C\<^sub>L B = id_cblinfun \ B o\<^sub>C\<^sub>L A = id_cblinfun)\ definition cblinfun_inv :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ ('b,'a) cblinfun\ where \cblinfun_inv A = (SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun)\ lemma assumes \iso_cblinfun A\ shows cblinfun_inv_left: \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and cblinfun_inv_right: \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ proof - from assms obtain B where AB: \A o\<^sub>C\<^sub>L B = id_cblinfun\ and BA: \B o\<^sub>C\<^sub>L A = id_cblinfun\ using iso_cblinfun_def by blast from BA have \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ by (metis (mono_tags, lifting) cblinfun_inv_def someI_ex) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB BA show \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ by auto qed lemma cblinfun_inv_uniq: assumes "A o\<^sub>C\<^sub>L B = id_cblinfun" and "B o\<^sub>C\<^sub>L A = id_cblinfun" shows "cblinfun_inv A = B" using assms by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) subsection \One-dimensional spaces\ instantiation cblinfun :: (one_dim, one_dim) complex_inner begin text \Once we have a theory for the trace, we could instead define the Hilbert-Schmidt inner product and relax the \<^class>\one_dim\-sort constraint to (\<^class>\cfinite_dim\,\<^class>\complex_normed_vector\) or similar\ definition "cinner_cblinfun (A::'a \\<^sub>C\<^sub>L 'b) (B::'a \\<^sub>C\<^sub>L 'b) = cnj (one_dim_iso (A *\<^sub>V 1)) * one_dim_iso (B *\<^sub>V 1)" instance proof intro_classes fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "\A, B\ = cnj \B, A\" unfolding cinner_cblinfun_def by auto show "\A + B, C\ = \A, C\ + \B, C\" by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) show "\c *\<^sub>C A, B\ = cnj c * \A, B\" by (simp add: cblinfun.scaleC_left cinner_cblinfun_def) show "0 \ \A, A\" unfolding cinner_cblinfun_def by auto have "bounded_clinear A \ A 1 = 0 \ A = (\_. 0)" for A::"'a \ 'b" proof (rule one_dim_clinear_eqI [where x = 1] , auto) show "clinear A" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: bounded_clinear.clinear) show "clinear ((\_. 0)::'a \ 'b)" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: complex_vector.module_hom_zero) qed hence "A *\<^sub>V 1 = 0 \ A = 0" by transfer hence "one_dim_iso (A *\<^sub>V 1) = 0 \ A = 0" by (metis one_dim_iso_of_zero one_dim_iso_inj) thus "(\A, A\ = 0) = (A = 0)" by (auto simp: cinner_cblinfun_def) show "norm A = sqrt (cmod \A, A\)" unfolding cinner_cblinfun_def apply transfer by (simp add: norm_mult abs_complex_def one_dim_onorm' cnj_x_x power2_eq_square bounded_clinear.clinear) qed end instantiation cblinfun :: (one_dim, one_dim) one_dim begin lift_definition one_cblinfun :: "'a \\<^sub>C\<^sub>L 'b" is "one_dim_iso" by (rule bounded_clinear_one_dim_iso) lift_definition times_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f g. f o one_dim_iso o g" by (simp add: comp_bounded_clinear) lift_definition inverse_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f. ((*) (one_dim_iso (inverse (f 1)))) o one_dim_iso" by (auto intro!: comp_bounded_clinear bounded_clinear_mult_right) definition divide_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" where "divide_cblinfun A B = A * inverse B" definition "canonical_basis_cblinfun = [1 :: 'a \\<^sub>C\<^sub>L 'b]" instance proof intro_classes let ?basis = "canonical_basis :: ('a \\<^sub>C\<^sub>L 'b) list" fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "distinct ?basis" unfolding canonical_basis_cblinfun_def by simp have "(1::'a \\<^sub>C\<^sub>L 'b) \ (0::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_one zero_neq_one) thus "cindependent (set ?basis)" unfolding canonical_basis_cblinfun_def by simp have "A \ cspan (set ?basis)" for A proof - define c :: complex where "c = one_dim_iso (A *\<^sub>V 1)" have "A x = one_dim_iso (A 1) *\<^sub>C one_dim_iso x" for x by (smt (z3) cblinfun.scaleC_right complex_vector.scale_left_commute one_dim_iso_idem one_dim_scaleC_1) hence "A = one_dim_iso (A *\<^sub>V 1) *\<^sub>C 1" apply transfer by metis thus "A \ cspan (set ?basis)" unfolding canonical_basis_cblinfun_def by (smt complex_vector.span_base complex_vector.span_scale list.set_intros(1)) qed thus "cspan (set ?basis) = UNIV" by auto have "A = (1::'a \\<^sub>C\<^sub>L 'b) \ norm (1::'a \\<^sub>C\<^sub>L 'b) = (1::real)" apply transfer by simp thus "A \ set ?basis \ norm A = 1" unfolding canonical_basis_cblinfun_def by simp show "?basis = [1]" unfolding canonical_basis_cblinfun_def by simp show "c *\<^sub>C 1 * c' *\<^sub>C 1 = (c * c') *\<^sub>C (1::'a\\<^sub>C\<^sub>L'b)" apply transfer by auto have "(1::'a \\<^sub>C\<^sub>L 'b) = (0::'a \\<^sub>C\<^sub>L 'b) \ False" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_zero' zero_neq_neg_one) thus "is_ortho_set (set ?basis)" unfolding is_ortho_set_def canonical_basis_cblinfun_def by auto show "A div B = A * inverse B" by (simp add: divide_cblinfun_def) show "inverse (c *\<^sub>C 1) = (1::'a\\<^sub>C\<^sub>L'b) /\<^sub>C c" apply transfer by (simp add: o_def one_dim_inverse) qed end lemma id_cblinfun_eq_1[simp]: \id_cblinfun = 1\ apply transfer by auto lemma one_dim_apply_is_times[simp]: fixes A :: "'a::one_dim \\<^sub>C\<^sub>L 'a" and B :: "'a \\<^sub>C\<^sub>L 'a" shows "A o\<^sub>C\<^sub>L B = A * B" apply transfer by simp lemma one_comp_one_cblinfun[simp]: "1 o\<^sub>C\<^sub>L 1 = 1" apply transfer unfolding o_def by simp lemma one_cblinfun_adj[simp]: "1* = 1" apply transfer by simp lemma scaleC_1_right[simp]: \scaleC x (1::'a::one_dim) = of_complex x\ unfolding of_complex_def by simp lemma scaleC_of_complex[simp]: \scaleC x (of_complex y) = of_complex (x * y)\ unfolding of_complex_def using scaleC_scaleC by blast lemma scaleC_1_apply[simp]: \(x *\<^sub>C 1) *\<^sub>V y = x *\<^sub>C y\ by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_apply_1_left[simp]: \1 *\<^sub>V y = y\ by (metis cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma of_complex_cblinfun_apply[simp]: \of_complex x *\<^sub>V y = x *\<^sub>C y\ unfolding of_complex_def by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_compose_1_left[simp]: \1 o\<^sub>C\<^sub>L x = x\ apply transfer by auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ apply transfer by auto lemma one_dim_iso_id_cblinfun: \one_dim_iso id_cblinfun = id_cblinfun\ by simp lemma one_dim_iso_id_cblinfun_eq_1: \one_dim_iso id_cblinfun = 1\ by simp lemma one_dim_iso_comp_distr[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a o\<^sub>C\<^sub>L one_dim_iso b\ by (smt (z3) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC) lemma one_dim_iso_comp_distr_times[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a * one_dim_iso b\ by (smt (verit, del_insts) mult.left_neutral mult_scaleC_left one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left) lemma one_dim_iso_adjoint[simp]: \one_dim_iso (A*) = (one_dim_iso A)*\ by (smt (z3) one_cblinfun_adj one_cinner_a_scaleC_one one_dim_iso_of_one one_dim_iso_scaleC scaleC_adj) lemma one_dim_iso_adjoint_complex[simp]: \one_dim_iso (A*) = cnj (one_dim_iso A)\ by (metis (mono_tags, lifting) one_cblinfun_adj one_dim_iso_idem one_dim_scaleC_1 scaleC_adj) lemma one_dim_cblinfun_compose_commute: \a o\<^sub>C\<^sub>L b = b o\<^sub>C\<^sub>L a\ for a b :: \('a::one_dim,'a) cblinfun\ by (simp add: one_dim_iso_inj) lemma one_cblinfun_apply_one[simp]: \1 *\<^sub>V 1 = 1\ by (simp add: one_cblinfun.rep_eq) subsection \Loewner order\ lift_definition heterogenous_cblinfun_id :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ is \if bounded_clinear (heterogenous_identity :: 'a::complex_normed_vector \ 'b::complex_normed_vector) then heterogenous_identity else (\_. 0)\ by auto lemma heterogenous_cblinfun_id_def'[simp]: "heterogenous_cblinfun_id = id_cblinfun" apply transfer by auto definition "heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'b::chilbert_space itself) \ unitary (heterogenous_cblinfun_id :: 'a \\<^sub>C\<^sub>L 'b) \ unitary (heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a)" lemma heterogenous_same_type_cblinfun[simp]: \heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'a::chilbert_space itself)\ unfolding heterogenous_same_type_cblinfun_def by auto instantiation cblinfun :: (chilbert_space, chilbert_space) ord begin definition less_eq_cblinfun :: \('a \\<^sub>C\<^sub>L 'b) \ ('a \\<^sub>C\<^sub>L 'b) \ bool\ where less_eq_cblinfun_def_heterogenous: \less_eq_cblinfun A B = (if heterogenous_same_type_cblinfun TYPE('a) TYPE('b) then \\::'b. cinner \ ((B-A) *\<^sub>V heterogenous_cblinfun_id *\<^sub>V \) \ 0 else (A=B))\ definition \less_cblinfun (A :: 'a \\<^sub>C\<^sub>L 'b) B \ A \ B \ \ B \ A\ instance.. end lemma less_eq_cblinfun_def: \A \ B \ (\\. cinner \ (A *\<^sub>V \) \ cinner \ (B *\<^sub>V \))\ unfolding less_eq_cblinfun_def_heterogenous by (auto simp del: less_eq_complex_def simp: cblinfun.diff_left cinner_diff_right) instantiation cblinfun :: (chilbert_space, chilbert_space) ordered_complex_vector begin instance proof intro_classes note less_eq_complex_def[simp del] fix x y z :: \'a \\<^sub>C\<^sub>L 'b\ fix a b :: complex define pos where \pos X \ (\\. cinner \ (X *\<^sub>V \) \ 0)\ for X :: \'b \\<^sub>C\<^sub>L 'b\ consider (unitary) \heterogenous_same_type_cblinfun TYPE('a) TYPE('b)\ \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B = pos ((B-A) o\<^sub>C\<^sub>L (heterogenous_cblinfun_id :: 'b\\<^sub>C\<^sub>L'a))\ | (trivial) \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B \ A = B\ apply atomize_elim by (auto simp: pos_def less_eq_cblinfun_def_heterogenous) note cases = this have [simp]: \pos 0\ unfolding pos_def by auto have pos_nondeg: \X = 0\ if \pos X\ and \pos (-X)\ for X apply (rule cblinfun_cinner_eqI, simp) using that by (metis (no_types, lifting) cblinfun.minus_left cinner_minus_right dual_order.antisym equation_minus_iff neg_le_0_iff_le pos_def) have pos_add: \pos (X+Y)\ if \pos X\ and \pos Y\ for X Y by (smt (z3) pos_def cblinfun.diff_left cinner_minus_right cinner_simps(3) diff_ge_0_iff_ge diff_minus_eq_add neg_le_0_iff_le order_trans that(1) that(2) uminus_cblinfun.rep_eq) have pos_scaleC: \pos (a *\<^sub>C X)\ if \a\0\ and \pos X\ for X a using that unfolding pos_def by (auto simp: cblinfun.scaleC_left) let ?id = \heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a\ show \x \ x\ apply (cases rule:cases) by auto show \(x < y) \ (x \ y \ \ y \ x)\ unfolding less_cblinfun_def by simp show \x \ z\ if \x \ y\ and \y \ z\ proof (cases rule:cases) case unitary define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ and \b = (z-y) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ with unitary that have \pos a\ and \pos b\ by auto then have \pos (a + b)\ by (rule pos_add) moreover have \a + b = (z - x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ unfolding a_def b_def by (metis (no_types, lifting) bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel ordered_field_class.sign_simps(2) ordered_field_class.sign_simps(8)) ultimately show ?thesis using unitary by auto next case trivial with that show ?thesis by auto qed show \x = y\ if \x \ y\ and \y \ x\ proof (cases rule:cases) case unitary then have \unitary ?id\ by (auto simp: heterogenous_same_type_cblinfun_def) define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L ?id\ and \b = (x-y) o\<^sub>C\<^sub>L ?id\ with unitary that have \pos a\ and \pos b\ by auto then have \a = 0\ apply (rule_tac pos_nondeg) apply (auto simp: a_def b_def) by (smt (verit, best) add.commute bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose cblinfun_compose_zero_left diff_0 diff_add_cancel group_cancel.rule0 group_cancel.sub1) then show ?thesis unfolding a_def using \unitary ?id\ by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_compose_zero_left eq_iff_diff_eq_0 unitaryD2) next case trivial with that show ?thesis by simp qed show \x + y \ x + z\ if \y \ z\ proof (cases rule:cases) case unitary with that show ?thesis by auto next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ a *\<^sub>C y\ if \x \ y\ and \0 \ a\ proof (cases rule:cases) case unitary with that pos_scaleC show ?thesis by (metis cblinfun_compose_scaleC_left complex_vector.scale_right_diff_distrib) next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ b *\<^sub>C x\ if \a \ b\ and \0 \ x\ proof (cases rule:cases) case unitary with that show ?thesis by (auto intro!: pos_scaleC simp flip: scaleC_diff_left) next case trivial with that show ?thesis by auto qed qed end lemma positive_id_cblinfun[simp]: "id_cblinfun \ 0" unfolding less_eq_cblinfun_def using cinner_ge_zero by auto lemma positive_hermitianI: \A = A*\ if \A \ 0\ apply (rule cinner_real_hermiteanI) using that by (auto simp: complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma positive_cblinfunI: \A \ 0\ if \\x. cinner x (A *\<^sub>V x) \ 0\ unfolding less_eq_cblinfun_def using that by auto (* Note: this does not require B to be a square operator *) lemma positive_cblinfun_squareI: \A = B* o\<^sub>C\<^sub>L B \ A \ 0\ apply (rule positive_cblinfunI) by (metis cblinfun_apply_cblinfun_compose cinner_adj_right cinner_ge_zero) lemma one_dim_loewner_order: \A \ B \ one_dim_iso A \ (one_dim_iso B :: complex)\ for A B :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ proof - note less_eq_complex_def[simp del] have A: \A = one_dim_iso A *\<^sub>C id_cblinfun\ by simp have B: \B = one_dim_iso B *\<^sub>C id_cblinfun\ by simp have \A \ B \ (\\. cinner \ (A \) \ cinner \ (B \))\ by (simp add: less_eq_cblinfun_def) also have \\ \ (\\::'a. one_dim_iso B * (\ \\<^sub>C \) \ one_dim_iso A * (\ \\<^sub>C \))\ apply (subst A, subst B) by (metis (no_types, opaque_lifting) cinner_scaleC_right id_cblinfun_apply scaleC_cblinfun.rep_eq) also have \\ \ one_dim_iso A \ (one_dim_iso B :: complex)\ by (auto intro!: mult_right_mono elim!: allE[where x=1]) finally show ?thesis by - qed lemma one_dim_positive: \A \ 0 \ one_dim_iso A \ (0::complex)\ for A :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ using one_dim_loewner_order[where B=0] by auto subsection \Embedding vectors to operators\ lift_definition vector_to_cblinfun :: \'a::complex_normed_vector \ 'b::one_dim \\<^sub>C\<^sub>L 'a\ is \\\ \. one_dim_iso \ *\<^sub>C \\ by (simp add: bounded_clinear_scaleC_const) lemma vector_to_cblinfun_cblinfun_apply: "vector_to_cblinfun (A *\<^sub>V \) = A o\<^sub>C\<^sub>L (vector_to_cblinfun \)" apply transfer unfolding comp_def bounded_clinear_def clinear_def Vector_Spaces.linear_def module_hom_def module_hom_axioms_def by simp lemma vector_to_cblinfun_add: \vector_to_cblinfun (x + y) = vector_to_cblinfun x + vector_to_cblinfun y\ apply transfer by (simp add: scaleC_add_right) lemma norm_vector_to_cblinfun[simp]: "norm (vector_to_cblinfun x) = norm x" proof transfer have "bounded_clinear (one_dim_iso::'a \ complex)" by simp moreover have "onorm (one_dim_iso::'a \ complex) * norm x = norm x" for x :: 'b by simp ultimately show "onorm (\\. one_dim_iso (\::'a) *\<^sub>C x) = norm x" for x :: 'b by (subst onorm_scaleC_left) qed lemma bounded_clinear_vector_to_cblinfun[bounded_clinear]: "bounded_clinear vector_to_cblinfun" apply (rule bounded_clinearI[where K=1]) apply (transfer, simp add: scaleC_add_right) apply (transfer, simp add: mult.commute) by simp lemma vector_to_cblinfun_scaleC[simp]: "vector_to_cblinfun (a *\<^sub>C \) = a *\<^sub>C vector_to_cblinfun \" for a::complex proof (subst asm_rl [of "a *\<^sub>C \ = (a *\<^sub>C id_cblinfun) *\<^sub>V \"]) show "a *\<^sub>C \ = a *\<^sub>C id_cblinfun *\<^sub>V \" by (simp add: scaleC_cblinfun.rep_eq) show "vector_to_cblinfun (a *\<^sub>C id_cblinfun *\<^sub>V \) = a *\<^sub>C (vector_to_cblinfun \::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun_id_cblinfun_apply cblinfun_compose_scaleC_left vector_to_cblinfun_cblinfun_apply) qed lemma vector_to_cblinfun_apply_one_dim[simp]: shows "vector_to_cblinfun \ *\<^sub>V \ = one_dim_iso \ *\<^sub>C \" apply transfer by (rule refl) lemma vector_to_cblinfun_adj_apply[simp]: shows "vector_to_cblinfun \* *\<^sub>V \ = of_complex (cinner \ \)" by (simp add: cinner_adj_right one_dim_iso_def one_dim_iso_inj) lemma vector_to_cblinfun_comp_one[simp]: "(vector_to_cblinfun s :: 'a::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L 1 = (vector_to_cblinfun s :: 'b::one_dim \\<^sub>C\<^sub>L _)" apply (transfer fixing: s) by fastforce lemma vector_to_cblinfun_0[simp]: "vector_to_cblinfun 0 = 0" by (metis cblinfun.zero_left cblinfun_compose_zero_left vector_to_cblinfun_cblinfun_apply) lemma image_vector_to_cblinfun[simp]: "vector_to_cblinfun x *\<^sub>S top = ccspan {x}" proof transfer show "closure (range (\\::'b. one_dim_iso \ *\<^sub>C x)) = closure (cspan {x})" for x :: 'a proof (rule arg_cong [where f = closure]) have "k *\<^sub>C x \ range (\\. one_dim_iso \ *\<^sub>C x)" for k by (smt (z3) id_apply one_dim_iso_id one_dim_iso_idem range_eqI) thus "range (\\. one_dim_iso (\::'b) *\<^sub>C x) = cspan {x}" unfolding complex_vector.span_singleton by auto qed qed lemma vector_to_cblinfun_adj_comp_vector_to_cblinfun[simp]: shows "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \ = cinner \ \ *\<^sub>C id_cblinfun" proof - have "one_dim_iso \ *\<^sub>C one_dim_iso (of_complex \\, \\) = \\, \\ *\<^sub>C one_dim_iso \" for \ :: "'c::one_dim" by (metis complex_vector.scale_left_commute of_complex_def one_dim_iso_of_one one_dim_iso_scaleC one_dim_scaleC_1) hence "one_dim_iso ((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = one_dim_iso ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by simp hence "((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by (rule one_dim_iso_inj) thus ?thesis using cblinfun_eqI[where x = "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \" and y = "\\, \\ *\<^sub>C id_cblinfun"] by auto qed lemma isometry_vector_to_cblinfun[simp]: assumes "norm x = 1" shows "isometry (vector_to_cblinfun x)" using assms cnorm_eq_1 isometry_def by force subsection \Butterflies (rank-1 projectors)\ definition butterfly_def: "butterfly (s::'a::complex_normed_vector) (t::'b::chilbert_space) = vector_to_cblinfun s o\<^sub>C\<^sub>L (vector_to_cblinfun t :: complex \\<^sub>C\<^sub>L _)*" abbreviation "selfbutter s \ butterfly s s" lemma butterfly_add_left: \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (simp add: butterfly_def vector_to_cblinfun_add cbilinear_add_left bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma butterfly_add_right: \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (simp add: butterfly_def adj_plus vector_to_cblinfun_add cblinfun_compose_add_right) lemma butterfly_def_one_dim: "butterfly s t = (vector_to_cblinfun s :: 'c::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L (vector_to_cblinfun t :: 'c \\<^sub>C\<^sub>L _)*" (is "_ = ?rhs") for s :: "'a::complex_normed_vector" and t :: "'b::chilbert_space" proof - let ?isoAC = "1 :: 'c \\<^sub>C\<^sub>L complex" let ?isoCA = "1 :: complex \\<^sub>C\<^sub>L 'c" let ?vector = "vector_to_cblinfun :: _ \ ('c \\<^sub>C\<^sub>L _)" have "butterfly s t = (?vector s o\<^sub>C\<^sub>L ?isoCA) o\<^sub>C\<^sub>L (?vector t o\<^sub>C\<^sub>L ?isoCA)*" unfolding butterfly_def vector_to_cblinfun_comp_one by simp also have "\ = ?vector s o\<^sub>C\<^sub>L (?isoCA o\<^sub>C\<^sub>L ?isoCA*) o\<^sub>C\<^sub>L (?vector t)*" by (metis (no_types, lifting) cblinfun_compose_assoc adj_cblinfun_compose) also have "\ = ?rhs" by simp finally show ?thesis by simp qed lemma butterfly_comp_cblinfun: "butterfly \ \ o\<^sub>C\<^sub>L a = butterfly \ (a* *\<^sub>V \)" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma cblinfun_comp_butterfly: "a o\<^sub>C\<^sub>L butterfly \ \ = butterfly (a *\<^sub>V \) \" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma butterfly_apply[simp]: "butterfly \ \' *\<^sub>V \ = \\', \\ *\<^sub>C \" by (simp add: butterfly_def scaleC_cblinfun.rep_eq) lemma butterfly_scaleC_left[simp]: "butterfly (c *\<^sub>C \) \ = c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleC_right[simp]: "butterfly \ (c *\<^sub>C \) = cnj c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleR_left[simp]: "butterfly (r *\<^sub>R \) \ = r *\<^sub>C butterfly \ \" by (simp add: scaleR_scaleC) lemma butterfly_scaleR_right[simp]: "butterfly \ (r *\<^sub>R \) = r *\<^sub>C butterfly \ \" by (simp add: butterfly_scaleC_right scaleR_scaleC) lemma butterfly_adjoint[simp]: "(butterfly \ \)* = butterfly \ \" unfolding butterfly_def by auto lemma butterfly_comp_butterfly[simp]: "butterfly \1 \2 o\<^sub>C\<^sub>L butterfly \3 \4 = \\2, \3\ *\<^sub>C butterfly \1 \4" by (simp add: butterfly_comp_cblinfun) lemma butterfly_0_left[simp]: "butterfly 0 a = 0" by (simp add: butterfly_def) lemma butterfly_0_right[simp]: "butterfly a 0 = 0" by (simp add: butterfly_def) lemma norm_butterfly: "norm (butterfly \ \) = norm \ * norm \" proof (cases "\=0") case True then show ?thesis by simp next case False show ?thesis unfolding norm_cblinfun.rep_eq thm onormI[OF _ False] proof (rule onormI[OF _ False]) fix x have "cmod \\, x\ * norm \ \ norm \ * norm \ * norm x" by (metis ab_semigroup_mult_class.mult_ac(1) complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult_left_mono norm_ge_zero) thus "norm (butterfly \ \ *\<^sub>V x) \ norm \ * norm \ * norm x" by (simp add: power2_eq_square) show "norm (butterfly \ \ *\<^sub>V \) = norm \ * norm \ * norm \" by (smt (z3) ab_semigroup_mult_class.mult_ac(1) butterfly_apply mult.commute norm_eq_sqrt_cinner norm_ge_zero norm_scaleC power2_eq_square real_sqrt_abs real_sqrt_eq_iff) qed qed lemma bounded_sesquilinear_butterfly[bounded_sesquilinear]: \bounded_sesquilinear (\(b::'b::chilbert_space) (a::'a::chilbert_space). butterfly a b)\ proof standard fix a a' :: 'a and b b' :: 'b and r :: complex show \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (rule butterfly_add_left) show \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (rule butterfly_add_right) show \butterfly (r *\<^sub>C a) b = r *\<^sub>C butterfly a b\ by simp show \butterfly a (r *\<^sub>C b) = cnj r *\<^sub>C butterfly a b\ by simp show \\K. \b a. norm (butterfly a b) \ norm b * norm a * K \ apply (rule exI[of _ 1]) by (simp add: norm_butterfly) qed lemma inj_selfbutter_upto_phase: assumes "selfbutter x = selfbutter y" shows "\c. cmod c = 1 \ x = c *\<^sub>C y" proof (cases "x = 0") case True from assms have "y = 0" using norm_butterfly by (metis True butterfly_0_left divisors_zero norm_eq_zero) with True show ?thesis using norm_one by fastforce next case False define c where "c = \y, x\ / \x, x\" have "\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x" by (simp add: butterfly_apply) also have "\ = selfbutter y *\<^sub>V x" using assms by simp also have "\ = \y, x\ *\<^sub>C y" by (simp add: butterfly_apply) finally have xcy: "x = c *\<^sub>C y" by (simp add: c_def ceq_vector_fraction_iff) have "cmod c * norm x = cmod c * norm y" using assms norm_butterfly by (smt (verit, ccfv_SIG) \\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x\ \selfbutter y *\<^sub>V x = \y, x\ *\<^sub>C y\ cinner_scaleC_right complex_vector.scale_left_commute complex_vector.scale_right_imp_eq mult_cancel_left norm_eq_sqrt_cinner norm_eq_zero scaleC_scaleC xcy) also have "cmod c * norm y = norm (c *\<^sub>C y)" by simp also have "\ = norm x" unfolding xcy[symmetric] by simp finally have c: "cmod c = 1" by (simp add: False) from c xcy show ?thesis by auto qed lemma butterfly_eq_proj: assumes "norm x = 1" shows "selfbutter x = proj x" proof - define B and \ :: "complex \\<^sub>C\<^sub>L 'a" where "B = selfbutter x" and "\ = vector_to_cblinfun x" then have B: "B = \ o\<^sub>C\<^sub>L \*" unfolding butterfly_def by simp have \adj\: "\* o\<^sub>C\<^sub>L \ = id_cblinfun" using \_def assms isometry_def isometry_vector_to_cblinfun by blast have "B o\<^sub>C\<^sub>L B = \ o\<^sub>C\<^sub>L (\* o\<^sub>C\<^sub>L \) o\<^sub>C\<^sub>L \*" by (simp add: B cblinfun_assoc_left(1)) also have "\ = B" unfolding \adj\ by (simp add: B) finally have idem: "B o\<^sub>C\<^sub>L B = B". have herm: "B = B*" unfolding B by simp from idem herm have BProj: "B = Proj (B *\<^sub>S top)" by (rule Proj_on_own_range'[symmetric]) have "B *\<^sub>S top = ccspan {x}" by (simp add: B \_def assms cblinfun_compose_image range_adjoint_isometry) with BProj show "B = proj x" by simp qed lemma butterfly_is_Proj: \norm x = 1 \ is_Proj (selfbutter x)\ by (subst butterfly_eq_proj, simp_all) lemma cspan_butterfly_UNIV: assumes \cspan basisA = UNIV\ assumes \cspan basisB = UNIV\ assumes \is_ortho_set basisB\ assumes \\b. b \ basisB \ norm b = 1\ shows \cspan {butterfly a b| (a::'a::{complex_normed_vector}) (b::'b::{chilbert_space,cfinite_dim}). a \ basisA \ b \ basisB} = UNIV\ proof - have F: \\F\{butterfly a b |a b. a \ basisA \ b \ basisB}. \b'\basisB. F *\<^sub>V b' = (if b' = b then a else 0)\ if \a \ basisA\ and \b \ basisB\ for a b apply (rule bexI[where x=\butterfly a b\]) using assms that by (auto simp: is_ortho_set_def cnorm_eq_1) show ?thesis apply (rule cblinfun_cspan_UNIV[where basisA=basisB and basisB=basisA]) using assms apply auto[2] using F by (smt (verit, ccfv_SIG) image_iff) qed lemma cindependent_butterfly: fixes basisA :: \'a::chilbert_space set\ and basisB :: \'b::chilbert_space set\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes normA: \\a. a\basisA \ norm a = 1\ and normB: \\b. b\basisB \ norm b = 1\ shows \cindependent {butterfly a b| a b. a\basisA \ b\basisB}\ proof (unfold complex_vector.independent_explicit_module, intro allI impI, rename_tac T f g) fix T :: \('b \\<^sub>C\<^sub>L 'a) set\ and f :: \'b \\<^sub>C\<^sub>L 'a \ complex\ and g :: \'b \\<^sub>C\<^sub>L 'a\ assume \finite T\ assume T_subset: \T \ {butterfly a b |a b. a \ basisA \ b \ basisB}\ define lin where \lin = (\g\T. f g *\<^sub>C g)\ assume \lin = 0\ assume \g \ T\ (* To show: f g = 0 *) then obtain a b where g: \g = butterfly a b\ and [simp]: \a \ basisA\ \b \ basisB\ using T_subset by auto have *: "(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0" if \g \ T - {butterfly a b}\ for g proof - from that obtain a' b' where g: \g = butterfly a' b'\ and [simp]: \a' \ basisA\ \b' \ basisB\ using T_subset by auto from that have \g \ butterfly a b\ by auto with g consider (a) \a\a'\ | (b) \b\b'\ by auto then show \(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0\ proof cases case a then show ?thesis using \is_ortho_set basisA\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) next case b then show ?thesis using \is_ortho_set basisB\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) qed qed have \0 = (vector_to_cblinfun a)* *\<^sub>V lin *\<^sub>V b\ using \lin = 0\ by auto also have \\ = (\g\T. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ unfolding lin_def apply (rule complex_vector.linear_sum) by (smt (z3) cblinfun.scaleC_left cblinfun.scaleC_right cblinfun.add_right clinearI plus_cblinfun.rep_eq) also have \\ = (\g\{butterfly a b}. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ apply (rule sum.mono_neutral_right) using \finite T\ * \g \ T\ g by auto also have \\ = (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b\ by (simp add: g) also have \\ = f g\ unfolding g using normA normB by (auto simp: butterfly_def scaleC_cblinfun.rep_eq cnorm_eq_1) finally show \f g = 0\ by simp qed lemma clinear_eq_butterflyI: fixes F G :: \('a::{chilbert_space,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_inner) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes \cspan basisA = UNIV\ \cspan basisB = UNIV\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes "\a b. a\basisA \ b\basisB \ F (butterfly a b) = G (butterfly a b)" assumes \\b. b\basisB \ norm b = 1\ shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_UNIV) using assms by auto subsection \Bifunctionals\ lift_definition bifunctional :: \'a::complex_normed_vector \\<^sub>C\<^sub>L (('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex)\ is \\x f. f *\<^sub>V x\ by (simp add: cblinfun.flip) lemma bifunctional_apply[simp]: \(bifunctional *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ by (transfer fixing: x f, simp) lemma bifunctional_isometric[simp]: \norm (bifunctional *\<^sub>V x) = norm x\ for x :: \'a::complex_inner\ proof - define f :: \'a \\<^sub>C\<^sub>L complex\ where \f = CBlinfun (\y. cinner x y)\ then have [simp]: \f *\<^sub>V y = cinner x y\ for y by (simp add: bounded_clinear_CBlinfun_apply bounded_clinear_cinner_right) then have [simp]: \norm f = norm x\ apply (auto intro!: norm_cblinfun_eqI[where x=x] simp: power2_norm_eq_cinner[symmetric]) apply (smt (verit, best) norm_eq_sqrt_cinner norm_ge_zero power2_norm_eq_cinner real_div_sqrt) using Cauchy_Schwarz_ineq2 by blast show ?thesis apply (auto intro!: norm_cblinfun_eqI[where x=f]) apply (metis norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) by (metis norm_cblinfun ordered_field_class.sign_simps(33)) qed lemma norm_bifunctional[simp]: \norm (bifunctional :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ proof - obtain x :: 'a where [simp]: \norm x = 1\ by (meson UNIV_not_singleton ex_norm1) show ?thesis by (auto intro!: norm_cblinfun_eqI[where x=x]) qed subsection \Banach-Steinhaus\ theorem cbanach_steinhaus: fixes F :: \'c \ 'a::cbanach \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\x. \M. \n. norm ((F n) *\<^sub>V x) \ M\ shows \\M. \ n. norm (F n) \ M\ using cblinfun_blinfun_transfer[transfer_rule] apply (rule TrueI)? (* Deletes current facts *) proof (use assms in transfer) fix F :: \'c \ 'a \\<^sub>L 'b\ assume \(\x. \M. \n. norm (F n *\<^sub>v x) \ M)\ hence \\x. bounded (range (\n. blinfun_apply (F n) x))\ by (metis (no_types, lifting) boundedI rangeE) hence \bounded (range F)\ by (simp add: banach_steinhaus) thus \\M. \n. norm (F n) \ M\ by (simp add: bounded_iff) qed subsection \Riesz-representation theorem\ theorem riesz_frechet_representation_cblinfun_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ shows \\t. \x. f *\<^sub>V x = \t, x\\ apply transfer by (rule riesz_frechet_representation_existence) lemma riesz_frechet_representation_cblinfun_unique: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::complex_inner \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ assumes \\x. f *\<^sub>V x = \u, x\\ shows \t = u\ using assms by (rule riesz_frechet_representation_unique) theorem riesz_frechet_representation_cblinfun_norm: includes notation_norm fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ shows \\f\ = \t\\ using assms proof transfer fix f::\'a \ complex\ and t assume \bounded_clinear f\ and \\x. f x = \t, x\\ from \\x. f x = \t, x\\ have \(norm (f x)) / (norm x) \ norm t\ for x proof(cases \norm x = 0\) case True thus ?thesis by simp next case False have \norm (f x) = norm (\t, x\)\ using \\x. f x = \t, x\\ by simp also have \norm \t, x\ \ norm t * norm x\ by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) finally have \norm (f x) \ norm t * norm x\ by blast thus ?thesis by (metis False linordered_field_class.divide_right_mono nonzero_mult_div_cancel_right norm_ge_zero) qed moreover have \(norm (f t)) / (norm t) = norm t\ proof(cases \norm t = 0\) case True thus ?thesis by simp next case False have \f t = \t, t\\ using \\x. f x = \t, x\\ by blast also have \\ = (norm t)^2\ by (meson cnorm_eq_square) also have \\ = (norm t)*(norm t)\ by (simp add: power2_eq_square) finally have \f t = (norm t)*(norm t)\ by blast thus ?thesis by (metis False Re_complex_of_real \\x. f x = cinner t x\ cinner_ge_zero complex_of_real_cmod nonzero_divide_eq_eq) qed ultimately have \Sup {(norm (f x)) / (norm x)| x. True} = norm t\ by (smt cSup_eq_maximum mem_Collect_eq) moreover have \Sup {(norm (f x)) / (norm x)| x. True} = (SUP x. (norm (f x)) / (norm x))\ by (simp add: full_SetCompr_eq) ultimately show \onorm f = norm t\ by (simp add: onorm_def) qed subsection \Extension of complex bounded operators\ definition cblinfun_extension where "cblinfun_extension S \ = (SOME B. \x\S. B *\<^sub>V x = \ x)" definition cblinfun_extension_exists where "cblinfun_extension_exists S \ = (\B. \x\S. B *\<^sub>V x = \ x)" lemma cblinfun_extension_existsI: assumes "\x. x\S \ B *\<^sub>V x = \ x" shows "cblinfun_extension_exists S \" using assms cblinfun_extension_exists_def by blast lemma cblinfun_extension_exists_finite_dim: fixes \::"'a::{complex_normed_vector,cfinite_dim} \ 'b::complex_normed_vector" assumes "cindependent S" and "cspan S = UNIV" shows "cblinfun_extension_exists S \" proof- define f::"'a \ 'b" where "f = complex_vector.construct S \" have "clinear f" by (simp add: complex_vector.linear_construct assms linear_construct f_def) have "bounded_clinear f" using \clinear f\ assms by auto then obtain B::"'a \\<^sub>C\<^sub>L 'b" where "B *\<^sub>V x = f x" for x using cblinfun_apply_cases by blast have "B *\<^sub>V x = \ x" if c1: "x\S" for x proof- have "B *\<^sub>V x = f x" by (simp add: \\x. B *\<^sub>V x = f x\) also have "\ = \ x" using assms complex_vector.construct_basis f_def that by (simp add: complex_vector.construct_basis) finally show?thesis by blast qed thus ?thesis unfolding cblinfun_extension_exists_def by blast qed lemma cblinfun_extension_exists_bounded_dense: fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes \closure S = UNIV\ assumes f_add: \\x y. x \ S \ y \ S \ f (x + y) = f x + f y\ assumes f_scale: \\c x y. x \ S \ f (c *\<^sub>C x) = c *\<^sub>C f x\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows \cblinfun_extension_exists S f\ proof - obtain B where bounded: \\x. x \ S \ norm (f x) \ B * norm x\ and \B > 0\ using bounded by (smt (z3) mult_mono norm_ge_zero) have \\xi. (xi \ x) \ (\i. xi i \ S)\ for x using assms(2) closure_sequential by blast then obtain seq :: \'a \ nat \ 'a\ where seq_lim: \seq x \ x\ and seq_S: \seq x i \ S\ for x i apply (atomize_elim, subst all_conj_distrib[symmetric]) apply (rule choice) by auto define g where \g x = lim (\i. f (seq x i))\ for x have \Cauchy (\i. f (seq x i))\ for x proof (rule CauchyI) fix e :: real assume \e > 0\ have \Cauchy (seq x)\ using LIMSEQ_imp_Cauchy seq_lim by blast then obtain M where less_eB: \norm (seq x m - seq x n) < e/B\ if \n \ M\ and \m \ M\ for n m apply atomize_elim by (meson CauchyD \0 < B\ \0 < e\ linordered_field_class.divide_pos_pos) have \norm (f (seq x m) - f (seq x n)) < e\ if \n \ M\ and \m \ M\ for n m proof - have \norm (f (seq x m) - f (seq x n)) = norm (f (seq x m - seq x n))\ using f_add f_scale seq_S by (metis add_diff_cancel assms(1) complex_vector.subspace_diff diff_add_cancel) also have \\ \ B * norm (seq x m - seq x n)\ apply (rule bounded) by (simp add: assms(1) complex_vector.subspace_diff seq_S) also from less_eB have \\ < B * (e/B)\ by (meson \0 < B\ linordered_semiring_strict_class.mult_strict_left_mono that) also have \\ \ e\ using \0 < B\ by auto finally show ?thesis by - qed then show \\M. \m\M. \n\M. norm (f (seq x m) - f (seq x n)) < e\ by auto qed then have f_seq_lim: \(\i. f (seq x i)) \ g x\ for x by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff g_def) have f_xi_lim: \(\i. f (xi i)) \ g x\ if \xi \ x\ and \\i. xi i \ S\ for xi x proof - from seq_lim that have \(\i. B * norm (xi i - seq x i)) \ 0\ by (metis (no_types) \0 < B\ cancel_comm_monoid_add_class.diff_cancel norm_not_less_zero norm_zero tendsto_diff tendsto_norm_zero_iff tendsto_zero_mult_left_iff) then have \(\i. f (xi i + (-1) *\<^sub>C seq x i)) \ 0\ apply (rule Lim_null_comparison[rotated]) using bounded by (simp add: assms(1) complex_vector.subspace_diff seq_S that(2)) then have \(\i. f (xi i) - f (seq x i)) \ 0\ apply (subst (asm) f_add) apply (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S)[2] apply (subst (asm) f_scale) by (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S) then show \(\i. f (xi i)) \ g x\ using Lim_transform f_seq_lim by fastforce qed have g_add: \g (x + y) = g x + g y\ for x y proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto obtain yi :: \nat \ 'a\ where \yi \ y\ and \yi i \ S\ for i using seq_S seq_lim by auto have \(\i. xi i + yi i) \ x + y\ using \xi \ x\ \yi \ y\ tendsto_add by blast then have lim1: \(\i. f (xi i + yi i)) \ g (x + y)\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ assms(1) complex_vector.subspace_add f_xi_lim) have \(\i. f (xi i + yi i)) = (\i. f (xi i) + f (yi i))\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ f_add) also have \\ \ g x + g y\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ \xi \ x\ \yi \ y\ f_xi_lim tendsto_add) finally show ?thesis using lim1 LIMSEQ_unique by blast qed have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ for c x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto have \(\i. c *\<^sub>C xi i) \ c *\<^sub>C x\ using \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at isCont_tendsto_compose by blast then have lim1: \(\i. f (c *\<^sub>C xi i)) \ g (c *\<^sub>C x)\ by (simp add: \\i. xi i \ S\ assms(1) complex_vector.subspace_scale f_xi_lim) have \(\i. f (c *\<^sub>C xi i)) = (\i. c *\<^sub>C f (xi i))\ by (simp add: \\i. xi i \ S\ f_scale) also have \\ \ c *\<^sub>C g x\ using \\i. xi i \ S\ \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at f_xi_lim isCont_tendsto_compose by blast finally show ?thesis using lim1 LIMSEQ_unique by blast qed have [simp]: \f x = g x\ if \x \ S\ for x proof - have \(\_. x) \ x\ by auto then have \(\_. f x) \ g x\ using that by (rule f_xi_lim) then show \f x = g x\ by (simp add: LIMSEQ_const_iff) qed have g_bounded: \norm (g x) \ B * norm x\ for x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto then have \(\i. f (xi i)) \ g x\ using f_xi_lim by presburger then have \(\i. norm (f (xi i))) \ norm (g x)\ by (metis tendsto_norm) moreover have \(\i. B * norm (xi i)) \ B * norm x\ by (simp add: \xi \ x\ tendsto_mult_left tendsto_norm) ultimately show \norm (g x) \ B * norm x\ apply (rule lim_mono[rotated]) using bounded using \xi _ \ S\ by blast qed have \bounded_clinear g\ using g_add g_scale apply (rule bounded_clinearI[where K=B]) using g_bounded by (simp add: ordered_field_class.sign_simps(5)) then have [simp]: \CBlinfun g *\<^sub>V x = g x\ for x by (subst CBlinfun_inverse, auto) show \cblinfun_extension_exists S f\ apply (rule cblinfun_extension_existsI[where B=\CBlinfun g\]) by auto qed lemma cblinfun_extension_apply: assumes "cblinfun_extension_exists S f" and "v \ S" shows "(cblinfun_extension S f) *\<^sub>V v = f v" by (smt assms cblinfun_extension_def cblinfun_extension_exists_def tfl_some) subsection \Notation\ bundle cblinfun_notation begin notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) notation cblinfun_apply (infixr "*\<^sub>V" 70) notation cblinfun_image (infixr "*\<^sub>S" 70) notation adj ("_*" [99] 100) end bundle no_cblinfun_notation begin no_notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) no_notation cblinfun_apply (infixr "*\<^sub>V" 70) no_notation cblinfun_image (infixr "*\<^sub>S" 70) no_notation adj ("_*" [99] 100) end bundle blinfun_notation begin notation blinfun_apply (infixr "*\<^sub>V" 70) end bundle no_blinfun_notation begin no_notation blinfun_apply (infixr "*\<^sub>V" 70) end unbundle no_cblinfun_notation +unbundle no_lattice_syntax end diff --git a/thys/Complex_Bounded_Operators/Complex_L2.thy b/thys/Complex_Bounded_Operators/Complex_L2.thy --- a/thys/Complex_Bounded_Operators/Complex_L2.thy +++ b/thys/Complex_Bounded_Operators/Complex_L2.thy @@ -1,1462 +1,1463 @@ section \\Complex_L2\ -- Hilbert space of square-summable functions\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_L2 imports Complex_Bounded_Linear_Function "HOL-Analysis.L2_Norm" "HOL-Library.Rewrite" "HOL-Analysis.Infinite_Sum" begin +unbundle lattice_syntax unbundle cblinfun_notation unbundle no_notation_blinfun_apply subsection \l2 norm of functions\ definition "has_ell2_norm (x::_\complex) \ (\i. (x i)\<^sup>2) abs_summable_on UNIV" lemma has_ell2_norm_bdd_above: \has_ell2_norm x \ bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (simp add: has_ell2_norm_def abs_summable_iff_bdd_above) lemma has_ell2_norm_L2_set: "has_ell2_norm x = bdd_above (L2_set (norm o x) ` Collect finite)" proof (rule iffI) have \mono sqrt\ using monoI real_sqrt_le_mono by blast assume \has_ell2_norm x\ then have *: \bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (subst (asm) has_ell2_norm_bdd_above) have \bdd_above ((\F. sqrt (sum (\xa. norm ((x xa)\<^sup>2)) F)) ` Collect finite)\ using bdd_above_image_mono[OF \mono sqrt\ *] by (auto simp: image_image) then show \bdd_above (L2_set (norm o x) ` Collect finite)\ by (auto simp: L2_set_def norm_power) next define p2 where \p2 x = (if x < 0 then 0 else x^2)\ for x :: real have \mono p2\ by (simp add: monoI p2_def) have [simp]: \p2 (L2_set f F) = (\i\F. (f i)\<^sup>2)\ for f and F :: \'a set\ by (smt (verit) L2_set_def L2_set_nonneg p2_def power2_less_0 real_sqrt_pow2 sum.cong sum_nonneg) assume *: \bdd_above (L2_set (norm o x) ` Collect finite)\ have \bdd_above (p2 ` L2_set (norm o x) ` Collect finite)\ using bdd_above_image_mono[OF \mono p2\ *] by auto then show \has_ell2_norm x\ apply (simp add: image_image has_ell2_norm_def abs_summable_iff_bdd_above) by (simp add: norm_power) qed definition ell2_norm :: \('a \ complex) \ real\ where \ell2_norm x = sqrt (\\<^sub>\i. norm (x i)^2)\ lemma ell2_norm_SUP: assumes \has_ell2_norm x\ shows "ell2_norm x = sqrt (SUP F\{F. finite F}. sum (\i. norm (x i)^2) F)" using assms apply (auto simp add: ell2_norm_def has_ell2_norm_def) apply (subst infsum_nonneg_is_SUPREMUM_real) by (auto simp: norm_power) lemma ell2_norm_L2_set: assumes "has_ell2_norm x" shows "ell2_norm x = (SUP F\{F. finite F}. L2_set (norm o x) F)" proof- have "sqrt (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)) = (SUP F\{F. finite F}. sqrt (\i\F. (cmod (x i))\<^sup>2))" proof (subst continuous_at_Sup_mono) show "mono sqrt" by (simp add: mono_def) show "continuous (at_left (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite))) sqrt" using continuous_at_split isCont_real_sqrt by blast show "sum (\i. (cmod (x i))\<^sup>2) ` Collect finite \ {}" by auto show "bdd_above (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)" using has_ell2_norm_bdd_above[THEN iffD1, OF assms] by (auto simp: norm_power) show "\ (sqrt ` sum (\i. (cmod (x i))\<^sup>2) ` Collect finite) = (SUP F\Collect finite. sqrt (\i\F. (cmod (x i))\<^sup>2))" by (metis image_image) qed thus ?thesis using assms by (auto simp: ell2_norm_SUP L2_set_def) qed lemma has_ell2_norm_finite[simp]: "has_ell2_norm (x::'a::finite\_)" unfolding has_ell2_norm_def by simp lemma ell2_norm_finite: "ell2_norm (x::'a::finite\complex) = sqrt (sum (\i. (norm(x i))^2) UNIV)" by (simp add: ell2_norm_def) lemma ell2_norm_finite_L2_set: "ell2_norm (x::'a::finite\complex) = L2_set (norm o x) UNIV" by (simp add: ell2_norm_finite L2_set_def) lemma ell2_ket: fixes a defines \f \ (\i. if a = i then 1 else 0)\ shows has_ell2_norm_ket: \has_ell2_norm f\ and ell2_norm_ket: \ell2_norm f = 1\ proof - have \(\x. (f x)\<^sup>2) abs_summable_on {a}\ apply (rule summable_on_finite) by simp then show \has_ell2_norm f\ unfolding has_ell2_norm_def apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1]) unfolding f_def by auto have \(\\<^sub>\x\{a}. (f x)\<^sup>2) = 1\ apply (subst infsum_finite) by (auto simp: f_def) then show \ell2_norm f = 1\ unfolding ell2_norm_def apply (subst infsum_cong_neutral[where T=\{a}\ and g=\\x. (cmod (f x))\<^sup>2\]) by (auto simp: f_def) qed lemma ell2_norm_geq0: \ell2_norm x \ 0\ by (auto simp: ell2_norm_def intro!: infsum_nonneg) lemma ell2_norm_point_bound: assumes \has_ell2_norm x\ shows \ell2_norm x \ cmod (x i)\ proof - have \(cmod (x i))\<^sup>2 = norm ((x i)\<^sup>2)\ by (simp add: norm_power) also have \norm ((x i)\<^sup>2) = sum (\i. (norm ((x i)\<^sup>2))) {i}\ by auto also have \\ = infsum (\i. (norm ((x i)\<^sup>2))) {i}\ by (rule infsum_finite[symmetric], simp) also have \\ \ infsum (\i. (norm ((x i)\<^sup>2))) UNIV\ apply (rule infsum_mono_neutral) using assms by (auto simp: has_ell2_norm_def) also have \\ = (ell2_norm x)\<^sup>2\ by (metis (no_types, lifting) ell2_norm_def ell2_norm_geq0 infsum_cong norm_power real_sqrt_eq_iff real_sqrt_unique) finally show ?thesis using ell2_norm_geq0 power2_le_imp_le by blast qed lemma ell2_norm_0: assumes "has_ell2_norm x" shows "(ell2_norm x = 0) = (x = (\_. 0))" proof assume u1: "x = (\_. 0)" have u2: "(SUP x::'a set\Collect finite. (0::real)) = 0" if "x = (\_. 0)" by (metis cSUP_const empty_Collect_eq finite.emptyI) show "ell2_norm x = 0" unfolding ell2_norm_def using u1 u2 by auto next assume norm0: "ell2_norm x = 0" show "x = (\_. 0)" proof fix i have \cmod (x i) \ ell2_norm x\ using assms by (rule ell2_norm_point_bound) also have \\ = 0\ by (fact norm0) finally show "x i = 0" by auto qed qed lemma ell2_norm_smult: assumes "has_ell2_norm x" shows "has_ell2_norm (\i. c * x i)" and "ell2_norm (\i. c * x i) = cmod c * ell2_norm x" proof - have L2_set_mul: "L2_set (cmod \ (\i. c * x i)) F = cmod c * L2_set (cmod \ x) F" for F proof- have "L2_set (cmod \ (\i. c * x i)) F = L2_set (\i. (cmod c * (cmod o x) i)) F" by (metis comp_def norm_mult) also have "\ = cmod c * L2_set (cmod o x) F" by (metis norm_ge_zero L2_set_right_distrib) finally show ?thesis . qed from assms obtain M where M: "M \ L2_set (cmod o x) F" if "finite F" for F unfolding has_ell2_norm_L2_set bdd_above_def by auto hence "cmod c * M \ L2_set (cmod o (\i. c * x i)) F" if "finite F" for F unfolding L2_set_mul by (simp add: ordered_comm_semiring_class.comm_mult_left_mono that) thus has: "has_ell2_norm (\i. c * x i)" unfolding has_ell2_norm_L2_set bdd_above_def using L2_set_mul[symmetric] by auto have "ell2_norm (\i. c * x i) = (SUP F \ Collect finite. (L2_set (cmod \ (\i. c * x i)) F))" by (simp add: ell2_norm_L2_set has) also have "\ = (SUP F \ Collect finite. (cmod c * L2_set (cmod \ x) F))" using L2_set_mul by auto also have "\ = cmod c * ell2_norm x" proof (subst ell2_norm_L2_set) show "has_ell2_norm x" by (simp add: assms) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = cmod c * \ (L2_set (cmod \ x) ` Collect finite)" proof (subst continuous_at_Sup_mono [where f = "\x. cmod c * x"]) show "mono ((*) (cmod c))" by (simp add: mono_def ordered_comm_semiring_class.comm_mult_left_mono) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) ((*) (cmod c))" proof (rule continuous_mult) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. cmod c)" by simp show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. x)" by simp qed show "L2_set (cmod \ x) ` Collect finite \ {}" by auto show "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms has_ell2_norm_L2_set) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = \ ((*) (cmod c) ` L2_set (cmod \ x) ` Collect finite)" by (metis image_image) qed qed finally show "ell2_norm (\i. c * x i) = cmod c * ell2_norm x". qed lemma ell2_norm_triangle: assumes "has_ell2_norm x" and "has_ell2_norm y" shows "has_ell2_norm (\i. x i + y i)" and "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" proof - have triangle: "L2_set (cmod \ (\i. x i + y i)) F \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" (is "?lhs\?rhs") if "finite F" for F proof - have "?lhs \ L2_set (\i. (cmod o x) i + (cmod o y) i) F" proof (rule L2_set_mono) show "(cmod \ (\i. x i + y i)) i \ (cmod \ x) i + (cmod \ y) i" if "i \ F" for i :: 'a using that norm_triangle_ineq by auto show "0 \ (cmod \ (\i. x i + y i)) i" if "i \ F" for i :: 'a using that by simp qed also have "\ \ ?rhs" by (rule L2_set_triangle_ineq) finally show ?thesis . qed obtain Mx My where Mx: "Mx \ L2_set (cmod o x) F" and My: "My \ L2_set (cmod o y) F" if "finite F" for F using assms unfolding has_ell2_norm_L2_set bdd_above_def by auto hence MxMy: "Mx + My \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" if "finite F" for F using that by fastforce hence bdd_plus: "bdd_above ((\xa. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa) ` Collect finite)" unfolding bdd_above_def by auto from MxMy have MxMy': "Mx + My \ L2_set (cmod \ (\i. x i + y i)) F" if "finite F" for F using triangle that by fastforce thus has: "has_ell2_norm (\i. x i + y i)" unfolding has_ell2_norm_L2_set bdd_above_def by auto have SUP_plus: "(SUP x\A. f x + g x) \ (SUP x\A. f x) + (SUP x\A. g x)" if notempty: "A\{}" and bddf: "bdd_above (f`A)"and bddg: "bdd_above (g`A)" for f g :: "'a set \ real" and A proof- have xleq: "x \ (SUP x\A. f x) + (SUP x\A. g x)" if x: "x \ (\x. f x + g x) ` A" for x proof - obtain a where aA: "a:A" and ax: "x = f a + g a" using x by blast have fa: "f a \ (SUP x\A. f x)" by (simp add: bddf aA cSUP_upper) moreover have "g a \ (SUP x\A. g x)" by (simp add: bddg aA cSUP_upper) ultimately have "f a + g a \ (SUP x\A. f x) + (SUP x\A. g x)" by simp with ax show ?thesis by simp qed have "(\x. f x + g x) ` A \ {}" using notempty by auto moreover have "x \ \ (f ` A) + \ (g ` A)" if "x \ (\x. f x + g x) ` A" for x :: real using that by (simp add: xleq) ultimately show ?thesis by (meson bdd_above_def cSup_le_iff) qed have a2: "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms(1) has_ell2_norm_L2_set) have a3: "bdd_above (L2_set (cmod \ y) ` Collect finite)" by (meson assms(2) has_ell2_norm_L2_set) have a1: "Collect finite \ {}" by auto have a4: "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ (SUP xa\Collect finite. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa)" by (metis (mono_tags, lifting) a1 bdd_plus cSUP_mono mem_Collect_eq triangle) have "\r. \ (L2_set (cmod \ (\a. x a + y a)) ` Collect finite) \ r \ \ (SUP A\Collect finite. L2_set (cmod \ x) A + L2_set (cmod \ y) A) \ r" using a4 by linarith hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ \ (L2_set (cmod \ x) ` Collect finite) + \ (L2_set (cmod \ y) ` Collect finite)" by (metis (no_types) SUP_plus a1 a2 a3) hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ ell2_norm x + ell2_norm y" by (simp add: assms(1) assms(2) ell2_norm_L2_set) thus "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" by (simp add: ell2_norm_L2_set has) qed lemma ell2_norm_uminus: assumes "has_ell2_norm x" shows \has_ell2_norm (\i. - x i)\ and \ell2_norm (\i. - x i) = ell2_norm x\ using assms by (auto simp: has_ell2_norm_def ell2_norm_def) subsection \The type \ell2\ of square-summable functions\ typedef 'a ell2 = "{x::'a\complex. has_ell2_norm x}" unfolding has_ell2_norm_def by (rule exI[of _ "\_.0"], auto) setup_lifting type_definition_ell2 instantiation ell2 :: (type)complex_vector begin lift_definition zero_ell2 :: "'a ell2" is "\_. 0" by (auto simp: has_ell2_norm_def) lift_definition uminus_ell2 :: "'a ell2 \ 'a ell2" is uminus by (simp add: has_ell2_norm_def) lift_definition plus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x + g x" by (rule ell2_norm_triangle) lift_definition minus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x - g x" apply (subst add_uminus_conv_diff[symmetric]) apply (rule ell2_norm_triangle) by (auto simp add: ell2_norm_uminus) lift_definition scaleR_ell2 :: "real \ 'a ell2 \ 'a ell2" is "\r f x. complex_of_real r * f x" by (rule ell2_norm_smult) lift_definition scaleC_ell2 :: "complex \ 'a ell2 \ 'a ell2" is "\c f x. c * f x" by (rule ell2_norm_smult) instance proof fix a b c :: "'a ell2" show "((*\<^sub>R) r::'a ell2 \ _) = (*\<^sub>C) (complex_of_real r)" for r apply (rule ext) apply transfer by auto show "a + b + c = a + (b + c)" by (transfer; rule ext; simp) show "a + b = b + a" by (transfer; rule ext; simp) show "0 + a = a" by (transfer; rule ext; simp) show "- a + a = 0" by (transfer; rule ext; simp) show "a - b = a + - b" by (transfer; rule ext; simp) show "r *\<^sub>C (a + b) = r *\<^sub>C a + r *\<^sub>C b" for r apply (transfer; rule ext) by (simp add: vector_space_over_itself.scale_right_distrib) show "(r + r') *\<^sub>C a = r *\<^sub>C a + r' *\<^sub>C a" for r r' apply (transfer; rule ext) by (simp add: ring_class.ring_distribs(2)) show "r *\<^sub>C r' *\<^sub>C a = (r * r') *\<^sub>C a" for r r' by (transfer; rule ext; simp) show "1 *\<^sub>C a = a" by (transfer; rule ext; simp) qed end instantiation ell2 :: (type)complex_normed_vector begin lift_definition norm_ell2 :: "'a ell2 \ real" is ell2_norm . declare norm_ell2_def[code del] definition "dist x y = norm (x - y)" for x y::"'a ell2" definition "sgn x = x /\<^sub>R norm x" for x::"'a ell2" definition [code del]: "uniformity = (INF e\{0<..}. principal {(x::'a ell2, y). norm (x - y) < e})" definition [code del]: "open U = (\x\U. \\<^sub>F (x', y) in INF e\{0<..}. principal {(x, y). norm (x - y) < e}. x' = x \ y \ U)" for U :: "'a ell2 set" instance proof fix a b :: "'a ell2" show "dist a b = norm (a - b)" by (simp add: dist_ell2_def) show "sgn a = a /\<^sub>R norm a" by (simp add: sgn_ell2_def) show "uniformity = (INF e\{0<..}. principal {(x, y). dist (x::'a ell2) y < e})" unfolding dist_ell2_def uniformity_ell2_def by simp show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. (x'::'a ell2) = x \ y \ U)" for U :: "'a ell2 set" unfolding uniformity_ell2_def open_ell2_def by simp_all show "(norm a = 0) = (a = 0)" apply transfer by (fact ell2_norm_0) show "norm (a + b) \ norm a + norm b" apply transfer by (fact ell2_norm_triangle) show "norm (r *\<^sub>R (a::'a ell2)) = \r\ * norm a" for r and a :: "'a ell2" apply transfer by (simp add: ell2_norm_smult(2)) show "norm (r *\<^sub>C a) = cmod r * norm a" for r apply transfer by (simp add: ell2_norm_smult(2)) qed end lemma norm_point_bound_ell2: "norm (Rep_ell2 x i) \ norm x" apply transfer by (simp add: ell2_norm_point_bound) lemma ell2_norm_finite_support: assumes \finite S\ \\ i. i \ S \ Rep_ell2 x i = 0\ shows \norm x = sqrt ((sum (\i. (cmod (Rep_ell2 x i))\<^sup>2)) S)\ proof (insert assms(2), transfer fixing: S) fix x :: \'a \ complex\ assume zero: \\i. i \ S \ x i = 0\ have \ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)\ by (auto simp: ell2_norm_def) also have \\ = sqrt (\\<^sub>\i\S. (cmod (x i))\<^sup>2)\ apply (subst infsum_cong_neutral[where g=\\i. (cmod (x i))\<^sup>2\ and S=UNIV and T=S]) using zero by auto also have \\ = sqrt (\i\S. (cmod (x i))\<^sup>2)\ using \finite S\ by simp finally show \ell2_norm x = sqrt (\i\S. (cmod (x i))\<^sup>2)\ by - qed instantiation ell2 :: (type) complex_inner begin lift_definition cinner_ell2 :: "'a ell2 \ 'a ell2 \ complex" is "\x y. infsum (\i. (cnj (x i) * y i)) UNIV" . declare cinner_ell2_def[code del] instance proof standard fix x y z :: "'a ell2" fix c :: complex show "cinner x y = cnj (cinner y x)" proof transfer fix x y :: "'a\complex" assume "has_ell2_norm x" and "has_ell2_norm y" have "(\\<^sub>\i. cnj (x i) * y i) = (\\<^sub>\i. cnj (cnj (y i) * x i))" by (metis complex_cnj_cnj complex_cnj_mult mult.commute) also have "\ = cnj (\\<^sub>\i. cnj (y i) * x i)" by (metis infsum_cnj) finally show "(\\<^sub>\i. cnj (x i) * y i) = cnj (\\<^sub>\i. cnj (y i) * x i)" . qed show "cinner (x + y) z = cinner x z + cinner y z" proof transfer fix x y z :: "'a \ complex" assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence cnj_y: "(\i. cnj (y i) * cnj (y i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm z" hence z: "(\i. z i * z i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_z:"(\i. cnj (x i) * z i) abs_summable_on UNIV" using cnj_x z by (rule abs_summable_product) have cnj_y_z:"(\i. cnj (y i) * z i) abs_summable_on UNIV" using cnj_y z by (rule abs_summable_product) show "(\\<^sub>\i. cnj (x i + y i) * z i) = (\\<^sub>\i. cnj (x i) * z i) + (\\<^sub>\i. cnj (y i) * z i)" apply (subst infsum_add [symmetric]) using cnj_x_z cnj_y_z by (auto simp add: summable_on_iff_abs_summable_on_complex distrib_left mult.commute) qed show "cinner (c *\<^sub>C x) y = cnj c * cinner x y" proof transfer fix x y :: "'a \ complex" and c :: complex assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence y: "(\i. y i * y i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_y:"(\i. cnj (x i) * y i) abs_summable_on UNIV" using cnj_x y by (rule abs_summable_product) thus "(\\<^sub>\i. cnj (c * x i) * y i) = cnj c * (\\<^sub>\i. cnj (x i) * y i)" by (auto simp flip: infsum_cmult_right simp add: abs_summable_summable mult.commute vector_space_over_itself.scale_left_commute) qed show "0 \ cinner x x" proof transfer fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence "(\i. cnj (x i) * x i) abs_summable_on UNIV" by auto hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square. have "0 = (\\<^sub>\i::'a. 0)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_complex) by (auto simp add: abs_summable_summable sum) finally show "0 \ (\\<^sub>\i. cnj (x i) * x i)" by assumption qed show "(cinner x x = 0) = (x = 0)" proof (transfer, auto) fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (smt (verit, del_insts) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_ge_zero norm_power real_norm_def summable_on_cong) hence cmod_x2: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square by simp assume eq0: "(\\<^sub>\i. cnj (x i) * x i) = 0" show "x = (\_. 0)" proof (rule ccontr) assume "x \ (\_. 0)" then obtain i where "x i \ 0" by auto hence "0 < cnj (x i) * x i" by (metis le_less cnj_x_x_geq0 complex_cnj_zero_iff vector_space_over_itself.scale_eq_0_iff) also have "\ = (\\<^sub>\i\{i}. cnj (x i) * x i)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_neutral_complex) by (auto simp add: abs_summable_summable cmod_x2) also from eq0 have "\ = 0" by assumption finally show False by simp qed qed show "norm x = sqrt (cmod (cinner x x))" proof transfer fix x :: "'a \ complex" assume x: "has_ell2_norm x" have "(\i::'a. cmod (x i) * cmod (x i)) abs_summable_on UNIV \ (\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" by (metis (no_types, lifting) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_power summable_on_cong x) from x have "ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)" unfolding ell2_norm_def by simp also have "\ = sqrt (\\<^sub>\i. cmod (cnj (x i) * x i))" unfolding norm_complex_def power2_eq_square by auto also have "\ = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by (auto simp: infsum_cmod abs_summable_summable sum) finally show "ell2_norm x = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by assumption qed qed end instance ell2 :: (type) chilbert_space proof fix X :: \nat \ 'a ell2\ define x where \x n a = Rep_ell2 (X n) a\ for n a have [simp]: \has_ell2_norm (x n)\ for n using Rep_ell2 x_def[abs_def] by simp assume \Cauchy X\ moreover have "dist (x n a) (x m a) \ dist (X n) (X m)" for n m a by (metis Rep_ell2 x_def dist_norm ell2_norm_point_bound mem_Collect_eq minus_ell2.rep_eq norm_ell2.rep_eq) ultimately have \Cauchy (\n. x n a)\ for a by (meson Cauchy_def le_less_trans) then obtain l where x_lim: \(\n. x n a) \ l a\ for a apply atomize_elim apply (rule choice) by (simp add: convergent_eq_Cauchy) define L where \L = Abs_ell2 l\ define normF where \normF F x = L2_set (cmod \ x) F\ for F :: \'a set\ and x have normF_triangle: \normF F (\a. x a + y a) \ normF F x + normF F y\ if \finite F\ for F x y proof - have \normF F (\a. x a + y a) = L2_set (\a. cmod (x a + y a)) F\ by (metis (mono_tags, lifting) L2_set_cong comp_apply normF_def) also have \\ \ L2_set (\a. cmod (x a) + cmod (y a)) F\ by (meson L2_set_mono norm_ge_zero norm_triangle_ineq) also have \\ \ L2_set (\a. cmod (x a)) F + L2_set (\a. cmod (y a)) F\ by (simp add: L2_set_triangle_ineq) also have \\ \ normF F x + normF F y\ by (smt (verit, best) L2_set_cong normF_def comp_apply) finally show ?thesis by - qed have normF_negate: \normF F (\a. - x a) = normF F x\ if \finite F\ for F x unfolding normF_def o_def by simp have normF_ell2norm: \normF F x \ ell2_norm x\ if \finite F\ and \has_ell2_norm x\ for F x apply (auto intro!: cSUP_upper2[where x=F] simp: that normF_def ell2_norm_L2_set) by (meson has_ell2_norm_L2_set that(2)) note Lim_bounded2[rotated, rule_format, trans] from \Cauchy X\ obtain I where cauchyX: \norm (X n - X m) \ \\ if \\>0\ \n\I \\ \m\I \\ for \ n m by (metis Cauchy_def dist_norm less_eq_real_def) have normF_xx: \normF F (\a. x n a - x m a) \ \\ if \finite F\ \\>0\ \n\I \\ \m\I \\ for \ n m F apply (subst asm_rl[of \(\a. x n a - x m a) = Rep_ell2 (X n - X m)\]) apply (simp add: x_def minus_ell2.rep_eq) using that cauchyX by (metis Rep_ell2 mem_Collect_eq normF_ell2norm norm_ell2.rep_eq order_trans) have normF_xl_lim: \(\m. normF F (\a. x m a - l a)) \ 0\ if \finite F\ for F proof - have \(\xa. cmod (x xa m - l m)) \ 0\ for m using x_lim by (simp add: LIM_zero_iff tendsto_norm_zero) then have \(\m. \i\F. ((cmod \ (\a. x m a - l a)) i)\<^sup>2) \ 0\ by (auto intro: tendsto_null_sum) then show ?thesis unfolding normF_def L2_set_def using tendsto_real_sqrt by force qed have normF_xl: \normF F (\a. x n a - l a) \ \\ if \n \ I \\ and \\ > 0\ and \finite F\ for n \ F proof - have \normF F (\a. x n a - l a) - \ \ normF F (\a. x n a - x m a) + normF F (\a. x m a - l a) - \\ for m using normF_triangle[OF \finite F\, where x=\(\a. x n a - x m a)\ and y=\(\a. x m a - l a)\] by auto also have \\ m \ normF F (\a. x m a - l a)\ if \m \ I \\ for m using normF_xx[OF \finite F\ \\>0\ \n \ I \\ \m \ I \\] by auto also have \(\m. \ m) \ 0\ using \finite F\ by (rule normF_xl_lim) finally show ?thesis by auto qed have \normF F l \ 1 + normF F (x (I 1))\ if [simp]: \finite F\ for F using normF_xl[where F=F and \=1 and n=\I 1\] using normF_triangle[where F=F and x=\x (I 1)\ and y=\\a. l a - x (I 1) a\] using normF_negate[where F=F and x=\(\a. x (I 1) a - l a)\] by auto also have \\ F \ 1 + ell2_norm (x (I 1))\ if \finite F\ for F using normF_ell2norm that by simp finally have [simp]: \has_ell2_norm l\ unfolding has_ell2_norm_L2_set by (auto intro!: bdd_aboveI simp flip: normF_def) then have \l = Rep_ell2 L\ by (simp add: Abs_ell2_inverse L_def) have [simp]: \has_ell2_norm (\a. x n a - l a)\ for n apply (subst diff_conv_add_uminus) apply (rule ell2_norm_triangle) by (auto intro!: ell2_norm_uminus) from normF_xl have ell2norm_xl: \ell2_norm (\a. x n a - l a) \ \\ if \n \ I \\ and \\ > 0\ for n \ apply (subst ell2_norm_L2_set) using that by (auto intro!: cSUP_least simp: normF_def) have \norm (X n - L) \ \\ if \n \ I \\ and \\ > 0\ for n \ using ell2norm_xl[OF that] by (simp add: x_def norm_ell2.rep_eq \l = Rep_ell2 L\ minus_ell2.rep_eq) then have \X \ L\ unfolding tendsto_iff apply (auto simp: dist_norm eventually_sequentially) by (meson field_lbound_gt_zero le_less_trans) then show \convergent X\ by (rule convergentI) qed instantiation ell2 :: (CARD_1) complex_algebra_1 begin lift_definition one_ell2 :: "'a ell2" is "\_. 1" by simp lift_definition times_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x * b x" by simp instance proof fix a b c :: "'a ell2" and r :: complex show "a * b * c = a * (b * c)" by (transfer, auto) show "(a + b) * c = a * c + b * c" apply (transfer, rule ext) by (simp add: distrib_left mult.commute) show "a * (b + c) = a * b + a * c" apply transfer by (simp add: ring_class.ring_distribs(1)) show "r *\<^sub>C a * b = r *\<^sub>C (a * b)" by (transfer, auto) show "(a::'a ell2) * r *\<^sub>C b = r *\<^sub>C (a * b)" by (transfer, auto) show "1 * a = a" by (transfer, rule ext, auto) show "a * 1 = a" by (transfer, rule ext, auto) show "(0::'a ell2) \ 1" apply transfer by (meson zero_neq_one) qed end instantiation ell2 :: (CARD_1) field begin lift_definition divide_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x / b x" by simp lift_definition inverse_ell2 :: "'a ell2 \ 'a ell2" is "\a x. inverse (a x)" by simp instance proof (intro_classes; transfer) fix a :: "'a \ complex" assume "a \ (\_. 0)" then obtain y where ay: "a y \ 0" by auto show "(\x. inverse (a x) * a x) = (\_. 1)" proof (rule ext) fix x have "x = y" by auto with ay have "a x \ 0" by metis then show "inverse (a x) * a x = 1" by auto qed qed (auto simp add: divide_complex_def mult.commute ring_class.ring_distribs) end subsection \Orthogonality\ lemma ell2_pointwise_ortho: assumes \\ i. Rep_ell2 x i = 0 \ Rep_ell2 y i = 0\ shows \is_orthogonal x y\ using assms apply transfer by (simp add: infsum_0) subsection \Truncated vectors\ lift_definition trunc_ell2:: \'a set \ 'a ell2 \ 'a ell2\ is \\ S x. (\ i. (if i \ S then x i else 0))\ proof (rename_tac S x) fix x :: \'a \ complex\ and S :: \'a set\ assume \has_ell2_norm x\ then have \(\i. (x i)\<^sup>2) abs_summable_on UNIV\ unfolding has_ell2_norm_def by - then have \(\i. (x i)\<^sup>2) abs_summable_on S\ using summable_on_subset_banach by blast then have \(\xa. (if xa \ S then x xa else 0)\<^sup>2) abs_summable_on UNIV\ apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1]) by auto then show \has_ell2_norm (\i. if i \ S then x i else 0)\ unfolding has_ell2_norm_def by - qed lemma trunc_ell2_empty[simp]: \trunc_ell2 {} x = 0\ apply transfer by simp lemma norm_id_minus_trunc_ell2: \(norm (x - trunc_ell2 S x))^2 = (norm x)^2 - (norm (trunc_ell2 S x))^2\ proof- have \Rep_ell2 (trunc_ell2 S x) i = 0 \ Rep_ell2 (x - trunc_ell2 S x) i = 0\ for i apply transfer by auto hence \\ (trunc_ell2 S x), (x - trunc_ell2 S x) \ = 0\ using ell2_pointwise_ortho by blast hence \(norm x)^2 = (norm (trunc_ell2 S x))^2 + (norm (x - trunc_ell2 S x))^2\ using pythagorean_theorem by fastforce thus ?thesis by simp qed lemma norm_trunc_ell2_finite: \finite S \ (norm (trunc_ell2 S x)) = sqrt ((sum (\i. (cmod (Rep_ell2 x i))\<^sup>2)) S)\ proof- assume \finite S\ moreover have \\ i. i \ S \ Rep_ell2 ((trunc_ell2 S x)) i = 0\ by (simp add: trunc_ell2.rep_eq) ultimately have \(norm (trunc_ell2 S x)) = sqrt ((sum (\i. (cmod (Rep_ell2 ((trunc_ell2 S x)) i))\<^sup>2)) S)\ using ell2_norm_finite_support by blast moreover have \\ i. i \ S \ Rep_ell2 ((trunc_ell2 S x)) i = Rep_ell2 x i\ by (simp add: trunc_ell2.rep_eq) ultimately show ?thesis by simp qed lemma trunc_ell2_lim_at_UNIV: \((\S. trunc_ell2 S \) \ \) (finite_subsets_at_top UNIV)\ proof - define f where \f i = (cmod (Rep_ell2 \ i))\<^sup>2\ for i have has: \has_ell2_norm (Rep_ell2 \)\ using Rep_ell2 by blast then have summable: "f abs_summable_on UNIV" by (smt (verit, del_insts) f_def has_ell2_norm_def norm_ge_zero norm_power real_norm_def summable_on_cong) have \norm \ = (ell2_norm (Rep_ell2 \))\ apply transfer by simp also have \\ = sqrt (infsum f UNIV)\ by (simp add: ell2_norm_def f_def[symmetric]) finally have norm\: \norm \ = sqrt (infsum f UNIV)\ by - have norm_trunc: \norm (trunc_ell2 S \) = sqrt (sum f S)\ if \finite S\ for S using f_def that norm_trunc_ell2_finite by fastforce have \(sum f \ infsum f UNIV) (finite_subsets_at_top UNIV)\ using f_def[abs_def] infsum_tendsto local.summable by fastforce then have \((\S. sqrt (sum f S)) \ sqrt (infsum f UNIV)) (finite_subsets_at_top UNIV)\ using tendsto_real_sqrt by blast then have \((\S. norm (trunc_ell2 S \)) \ norm \) (finite_subsets_at_top UNIV)\ apply (subst tendsto_cong[where g=\\S. sqrt (sum f S)\]) by (auto simp add: eventually_finite_subsets_at_top_weakI norm_trunc norm\) then have \((\S. (norm (trunc_ell2 S \))\<^sup>2) \ (norm \)\<^sup>2) (finite_subsets_at_top UNIV)\ by (simp add: tendsto_power) then have \((\S. (norm \)\<^sup>2 - (norm (trunc_ell2 S \))\<^sup>2) \ 0) (finite_subsets_at_top UNIV)\ apply (rule tendsto_diff[where a=\(norm \)^2\ and b=\(norm \)^2\, simplified, rotated]) by auto then have \((\S. (norm (\ - trunc_ell2 S \))\<^sup>2) \ 0) (finite_subsets_at_top UNIV)\ unfolding norm_id_minus_trunc_ell2 by simp then have \((\S. norm (\ - trunc_ell2 S \)) \ 0) (finite_subsets_at_top UNIV)\ by auto then have \((\S. \ - trunc_ell2 S \) \ 0) (finite_subsets_at_top UNIV)\ by (rule tendsto_norm_zero_cancel) then show ?thesis apply (rule Lim_transform2[where f=\\_. \\, rotated]) by simp qed subsection \Kets and bras\ lift_definition ket :: "'a \ 'a ell2" is "\x y. if x=y then 1 else 0" by (rule has_ell2_norm_ket) abbreviation bra :: "'a \ (_,complex) cblinfun" where "bra i \ vector_to_cblinfun (ket i)*" for i instance ell2 :: (type) not_singleton proof standard have "ket undefined \ (0::'a ell2)" proof transfer show "(\y. if (undefined::'a) = y then 1::complex else 0) \ (\_. 0)" by (meson one_neq_zero) qed thus \\x y::'a ell2. x \ y\ by blast qed lemma cinner_ket_left: \\ket i, \\ = Rep_ell2 \ i\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto lemma cinner_ket_right: \\\, ket i\ = cnj (Rep_ell2 \ i)\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto lemma cinner_ket_eqI: assumes \\i. cinner (ket i) \ = cinner (ket i) \\ shows \\ = \\ by (metis Rep_ell2_inject assms cinner_ket_left ext) lemma norm_ket[simp]: "norm (ket i) = 1" apply transfer by (rule ell2_norm_ket) lemma cinner_ket_same[simp]: \\ket i, ket i\ = 1\ proof- have \norm (ket i) = 1\ by simp hence \sqrt (cmod \ket i, ket i\) = 1\ by (metis norm_eq_sqrt_cinner) hence \cmod \ket i, ket i\ = 1\ using real_sqrt_eq_1_iff by blast moreover have \\ket i, ket i\ = cmod \ket i, ket i\\ proof- have \\ket i, ket i\ \ \\ by (simp add: cinner_real) thus ?thesis by (metis cinner_ge_zero complex_of_real_cmod) qed ultimately show ?thesis by simp qed lemma orthogonal_ket[simp]: \is_orthogonal (ket i) (ket j) \ i \ j\ by (simp add: cinner_ket_left ket.rep_eq) lemma cinner_ket: \\ket i, ket j\ = (if i=j then 1 else 0)\ by (simp add: cinner_ket_left ket.rep_eq) lemma ket_injective[simp]: \ket i = ket j \ i = j\ by (metis cinner_ket one_neq_zero) lemma inj_ket[simp]: \inj ket\ by (simp add: inj_on_def) lemma trunc_ell2_ket_cspan: \trunc_ell2 S x \ (cspan (range ket))\ if \finite S\ proof (use that in induction) case empty then show ?case by (auto intro: complex_vector.span_zero) next case (insert a F) from insert.hyps have \trunc_ell2 (insert a F) x = trunc_ell2 F x + Rep_ell2 x a *\<^sub>C ket a\ apply (transfer fixing: F a) by auto with insert.IH show ?case by (simp add: complex_vector.span_add_eq complex_vector.span_base complex_vector.span_scale) qed lemma closed_cspan_range_ket[simp]: \closure (cspan (range ket)) = UNIV\ proof (intro set_eqI iffI UNIV_I closure_approachable[THEN iffD2] allI impI) fix \ :: \'a ell2\ fix e :: real assume \e > 0\ have \((\S. trunc_ell2 S \) \ \) (finite_subsets_at_top UNIV)\ by (rule trunc_ell2_lim_at_UNIV) then obtain F where \finite F\ and \dist (trunc_ell2 F \) \ < e\ apply (drule_tac tendstoD[OF _ \e > 0\]) by (auto dest: simp: eventually_finite_subsets_at_top) moreover have \trunc_ell2 F \ \ cspan (range ket)\ using \finite F\ trunc_ell2_ket_cspan by blast ultimately show \\\\cspan (range ket). dist \ \ < e\ by auto qed lemma ccspan_range_ket[simp]: "ccspan (range ket) = (top::('a ell2 ccsubspace))" proof- have \closure (complex_vector.span (range ket)) = (UNIV::'a ell2 set)\ using Complex_L2.closed_cspan_range_ket by blast thus ?thesis by (simp add: ccspan.abs_eq top_ccsubspace.abs_eq) qed lemma cspan_range_ket_finite[simp]: "cspan (range ket :: 'a::finite ell2 set) = UNIV" by (metis closed_cspan_range_ket closure_finite_cspan finite_class.finite_UNIV finite_imageI) instance ell2 :: (finite) cfinite_dim proof define basis :: \'a ell2 set\ where \basis = range ket\ have \finite basis\ unfolding basis_def by simp moreover have \cspan basis = UNIV\ by (simp add: basis_def) ultimately show \\basis::'a ell2 set. finite basis \ cspan basis = UNIV\ by auto qed instantiation ell2 :: (enum) onb_enum begin definition "canonical_basis_ell2 = map ket Enum.enum" instance proof show "distinct (canonical_basis::'a ell2 list)" proof- have \finite (UNIV::'a set)\ by simp have \distinct (enum_class.enum::'a list)\ using enum_distinct by blast moreover have \inj_on ket (set enum_class.enum)\ by (meson inj_onI ket_injective) ultimately show ?thesis unfolding canonical_basis_ell2_def using distinct_map by blast qed show "is_ortho_set (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (z3) norm_ket f_inv_into_f is_ortho_set_def orthogonal_ket norm_zero) show "cindependent (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (verit, best) norm_ket f_inv_into_f is_ortho_set_def is_ortho_set_cindependent orthogonal_ket norm_zero) show "cspan (set (canonical_basis::'a ell2 list)) = UNIV" by (auto simp: canonical_basis_ell2_def enum_UNIV) show "norm (x::'a ell2) = 1" if "(x::'a ell2) \ set canonical_basis" for x :: "'a ell2" using that unfolding canonical_basis_ell2_def by auto qed end lemma canonical_basis_length_ell2[code_unfold, simp]: "length (canonical_basis ::'a::enum ell2 list) = CARD('a)" unfolding canonical_basis_ell2_def apply simp using card_UNIV_length_enum by metis lemma ket_canonical_basis: "ket x = canonical_basis ! enum_idx x" proof- have "x = (enum_class.enum::'a list) ! enum_idx x" using enum_idx_correct[where i = x] by simp hence p1: "ket x = ket ((enum_class.enum::'a list) ! enum_idx x)" by simp have "enum_idx x < length (enum_class.enum::'a list)" using enum_idx_bound[where x = x]. hence "(map ket (enum_class.enum::'a list)) ! enum_idx x = ket ((enum_class.enum::'a list) ! enum_idx x)" by auto thus ?thesis unfolding canonical_basis_ell2_def using p1 by auto qed lemma clinear_equal_ket: fixes f g :: \'a::finite ell2 \ _\ assumes \clinear f\ assumes \clinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule complex_vector.linear_eq_on_span[where f=f and g=g and B=\range ket\]) using assms by auto lemma equal_ket: fixes A B :: \('a ell2, 'b::complex_normed_vector) cblinfun\ assumes \\ x. cblinfun_apply A (ket x) = cblinfun_apply B (ket x)\ shows \A = B\ apply (rule cblinfun_eq_gen_eqI[where G=\range ket\]) using assms by auto lemma antilinear_equal_ket: fixes f g :: \'a::finite ell2 \ _\ assumes \antilinear f\ assumes \antilinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ proof - have [simp]: \clinear (f \ from_conjugate_space)\ apply (rule antilinear_o_antilinear) using assms by (simp_all add: antilinear_from_conjugate_space) have [simp]: \clinear (g \ from_conjugate_space)\ apply (rule antilinear_o_antilinear) using assms by (simp_all add: antilinear_from_conjugate_space) have [simp]: \cspan (to_conjugate_space ` (range ket :: 'a ell2 set)) = UNIV\ by simp have "f o from_conjugate_space = g o from_conjugate_space" apply (rule ext) apply (rule complex_vector.linear_eq_on_span[where f="f o from_conjugate_space" and g="g o from_conjugate_space" and B=\to_conjugate_space ` range ket\]) apply (simp, simp) using assms(3) by (auto simp: to_conjugate_space_inverse) then show "f = g" by (smt (verit) UNIV_I from_conjugate_space_inverse surj_def surj_fun_eq to_conjugate_space_inject) qed lemma cinner_ket_adjointI: fixes F::"'a ell2 \\<^sub>C\<^sub>L _" and G::"'b ell2 \\<^sub>C\<^sub>L_" assumes "\ i j. \F *\<^sub>V ket i, ket j\ = \ket i, G *\<^sub>V ket j\" shows "F = G*" proof - from assms have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ if \x \ range ket\ and \y \ range ket\ for x y using that by auto then have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ if \x \ range ket\ for x y apply (rule bounded_clinear_eq_on[where G=\range ket\ and t=y, rotated 2]) using that by (auto intro!: bounded_linear_intros) then have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ for x y apply (rule bounded_antilinear_eq_on[where G=\range ket\ and t=x, rotated 2]) by (auto intro!: bounded_linear_intros) then show ?thesis by (rule adjoint_eqI) qed lemma ket_nonzero[simp]: "ket i \ 0" using norm_ket[of i] by force lemma cindependent_ket: "cindependent (range (ket::'a\_))" proof- define S where "S = range (ket::'a\_)" have "is_ortho_set S" unfolding S_def is_ortho_set_def by auto moreover have "0 \ S" unfolding S_def using ket_nonzero by (simp add: image_iff) ultimately show ?thesis using is_ortho_set_cindependent[where A = S] unfolding S_def by blast qed lemma cdim_UNIV_ell2[simp]: \cdim (UNIV::'a::finite ell2 set) = CARD('a)\ apply (subst cspan_range_ket_finite[symmetric]) by (metis card_image cindependent_ket complex_vector.dim_span_eq_card_independent inj_ket) lemma is_ortho_set_ket[simp]: \is_ortho_set (range ket)\ using is_ortho_set_def by fastforce subsection \Butterflies\ lemma cspan_butterfly_ket: \cspan {butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = UNIV\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cspan_butterfly_UNIV) by auto qed lemma cindependent_butterfly_ket: \cindependent {butterfly (ket i) (ket j)| (i::'b) (j::'a). True}\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b) (j::'a). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cindependent_butterfly) by auto qed lemma clinear_eq_butterfly_ketI: fixes F G :: \('a::finite ell2 \\<^sub>C\<^sub>L 'b::finite ell2) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes "\i j. F (butterfly (ket i) (ket j)) = G (butterfly (ket i) (ket j))" shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_ket) using assms by auto lemma sum_butterfly_ket[simp]: \(\(i::'a::finite)\UNIV. butterfly (ket i) (ket i)) = id_cblinfun\ apply (rule equal_ket) apply (subst complex_vector.linear_sum[where f=\\y. y *\<^sub>V ket _\]) apply (auto simp add: scaleC_cblinfun.rep_eq cblinfun.add_left clinearI butterfly_def cblinfun_compose_image cinner_ket) apply (subst sum.mono_neutral_cong_right[where S=\{_}\]) by auto subsection \One-dimensional spaces\ instantiation ell2 :: ("{enum,CARD_1}") one_dim begin text \Note: enum is not needed logically, but without it this instantiation clashes with \instantiation ell2 :: (enum) onb_enum\\ instance proof show "canonical_basis = [1::'a ell2]" unfolding canonical_basis_ell2_def apply transfer by (simp add: enum_CARD_1[of undefined]) show "a *\<^sub>C 1 * b *\<^sub>C 1 = (a * b) *\<^sub>C (1::'a ell2)" for a b apply (transfer fixing: a b) by simp show "x / y = x * inverse y" for x y :: "'a ell2" by (simp add: divide_inverse) show "inverse (c *\<^sub>C 1) = inverse c *\<^sub>C (1::'a ell2)" for c :: complex apply transfer by auto qed end subsection \Classical operators\ text \We call an operator mapping \<^term>\ket x\ to \<^term>\ket (\ x)\ or \<^term>\0\ "classical". (The meaning is inspired by the fact that in quantum mechanics, such operators usually correspond to operations with classical interpretation (such as Pauli-X, CNOT, measurement in the computational basis, etc.))\ definition classical_operator :: "('a\'b option) \ 'a ell2 \\<^sub>C\<^sub>L'b ell2" where "classical_operator \ = (let f = (\t. (case \ (inv (ket::'a\_) t) of None \ (0::'b ell2) | Some i \ ket i)) in cblinfun_extension (range (ket::'a\_)) f)" definition "classical_operator_exists \ \ cblinfun_extension_exists (range ket) (\t. case \ (inv ket t) of None \ 0 | Some i \ ket i)" lemma classical_operator_existsI: assumes "\x. B *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" shows "classical_operator_exists \" unfolding classical_operator_exists_def apply (rule cblinfun_extension_existsI[of _ B]) using assms by (auto simp: inv_f_f[OF inj_ket]) lemma classical_operator_exists_inj: assumes "inj_map \" shows "classical_operator_exists \" proof - define f where \f t = (case \ (inv ket t) of None \ 0 | Some x \ ket x)\ for t define g where \g = cconstruct (range ket) f\ have g_f: \g (ket x) = f (ket x)\ for x unfolding g_def apply (rule complex_vector.construct_basis) using cindependent_ket by auto have \clinear g\ unfolding g_def apply (rule complex_vector.linear_construct) using cindependent_ket by blast then have \g (x + y) = g x + g y\ if \x \ cspan (range ket)\ and \y \ cspan (range ket)\ for x y using clinear_iff by blast moreover from \clinear g\ have \g (c *\<^sub>C x) = c *\<^sub>C g x\ if \x \ cspan (range ket)\ for x c by (simp add: complex_vector.linear_scale) moreover have \norm (g x) \ norm x\ if \x \ cspan (range ket)\ for x proof - from that obtain t r where x_sum: \x = (\a\t. r a *\<^sub>C a)\ and \finite t\ and \t \ range ket\ unfolding complex_vector.span_explicit by auto then obtain T where tT: \t = ket ` T\ and [simp]: \finite T\ by (meson finite_subset_image) define R where \R i = r (ket i)\ for i have x_sum: \x = (\i\T. R i *\<^sub>C ket i)\ unfolding R_def tT x_sum apply (rule sum.reindex_cong) by (auto simp add: inj_on_def) define T' \' \T \R where \T' = {i\T. \ i \ None}\ and \\' = the o \\ and \\T = \' ` T'\ and \\R i = R (inv_into T' \' i)\ for i have \inj_on \' T'\ by (smt (z3) T'_def \'_def assms comp_apply inj_map_def inj_on_def mem_Collect_eq option.expand) have [simp]: \finite \T\ by (simp add: T'_def \T_def) have \g x = (\i\T. R i *\<^sub>C g (ket i))\ by (smt (verit, ccfv_threshold) \clinear g\ complex_vector.linear_scale complex_vector.linear_sum sum.cong x_sum) also have \\ = (\i\T. R i *\<^sub>C f (ket i))\ using g_f by presburger also have \\ = (\i\T. R i *\<^sub>C (case \ i of None \ 0 | Some x \ ket x))\ unfolding f_def by auto also have \\ = (\i\T'. R i *\<^sub>C ket (\' i))\ apply (rule sum.mono_neutral_cong_right) unfolding T'_def \'_def by auto also have \\ = (\i\\' ` T'. R (inv_into T' \' i) *\<^sub>C ket i)\ apply (subst sum.reindex) using \inj_on \' T'\ apply assumption apply (rule sum.cong) using \inj_on \' T'\ by auto finally have gx_sum: \g x = (\i\\T. \R i *\<^sub>C ket i)\ using \R_def \T_def by auto have \(norm (g x))\<^sup>2 = (\a\\T. (cmod (\R a))\<^sup>2)\ unfolding gx_sum apply (subst pythagorean_theorem_sum) by auto also have \\ = (\i\T'. (cmod (R i))\<^sup>2)\ unfolding \R_def \T_def apply (subst sum.reindex) using \inj_on \' T'\ apply assumption apply (rule sum.cong) using \inj_on \' T'\ by auto also have \\ \ (\a\T. (cmod (R a))\<^sup>2)\ apply (rule sum_mono2) using T'_def by auto also have \\ = (norm x)\<^sup>2\ unfolding x_sum apply (subst pythagorean_theorem_sum) using \finite T\ by auto finally show \norm (g x) \ norm x\ by auto qed ultimately have \cblinfun_extension_exists (cspan (range ket)) g\ apply (rule_tac cblinfun_extension_exists_bounded_dense[where B=1]) by auto then have \cblinfun_extension_exists (range ket) f\ by (metis (mono_tags, opaque_lifting) g_f cblinfun_extension_apply cblinfun_extension_existsI complex_vector.span_base rangeE) then show \classical_operator_exists \\ unfolding classical_operator_exists_def f_def by simp qed lemma classical_operator_exists_finite[simp]: "classical_operator_exists (\ :: _::finite \ _)" unfolding classical_operator_exists_def apply (rule cblinfun_extension_exists_finite_dim) using cindependent_ket apply blast using finite_class.finite_UNIV finite_imageI closed_cspan_range_ket closure_finite_cspan by blast lemma classical_operator_ket: assumes "classical_operator_exists \" shows "(classical_operator \) *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" unfolding classical_operator_def using f_inv_into_f ket_injective rangeI by (metis assms cblinfun_extension_apply classical_operator_exists_def) lemma classical_operator_ket_finite: "(classical_operator \) *\<^sub>V (ket (x::'a::finite)) = (case \ x of Some i \ ket i | None \ 0)" by (rule classical_operator_ket, simp) lemma classical_operator_adjoint[simp]: fixes \ :: "'a \ 'b option" assumes a1: "inj_map \" shows "(classical_operator \)* = classical_operator (inv_map \)" proof- define F where "F = classical_operator (inv_map \)" define G where "G = classical_operator \" have "\F *\<^sub>V ket i, ket j\ = \ket i, G *\<^sub>V ket j\" for i j proof- have w1: "(classical_operator (inv_map \)) *\<^sub>V (ket i) = (case inv_map \ i of Some k \ ket k | None \ 0)" by (simp add: classical_operator_ket classical_operator_exists_inj) have w2: "(classical_operator \) *\<^sub>V (ket j) = (case \ j of Some k \ ket k | None \ 0)" by (simp add: assms classical_operator_ket classical_operator_exists_inj) have "\F *\<^sub>V ket i, ket j\ = \classical_operator (inv_map \) *\<^sub>V ket i, ket j\" unfolding F_def by blast also have "\ = \(case inv_map \ i of Some k \ ket k | None \ 0), ket j\" using w1 by simp also have "\ = \ket i, (case \ j of Some k \ ket k | None \ 0)\" proof(induction "inv_map \ i") case None hence pi1: "None = inv_map \ i". show ?case proof (induction "\ j") case None thus ?case using pi1 by auto next case (Some c) have "c \ i" proof(rule classical) assume "\(c \ i)" hence "c = i" by blast hence "inv_map \ c = inv_map \ i" by simp hence "inv_map \ c = None" by (simp add: pi1) moreover have "inv_map \ c = Some j" using Some.hyps unfolding inv_map_def apply auto by (metis a1 f_inv_into_f inj_map_def option.distinct(1) rangeI) ultimately show ?thesis by simp qed thus ?thesis by (metis None.hyps Some.hyps cinner_zero_left orthogonal_ket option.simps(4) option.simps(5)) qed next case (Some d) hence s1: "Some d = inv_map \ i". show "\case inv_map \ i of None \ 0 | Some a \ ket a, ket j\ = \ket i, case \ j of None \ 0 | Some a \ ket a\" proof(induction "\ j") case None have "d \ j" proof(rule classical) assume "\(d \ j)" hence "d = j" by blast hence "\ d = \ j" by simp hence "\ d = None" by (simp add: None.hyps) moreover have "\ d = Some i" using Some.hyps unfolding inv_map_def apply auto by (metis f_inv_into_f option.distinct(1) option.inject) ultimately show ?thesis by simp qed thus ?case by (metis None.hyps Some.hyps cinner_zero_right orthogonal_ket option.case_eq_if option.simps(5)) next case (Some c) hence s2: "\ j = Some c" by simp have "\ket d, ket j\ = \ket i, ket c\" proof(cases "\ j = Some i") case True hence ij: "Some j = inv_map \ i" unfolding inv_map_def apply auto apply (metis a1 f_inv_into_f inj_map_def option.discI range_eqI) by (metis range_eqI) have "i = c" using True s2 by auto moreover have "j = d" by (metis option.inject s1 ij) ultimately show ?thesis by (simp add: cinner_ket_same) next case False moreover have "\ d = Some i" using s1 unfolding inv_map_def by (metis f_inv_into_f option.distinct(1) option.inject) ultimately have "j \ d" by auto moreover have "i \ c" using False s2 by auto ultimately show ?thesis by (metis orthogonal_ket) qed hence "\case Some d of None \ 0 | Some a \ ket a, ket j\ = \ket i, case Some c of None \ 0 | Some a \ ket a\" by simp thus "\case inv_map \ i of None \ 0 | Some a \ ket a, ket j\ = \ket i, case \ j of None \ 0 | Some a \ ket a\" by (simp add: Some.hyps s1) qed qed also have "\ = \ket i, classical_operator \ *\<^sub>V ket j\" by (simp add: w2) also have "\ = \ket i, G *\<^sub>V ket j\" unfolding G_def by blast finally show ?thesis . qed hence "G* = F" using cinner_ket_adjointI by auto thus ?thesis unfolding G_def F_def . qed lemma fixes \::"'b \ 'c option" and \::"'a \ 'b option" assumes "classical_operator_exists \" assumes "classical_operator_exists \" shows classical_operator_exists_comp[simp]: "classical_operator_exists (\ \\<^sub>m \)" and classical_operator_mult[simp]: "classical_operator \ o\<^sub>C\<^sub>L classical_operator \ = classical_operator (\ \\<^sub>m \)" proof - define C\ C\ C\\ where "C\ = classical_operator \" and "C\ = classical_operator \" and "C\\ = classical_operator (\ \\<^sub>m \)" have C\x: "C\ *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" for x unfolding C\_def using \classical_operator_exists \\ by (rule classical_operator_ket) have C\x: "C\ *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" for x unfolding C\_def using \classical_operator_exists \\ by (rule classical_operator_ket) have C\\x': "(C\ o\<^sub>C\<^sub>L C\) *\<^sub>V (ket x) = (case (\ \\<^sub>m \) x of Some i \ ket i | None \ 0)" for x apply (simp add: scaleC_cblinfun.rep_eq C\x) apply (cases "\ x") by (auto simp: C\x) thus \classical_operator_exists (\ \\<^sub>m \)\ by (rule classical_operator_existsI) hence "C\\ *\<^sub>V (ket x) = (case (\ \\<^sub>m \) x of Some i \ ket i | None \ 0)" for x unfolding C\\_def by (rule classical_operator_ket) with C\\x' have "(C\ o\<^sub>C\<^sub>L C\) *\<^sub>V (ket x) = C\\ *\<^sub>V (ket x)" for x by simp thus "C\ o\<^sub>C\<^sub>L C\ = C\\" by (simp add: equal_ket) qed lemma classical_operator_Some[simp]: "classical_operator (Some::'a\_) = id_cblinfun" proof- have "(classical_operator Some) *\<^sub>V (ket i) = id_cblinfun *\<^sub>V (ket i)" for i::'a apply (subst classical_operator_ket) apply (rule classical_operator_exists_inj) by auto thus ?thesis using equal_ket[where A = "classical_operator (Some::'a \ _ option)" and B = "id_cblinfun::'a ell2 \\<^sub>C\<^sub>L _"] by blast qed lemma isometry_classical_operator[simp]: fixes \::"'a \ 'b" assumes a1: "inj \" shows "isometry (classical_operator (Some o \))" proof - have b0: "inj_map (Some \ \)" by (simp add: a1) have b0': "inj_map (inv_map (Some \ \))" by simp have b1: "inv_map (Some \ \) \\<^sub>m (Some \ \) = Some" apply (rule ext) unfolding inv_map_def o_def using assms unfolding inj_def inv_def by auto have b3: "classical_operator (inv_map (Some \ \)) o\<^sub>C\<^sub>L classical_operator (Some \ \) = classical_operator (inv_map (Some \ \) \\<^sub>m (Some \ \))" by (metis b0 b0' b1 classical_operator_Some classical_operator_exists_inj classical_operator_mult) show ?thesis unfolding isometry_def apply (subst classical_operator_adjoint) using b0 by (auto simp add: b1 b3) qed lemma unitary_classical_operator[simp]: fixes \::"'a \ 'b" assumes a1: "bij \" shows "unitary (classical_operator (Some o \))" proof (unfold unitary_def, rule conjI) have "inj \" using a1 bij_betw_imp_inj_on by auto hence "isometry (classical_operator (Some o \))" by simp hence "classical_operator (Some \ \)* o\<^sub>C\<^sub>L classical_operator (Some \ \) = id_cblinfun" unfolding isometry_def by simp thus \classical_operator (Some \ \)* o\<^sub>C\<^sub>L classical_operator (Some \ \) = id_cblinfun\ by simp next have "inj \" by (simp add: assms bij_is_inj) have comp: "Some \ \ \\<^sub>m inv_map (Some \ \) = Some" apply (rule ext) unfolding inv_map_def o_def map_comp_def unfolding inv_def apply auto apply (metis \inj \\ inv_def inv_f_f) using bij_def image_iff range_eqI by (metis a1) have "classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (Some \ \)* = classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (inv_map (Some \ \))" by (simp add: \inj \\) also have "\ = classical_operator ((Some \ \) \\<^sub>m (inv_map (Some \ \)))" by (simp add: \inj \\ classical_operator_exists_inj) also have "\ = classical_operator (Some::'b\_)" using comp by simp also have "\ = (id_cblinfun:: 'b ell2 \\<^sub>C\<^sub>L _)" by simp finally show "classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (Some \ \)* = id_cblinfun". qed - +unbundle no_lattice_syntax unbundle no_cblinfun_notation end diff --git a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy --- a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy +++ b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy @@ -1,2820 +1,2824 @@ section \\Complex_Vector_Spaces\ -- Complex Vector Spaces\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Vector_Spaces imports "HOL-Analysis.Elementary_Topology" "HOL-Analysis.Operator_Norm" "HOL-Analysis.Elementary_Normed_Spaces" "HOL-Library.Set_Algebras" "HOL-Analysis.Starlike" "HOL-Types_To_Sets.Types_To_Sets" "Complex_Bounded_Operators.Extra_Vector_Spaces" "Complex_Bounded_Operators.Extra_Ordered_Fields" "Complex_Bounded_Operators.Extra_Lattice" "Complex_Bounded_Operators.Extra_General" Complex_Vector_Spaces0 begin bundle notation_norm begin notation norm ("\_\") end +unbundle lattice_syntax + subsection \Misc\ lemma (in scaleC) scaleC_real: assumes "r\\" shows "r *\<^sub>C x = Re r *\<^sub>R x" unfolding scaleR_scaleC using assms by simp lemma of_complex_of_real_eq [simp]: "of_complex (of_real n) = of_real n" unfolding of_complex_def of_real_def unfolding scaleR_scaleC by simp lemma Complexs_of_real [simp]: "of_real r \ \" unfolding Complexs_def of_real_def of_complex_def apply (subst scaleR_scaleC) by simp lemma Reals_in_Complexs: "\ \ \" unfolding Reals_def by auto lemma (in clinear) "linear f" apply standard by (simp_all add: add scaleC scaleR_scaleC) lemma (in bounded_clinear) bounded_linear: "bounded_linear f" by (simp add: add bounded bounded_linear.intro bounded_linear_axioms.intro linearI scaleC scaleR_scaleC) lemma clinear_times: "clinear (\x. c * x)" for c :: "'a::complex_algebra" by (auto simp: clinearI distrib_left) lemma (in clinear) linear: shows \linear f\ by (simp add: add linearI scaleC scaleR_scaleC) lemma bounded_clinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_clinear f" using assms by (auto intro!: exI bounded_clinear.intro clinearI simp: bounded_clinear_axioms_def) lemma bounded_clinear_id[simp]: \bounded_clinear id\ by (simp add: id_def) (* The following would be a natural inclusion of locales, but unfortunately it leads to name conflicts upon interpretation of bounded_cbilinear *) (* sublocale bounded_cbilinear \ bounded_bilinear by (rule bounded_bilinear) *) definition cbilinear :: \('a::complex_vector \ 'b::complex_vector \ 'c::complex_vector) \ bool\ where \cbilinear = (\ f. (\ y. clinear (\ x. f x y)) \ (\ x. clinear (\ y. f x y)) )\ lemma cbilinear_add_left: assumes \cbilinear f\ shows \f (a + b) c = f a c + f b c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_add_right: assumes \cbilinear f\ shows \f a (b + c) = f a b + f a c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_times: fixes g' :: \'a::complex_vector \ complex\ and g :: \'b::complex_vector \ complex\ assumes \\ x y. h x y = (g' x)*(g y)\ and \clinear g\ and \clinear g'\ shows \cbilinear h\ proof - have w1: "h (b1 + b2) y = h b1 y + h b2 y" for b1 :: 'a and b2 :: 'a and y proof- have \h (b1 + b2) y = g' (b1 + b2) * g y\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = (g' b1 + g' b2) * g y\ using \clinear g'\ unfolding clinear_def by (simp add: assms(3) complex_vector.linear_add) also have \\ = g' b1 * g y + g' b2 * g y\ by (simp add: ring_class.ring_distribs(2)) also have \\ = h b1 y + h b2 y\ using assms(1) by auto finally show ?thesis by blast qed have w2: "h (r *\<^sub>C b) y = r *\<^sub>C h b y" for r :: complex and b :: 'a and y proof- have \h (r *\<^sub>C b) y = g' (r *\<^sub>C b) * g y\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' b * g y)\ by (simp add: assms(3) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h b y)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "clinear (\x. h x y)" for y :: 'b unfolding clinear_def by (meson clinearI clinear_def w1 w2) hence t2: "\y. clinear (\x. h x y)" by simp have v1: "h x (b1 + b2) = h x b1 + h x b2" for b1 :: 'b and b2 :: 'b and x proof- have \h x (b1 + b2) = g' x * g (b1 + b2)\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = g' x * (g b1 + g b2)\ using \clinear g'\ unfolding clinear_def by (simp add: assms(2) complex_vector.linear_add) also have \\ = g' x * g b1 + g' x * g b2\ by (simp add: ring_class.ring_distribs(1)) also have \\ = h x b1 + h x b2\ using assms(1) by auto finally show ?thesis by blast qed have v2: "h x (r *\<^sub>C b) = r *\<^sub>C h x b" for r :: complex and b :: 'b and x proof- have \h x (r *\<^sub>C b) = g' x * g (r *\<^sub>C b)\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' x * g b)\ by (simp add: assms(2) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h x b)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) (h x)" for x :: 'a using v1 v2 by (meson clinearI clinear_def) hence t1: "\x. clinear (h x)" unfolding clinear_def by simp show ?thesis unfolding cbilinear_def by (simp add: t1 t2) qed lemma csubspace_is_subspace: "csubspace A \ subspace A" apply (rule subspaceI) by (auto simp: complex_vector.subspace_def scaleR_scaleC) lemma span_subset_cspan: "span A \ cspan A" unfolding span_def complex_vector.span_def by (simp add: csubspace_is_subspace hull_antimono) lemma cindependent_implies_independent: assumes "cindependent (S::'a::complex_vector set)" shows "independent S" using assms unfolding dependent_def complex_vector.dependent_def using span_subset_cspan by blast lemma cspan_singleton: "cspan {x} = {\ *\<^sub>C x| \. True}" proof - have \cspan {x} = {y. y\cspan {x}}\ by auto also have \\ = {\ *\<^sub>C x| \. True}\ apply (subst complex_vector.span_breakdown_eq) by auto finally show ?thesis by - qed lemma cspan_as_span: "cspan (B::'a::complex_vector set) = span (B \ scaleC \ ` B)" proof auto let ?cspan = complex_vector.span let ?rspan = real_vector.span fix \ assume cspan: "\ \ ?cspan B" have "\B' r. finite B' \ B' \ B \ \ = (\b\B'. r b *\<^sub>C b)" using complex_vector.span_explicit[of B] cspan by auto then obtain B' r where "finite B'" and "B' \ B" and \_explicit: "\ = (\b\B'. r b *\<^sub>C b)" by atomize_elim define R where "R = B \ scaleC \ ` B" have x2: "(case x of (b, i) \ if i then Im (r b) *\<^sub>R \ *\<^sub>C b else Re (r b) *\<^sub>R b) \ span (B \ (*\<^sub>C) \ ` B)" if "x \ B' \ (UNIV::bool set)" for x :: "'a \ bool" using that \B' \ B\ by (auto simp add: real_vector.span_base real_vector.span_scale subset_iff) have x1: "\ = (\x\B'. \i\UNIV. if i then Im (r x) *\<^sub>R \ *\<^sub>C x else Re (r x) *\<^sub>R x)" if "\b. r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" using that by (simp add: UNIV_bool \_explicit) moreover have "r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" for b using complex_eq scaleC_add_left scaleC_scaleC scaleR_scaleC by (metis (no_types, lifting) complex_of_real_i i_complex_of_real) ultimately have "\ = (\(b,i)\(B'\UNIV). if i then Im (r b) *\<^sub>R (\ *\<^sub>C b) else Re (r b) *\<^sub>R b)" by (simp add: sum.cartesian_product) also have "\ \ ?rspan R" unfolding R_def using x2 by (rule real_vector.span_sum) finally show "\ \ ?rspan R" by - next let ?cspan = complex_vector.span let ?rspan = real_vector.span define R where "R = B \ scaleC \ ` B" fix \ assume rspan: "\ \ ?rspan R" have "subspace {a. a \ cspan B}" by (rule real_vector.subspaceI, auto simp add: complex_vector.span_zero complex_vector.span_add_eq2 complex_vector.span_scale scaleR_scaleC) moreover have "x \ cspan B" if "x \ R" for x :: 'a using that R_def complex_vector.span_base complex_vector.span_scale by fastforce ultimately show "\ \ ?cspan B" using real_vector.span_induct rspan by blast qed lemma isomorphic_equal_cdim: assumes lin_f: \clinear f\ assumes inj_f: \inj_on f (cspan S)\ assumes im_S: \f ` S = T\ shows \cdim S = cdim T\ proof - obtain SB where SB_span: "cspan SB = cspan S" and indep_SB: \cindependent SB\ by (metis complex_vector.basis_exists complex_vector.span_mono complex_vector.span_span subset_antisym) with lin_f inj_f have indep_fSB: \cindependent (f ` SB)\ apply (rule_tac complex_vector.linear_independent_injective_image) by auto from lin_f have \cspan (f ` SB) = f ` cspan SB\ by (meson complex_vector.linear_span_image) also from SB_span lin_f have \\ = cspan T\ by (metis complex_vector.linear_span_image im_S) finally have \cdim T = card (f ` SB)\ using indep_fSB complex_vector.dim_eq_card by blast also have \\ = card SB\ apply (rule card_image) using inj_f by (metis SB_span complex_vector.linear_inj_on_span_iff_independent_image indep_fSB lin_f) also have \\ = cdim S\ using indep_SB SB_span by (metis complex_vector.dim_eq_card) finally show ?thesis by simp qed lemma cindependent_inter_scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 1" shows "B \ (*\<^sub>C) c ` B = {}" proof (rule classical, cases \c = 0\) case True then show ?thesis using a1 by (auto simp add: complex_vector.dependent_zero) next case False assume "\(B \ (*\<^sub>C) c ` B = {})" hence "B \ (*\<^sub>C) c ` B \ {}" by blast then obtain x where u1: "x \ B \ (*\<^sub>C) c ` B" by blast then obtain b where u2: "x = b" and u3: "b\B" by blast then obtain b' where u2': "x = c *\<^sub>C b'" and u3': "b'\B" using u1 by blast have g1: "b = c *\<^sub>C b'" using u2 and u2' by simp hence "b \ complex_vector.span {b'}" using False by (simp add: complex_vector.span_base complex_vector.span_scale) hence "b = b'" by (metis u3' a1 complex_vector.dependent_def complex_vector.span_base complex_vector.span_scale insertE insert_Diff u2 u2' u3) hence "b' = c *\<^sub>C b'" using g1 by blast thus ?thesis by (metis a1 a3 complex_vector.dependent_zero complex_vector.scale_right_imp_eq mult_cancel_right2 scaleC_scaleC u3') qed lemma real_independent_from_complex_independent: assumes "cindependent (B::'a::complex_vector set)" defines "B' == ((*\<^sub>C) \ ` B)" shows "independent (B \ B')" proof (rule notI) assume \dependent (B \ B')\ then obtain T f0 x where [simp]: \finite T\ and \T \ B \ B'\ and f0_sum: \(\v\T. f0 v *\<^sub>R v) = 0\ and x: \x \ T\ and f0_x: \f0 x \ 0\ by (auto simp: real_vector.dependent_explicit) define f T1 T2 T' f' x' where \f v = (if v \ T then f0 v else 0)\ and \T1 = T \ B\ and \T2 = scaleC (-\) ` (T \ B')\ and \T' = T1 \ T2\ and \f' v = f v + \ * f (\ *\<^sub>C v)\ and \x' = (if x \ T1 then x else -\ *\<^sub>C x)\ for v have \B \ B' = {}\ by (simp add: assms cindependent_inter_scaleC_cindependent) have \T' \ B\ by (auto simp: T'_def T1_def T2_def B'_def) have [simp]: \finite T'\ \finite T1\ \finite T2\ by (auto simp add: T'_def T1_def T2_def) have f_sum: \(\v\T. f v *\<^sub>R v) = 0\ unfolding f_def using f0_sum by auto have f_x: \f x \ 0\ using f0_x x by (auto simp: f_def) have f'_sum: \(\v\T'. f' v *\<^sub>C v) = 0\ proof - have \(\v\T'. f' v *\<^sub>C v) = (\v\T'. complex_of_real (f v) *\<^sub>C v) + (\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ by (auto simp: f'_def sum.distrib scaleC_add_left) also have \(\v\T'. complex_of_real (f v) *\<^sub>C v) = (\v\T1. f v *\<^sub>R v)\ (is \_ = ?left\) apply (auto simp: T'_def scaleR_scaleC intro!: sum.mono_neutral_cong_right) using T'_def T1_def \T' \ B\ f_def by auto also have \(\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v) = (\v\T2. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ (is \_ = ?right\) apply (auto simp: T'_def intro!: sum.mono_neutral_cong_right) by (smt (z3) B'_def IntE IntI T1_def T2_def \f \ \v. if v \ T then f0 v else 0\ add.inverse_inverse complex_vector.vector_space_axioms i_squared imageI mult_minus_left vector_space.vector_space_assms(3) vector_space.vector_space_assms(4)) also have \?right = (\v\T\B'. f v *\<^sub>R v)\ (is \_ = ?right\) apply (rule sum.reindex_cong[symmetric, where l=\scaleC \\]) apply (auto simp: T2_def image_image scaleR_scaleC) using inj_on_def by fastforce also have \?left + ?right = (\v\T. f v *\<^sub>R v)\ apply (subst sum.union_disjoint[symmetric]) using \B \ B' = {}\ \T \ B \ B'\ apply (auto simp: T1_def) by (metis Int_Un_distrib Un_Int_eq(4) sup.absorb_iff1) also have \\ = 0\ by (rule f_sum) finally show ?thesis by - qed have x': \x' \ T'\ using \T \ B \ B'\ x by (auto simp: x'_def T'_def T1_def T2_def) have f'_x': \f' x' \ 0\ using Complex_eq Complex_eq_0 f'_def f_x x'_def by auto from \finite T'\ \T' \ B\ f'_sum x' f'_x' have \cdependent B\ using complex_vector.independent_explicit_module by blast with assms show False by auto qed lemma crepresentation_from_representation: assumes a1: "cindependent B" and a2: "b \ B" and a3: "finite B" shows "crepresentation B \ b = (representation (B \ (*\<^sub>C) \ ` B) \ b) + \ *\<^sub>C (representation (B \ (*\<^sub>C) \ ` B) \ (\ *\<^sub>C b))" proof (cases "\ \ cspan B") define B' where "B' = B \ (*\<^sub>C) \ ` B" case True define r where "r v = real_vector.representation B' \ v" for v define r' where "r' v = real_vector.representation B' \ (\ *\<^sub>C v)" for v define f where "f v = r v + \ *\<^sub>C r' v" for v define g where "g v = crepresentation B \ v" for v have "(\v | g v \ 0. g v *\<^sub>C v) = \" unfolding g_def using Collect_cong Collect_mono_iff DiffD1 DiffD2 True a1 complex_vector.finite_representation complex_vector.sum_nonzero_representation_eq sum.mono_neutral_cong_left by fastforce moreover have "finite {v. g v \ 0}" unfolding g_def by (simp add: complex_vector.finite_representation) moreover have "v \ B" if "g v \ 0" for v using that unfolding g_def by (simp add: complex_vector.representation_ne_zero) ultimately have rep1: "(\v\B. g v *\<^sub>C v) = \" unfolding g_def using a3 True a1 complex_vector.sum_representation_eq by blast have l0': "inj ((*\<^sub>C) \::'a \'a)" unfolding inj_def by simp have l0: "inj ((*\<^sub>C) (- \)::'a \'a)" unfolding inj_def by simp have l1: "(*\<^sub>C) (- \) ` B \ B = {}" using cindependent_inter_scaleC_cindependent[where B=B and c = "- \"] by (metis Int_commute a1 add.inverse_inverse complex_i_not_one i_squared mult_cancel_left1 neg_equal_0_iff_equal) have l2: "B \ (*\<^sub>C) \ ` B = {}" by (simp add: a1 cindependent_inter_scaleC_cindependent) have rr1: "r (\ *\<^sub>C v) = r' v" for v unfolding r_def r'_def by simp have k1: "independent B'" unfolding B'_def using a1 real_independent_from_complex_independent by simp have "\ \ span B'" using B'_def True cspan_as_span by blast have "v \ B'" if "r v \ 0" for v unfolding r_def using r_def real_vector.representation_ne_zero that by auto have "finite B'" unfolding B'_def using a3 by simp have "(\v\B'. r v *\<^sub>R v) = \" unfolding r_def using True Real_Vector_Spaces.real_vector.sum_representation_eq[where B = B' and basis = B' and v = \] by (smt Real_Vector_Spaces.dependent_raw_def \\ \ Real_Vector_Spaces.span B'\ \finite B'\ equalityD2 k1) have d1: "(\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v)) = (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using l0' by (metis (mono_tags, lifting) inj_eq inj_on_def sum.reindex_cong) have "(\v\B. (r v + \ * (r' v)) *\<^sub>C v) = (\v\B. r v *\<^sub>C v + (\ * r' v) *\<^sub>C v)" by (meson scaleC_left.add) also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. (\ * r' v) *\<^sub>C v)" using sum.distrib by fastforce also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. \ *\<^sub>C (r' v *\<^sub>C v))" by auto also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. \ *\<^sub>C (r (\ *\<^sub>C v) *\<^sub>R v))" unfolding r'_def r_def by (metis (mono_tags, lifting) scaleR_scaleC sum.cong) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v))" by (metis (no_types, lifting) complex_vector.scale_left_commute scaleR_scaleC) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using d1 by simp also have "\ = \" using l2 \(\v\B'. r v *\<^sub>R v) = \\ unfolding B'_def by (simp add: a3 sum.union_disjoint) finally have "(\v\B. f v *\<^sub>C v) = \" unfolding r'_def r_def f_def by simp hence "0 = (\v\B. f v *\<^sub>C v) - (\v\B. crepresentation B \ v *\<^sub>C v)" using rep1 unfolding g_def by simp also have "\ = (\v\B. f v *\<^sub>C v - crepresentation B \ v *\<^sub>C v)" by (simp add: sum_subtractf) also have "\ = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)" by (metis scaleC_left.diff) finally have "0 = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)". hence "(\v\B. (f v - crepresentation B \ v) *\<^sub>C v) = 0" by simp hence "f b - crepresentation B \ b = 0" using a1 a2 a3 complex_vector.independentD[where s = B and t = B and u = "\v. f v - crepresentation B \ v" and v = b] order_refl by smt hence "crepresentation B \ b = f b" by simp thus ?thesis unfolding f_def r_def r'_def B'_def by auto next define B' where "B' = B \ (*\<^sub>C) \ ` B" case False have b2: "\ \ real_vector.span B'" unfolding B'_def using False cspan_as_span by auto have "\ \ complex_vector.span B" using False by blast have "crepresentation B \ b = 0" unfolding complex_vector.representation_def by (simp add: False) moreover have "real_vector.representation B' \ b = 0" unfolding real_vector.representation_def by (simp add: b2) moreover have "real_vector.representation B' \ ((*\<^sub>C) \ b) = 0" unfolding real_vector.representation_def by (simp add: b2) ultimately show ?thesis unfolding B'_def by simp qed lemma CARD_1_vec_0[simp]: \(\ :: _ ::{complex_vector,CARD_1}) = 0\ by auto lemma scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 0" shows "cindependent ((*\<^sub>C) c ` B)" proof- have "u y = 0" if g1: "y\S" and g2: "(\x\S. u x *\<^sub>C x) = 0" and g3: "finite S" and g4: "S\(*\<^sub>C) c ` B" for u y S proof- define v where "v x = u (c *\<^sub>C x)" for x obtain S' where "S'\B" and S_S': "S = (*\<^sub>C) c ` S'" by (meson g4 subset_imageE) have "inj ((*\<^sub>C) c::'a\_)" unfolding inj_def using a3 by auto hence "finite S'" using S_S' finite_imageD g3 subset_inj_on by blast have "t \ (*\<^sub>C) (inverse c) ` S" if "t \ S'" for t proof- have "c *\<^sub>C t \ S" using \S = (*\<^sub>C) c ` S'\ that by blast hence "(inverse c) *\<^sub>C (c *\<^sub>C t) \ (*\<^sub>C) (inverse c) ` S" by blast moreover have "(inverse c) *\<^sub>C (c *\<^sub>C t) = t" by (simp add: a3) ultimately show ?thesis by simp qed moreover have "t \ S'" if "t \ (*\<^sub>C) (inverse c) ` S" for t proof- obtain t' where "t = (inverse c) *\<^sub>C t'" and "t' \ S" using \t \ (*\<^sub>C) (inverse c) ` S\ by auto have "c *\<^sub>C t = c *\<^sub>C ((inverse c) *\<^sub>C t')" using \t = (inverse c) *\<^sub>C t'\ by simp also have "\ = (c * (inverse c)) *\<^sub>C t'" by simp also have "\ = t'" by (simp add: a3) finally have "c *\<^sub>C t = t'". thus ?thesis using \t' \ S\ using \S = (*\<^sub>C) c ` S'\ a3 complex_vector.scale_left_imp_eq by blast qed ultimately have "S' = (*\<^sub>C) (inverse c) ` S" by blast hence "inverse c *\<^sub>C y \ S'" using that(1) by blast have t: "inj (((*\<^sub>C) c)::'a \ _)" using a3 complex_vector.injective_scale[where c = c] by blast have "0 = (\x\(*\<^sub>C) c ` S'. u x *\<^sub>C x)" using \S = (*\<^sub>C) c ` S'\ that(2) by auto also have "\ = (\x\S'. v x *\<^sub>C (c *\<^sub>C x))" unfolding v_def using t Groups_Big.comm_monoid_add_class.sum.reindex[where h = "((*\<^sub>C) c)" and A = S' and g = "\x. u x *\<^sub>C x"] subset_inj_on by auto also have "\ = c *\<^sub>C (\x\S'. v x *\<^sub>C x)" by (metis (mono_tags, lifting) complex_vector.scale_left_commute scaleC_right.sum sum.cong) finally have "0 = c *\<^sub>C (\x\S'. v x *\<^sub>C x)". hence "(\x\S'. v x *\<^sub>C x) = 0" using a3 by auto hence "v (inverse c *\<^sub>C y) = 0" using \inverse c *\<^sub>C y \ S'\ \finite S'\ \S' \ B\ a1 complex_vector.independentD by blast thus "u y = 0" unfolding v_def by (simp add: a3) qed thus ?thesis using complex_vector.dependent_explicit by (simp add: complex_vector.dependent_explicit ) qed subsection \Antilinear maps and friends\ locale antilinear = additive f for f :: "'a::complex_vector \ 'b::complex_vector" + assumes scaleC: "f (scaleC r x) = cnj r *\<^sub>C f x" sublocale antilinear \ linear proof (rule linearI) show "f (b1 + b2) = f b1 + f b2" for b1 :: 'a and b2 :: 'a by (simp add: add) show "f (r *\<^sub>R b) = r *\<^sub>R f b" for r :: real and b :: 'a unfolding scaleR_scaleC by (subst scaleC, simp) qed lemma antilinear_imp_scaleC: fixes D :: "complex \ 'a::complex_vector" assumes "antilinear D" obtains d where "D = (\x. cnj x *\<^sub>C d)" proof - interpret clinear "D o cnj" apply standard apply auto apply (simp add: additive.add assms antilinear.axioms(1)) using assms antilinear.scaleC by fastforce obtain d where "D o cnj = (\x. x *\<^sub>C d)" using clinear_axioms complex_vector.linear_imp_scale by blast then have \D = (\x. cnj x *\<^sub>C d)\ by (metis comp_apply complex_cnj_cnj) then show ?thesis by (rule that) qed corollary complex_antilinearD: fixes f :: "complex \ complex" assumes "antilinear f" obtains c where "f = (\x. c * cnj x)" by (rule antilinear_imp_scaleC [OF assms]) (force simp: scaleC_conv_of_complex) lemma antilinearI: assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>C x) = cnj c *\<^sub>C f x" shows "antilinear f" by standard (rule assms)+ lemma antilinear_o_antilinear: "antilinear f \ antilinear g \ clinear (g o f)" apply (rule clinearI) apply (simp add: additive.add antilinear_def) by (simp add: antilinear.scaleC) lemma clinear_o_antilinear: "antilinear f \ clinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) lemma antilinear_o_clinear: "clinear f \ antilinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) locale bounded_antilinear = antilinear f for f :: "'a::complex_normed_vector \ 'b::complex_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" lemma bounded_antilinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = cnj r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_antilinear f" using assms by (auto intro!: exI bounded_antilinear.intro antilinearI simp: bounded_antilinear_axioms_def) sublocale bounded_antilinear \ bounded_linear apply standard by (fact bounded) lemma (in bounded_antilinear) bounded_linear: "bounded_linear f" by (fact bounded_linear) lemma (in bounded_antilinear) antilinear: "antilinear f" by (fact antilinear_axioms) lemma bounded_antilinear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleC r x) = scaleC (cnj r) (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_antilinear f" by standard (blast intro: assms)+ lemma bounded_antilinear_0[simp]: \bounded_antilinear (\_. 0)\ by (rule bounded_antilinear_intro[where K=0], auto) lemma cnj_bounded_antilinear[simp]: "bounded_antilinear cnj" apply (rule bounded_antilinear_intro [where K = 1]) by auto lemma bounded_antilinear_o_bounded_antilinear: assumes "bounded_antilinear f" and "bounded_antilinear g" shows "bounded_clinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_antilinear g by fact fix b1 b2 b r show "f (g (b1 + b2)) = f (g b1) + f (g b2)" by (simp add: f.add g.add) show "f (g (r *\<^sub>C b)) = r *\<^sub>C f (g b)" by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_antilinear_o_bounded_clinear: assumes "bounded_antilinear f" and "bounded_clinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_clinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_clinear_o_bounded_antilinear: assumes "bounded_clinear f" and "bounded_antilinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_clinear f by fact interpret g: bounded_antilinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x using f.scaleC g.scaleC by fastforce have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bij_clinear_imp_inv_clinear: "clinear (inv f)" if a1: "clinear f" and a2: "bij f" proof fix b1 b2 r b show "inv f (b1 + b2) = inv f b1 + inv f b2" by (simp add: a1 a2 bij_is_inj bij_is_surj complex_vector.linear_add inv_f_eq surj_f_inv_f) show "inv f (r *\<^sub>C b) = r *\<^sub>C inv f b" using that by (smt bij_inv_eq_iff clinear_def complex_vector.linear_scale) qed locale bounded_sesquilinear = fixes prod :: "'a::complex_normed_vector \ 'b::complex_normed_vector \ 'c::complex_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleC_left: "prod (r *\<^sub>C a) b = (cnj r) *\<^sub>C (prod a b)" and scaleC_right: "prod a (r *\<^sub>C b) = r *\<^sub>C (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" sublocale bounded_sesquilinear \ bounded_bilinear apply standard by (auto simp: add_left add_right scaleC_left scaleC_right bounded scaleR_scaleC) lemma (in bounded_sesquilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by (fact bounded_bilinear_axioms) lemma (in bounded_sesquilinear) bounded_antilinear_left: "bounded_antilinear (\a. prod a b)" apply standard apply (auto simp add: scaleC_left add_left) by (metis ab_semigroup_mult_class.mult_ac(1) bounded) lemma (in bounded_sesquilinear) bounded_clinear_right: "bounded_clinear (\b. prod a b)" apply standard apply (auto simp add: scaleC_right add_right) by (metis ab_semigroup_mult_class.mult_ac(1) ordered_field_class.sign_simps(34) pos_bounded) lemma (in bounded_sesquilinear) comp1: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x. prod (g x))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (g (a + a')) b = prod (g a) b + prod (g a') b" by (simp add: add add_left) show "prod (g a) (b + b') = prod (g a) b + prod (g a) b'" by (simp add: add add_right) show "prod (g (r *\<^sub>C a)) b = cnj r *\<^sub>C prod (g a) b" by (simp add: scaleC scaleC_left) show "prod (g a) (r *\<^sub>C b) = r *\<^sub>C prod (g a) b" by (simp add: scaleC_right) interpret bounded_bilinear \(\x. prod (g x))\ by (simp add: bounded_linear comp1) show "\K. \a b. norm (prod (g a) b) \ norm a * norm b * K" using bounded by blast qed lemma (in bounded_sesquilinear) comp2: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x y. prod x (g y))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (a + a') (g b) = prod a (g b) + prod a' (g b)" by (simp add: add add_left) show "prod a (g (b + b')) = prod a (g b) + prod a (g b')" by (simp add: add add_right) show "prod (r *\<^sub>C a) (g b) = cnj r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_left) show "prod a (g (r *\<^sub>C b)) = r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_right) interpret bounded_bilinear \(\x y. prod x (g y))\ apply (rule bounded_bilinear.flip) using _ bounded_linear apply (rule bounded_bilinear.comp1) using bounded_bilinear by (rule bounded_bilinear.flip) show "\K. \a b. norm (prod a (g b)) \ norm a * norm b * K" using bounded by blast qed lemma (in bounded_sesquilinear) comp: "bounded_clinear f \ bounded_clinear g \ bounded_sesquilinear (\x y. prod (f x) (g y))" using comp1 bounded_sesquilinear.comp2 by auto lemma bounded_clinear_const_scaleR: fixes c :: real assumes \bounded_clinear f\ shows \bounded_clinear (\ x. c *\<^sub>R f x )\ proof- have \bounded_clinear (\ x. (complex_of_real c) *\<^sub>C f x )\ by (simp add: assms bounded_clinear_const_scaleC) thus ?thesis by (simp add: scaleR_scaleC) qed lemma bounded_linear_bounded_clinear: \bounded_linear A \ \c x. A (c *\<^sub>C x) = c *\<^sub>C A x \ bounded_clinear A\ apply standard by (simp_all add: linear_simps bounded_linear.bounded) lemma comp_bounded_clinear: fixes A :: \'b::complex_normed_vector \ 'c::complex_normed_vector\ and B :: \'a::complex_normed_vector \ 'b\ assumes \bounded_clinear A\ and \bounded_clinear B\ shows \bounded_clinear (A \ B)\ by (metis clinear_compose assms(1) assms(2) bounded_clinear_axioms_def bounded_clinear_compose bounded_clinear_def o_def) lemmas isCont_scaleC [simp] = bounded_bilinear.isCont [OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] subsection \Misc 2\ lemmas sums_of_complex = bounded_linear.sums [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas summable_of_complex = bounded_linear.summable [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas suminf_of_complex = bounded_linear.suminf [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_left = bounded_linear.sums[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_left = bounded_linear.summable[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_left = bounded_linear.suminf[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_right = bounded_linear.sums[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_right = bounded_linear.summable[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_right = bounded_linear.suminf[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemma closed_scaleC: fixes S::\'a::complex_normed_vector set\ and a :: complex assumes \closed S\ shows \closed ((*\<^sub>C) a ` S)\ proof (cases \a = 0\) case True then show ?thesis apply (cases \S = {}\) by (auto simp: image_constant) next case False then have \(*\<^sub>C) a ` S = (*\<^sub>C) (inverse a) -` S\ by (auto simp add: rev_image_eqI) moreover have \closed ((*\<^sub>C) (inverse a) -` S)\ by (simp add: assms continuous_closed_vimage) ultimately show ?thesis by simp qed lemma closure_scaleC: fixes S::\'a::complex_normed_vector set\ shows \closure ((*\<^sub>C) a ` S) = (*\<^sub>C) a ` closure S\ proof have \closed (closure S)\ by simp show "closure ((*\<^sub>C) a ` S) \ (*\<^sub>C) a ` closure S" by (simp add: closed_scaleC closure_minimal closure_subset image_mono) have "x \ closure ((*\<^sub>C) a ` S)" if "x \ (*\<^sub>C) a ` closure S" for x :: 'a proof- obtain t where \x = ((*\<^sub>C) a) t\ and \t \ closure S\ using \x \ (*\<^sub>C) a ` closure S\ by auto have \\s. (\n. s n \ S) \ s \ t\ using \t \ closure S\ Elementary_Topology.closure_sequential by blast then obtain s where \\n. s n \ S\ and \s \ t\ by blast have \(\ n. scaleC a (s n) \ ((*\<^sub>C) a ` S))\ using \\n. s n \ S\ by blast moreover have \(\ n. scaleC a (s n)) \ x\ proof- have \isCont (scaleC a) t\ by simp thus ?thesis using \s \ t\ \x = ((*\<^sub>C) a) t\ by (simp add: isCont_tendsto_compose) qed ultimately show ?thesis using Elementary_Topology.closure_sequential by metis qed thus "(*\<^sub>C) a ` closure S \ closure ((*\<^sub>C) a ` S)" by blast qed lemma onorm_scalarC: fixes f :: \'a::complex_normed_vector \ 'b::complex_normed_vector\ assumes a1: \bounded_clinear f\ shows \onorm (\ x. r *\<^sub>C (f x)) = (cmod r) * onorm f\ proof- have \(norm (f x)) / norm x \ onorm f\ for x using a1 by (simp add: bounded_clinear.bounded_linear le_onorm) hence t2: \bdd_above {(norm (f x)) / norm x | x. True}\ by fastforce have \continuous_on UNIV ( (*) w ) \ for w::real by simp hence \isCont ( ((*) (cmod r)) ) x\ for x by simp hence t3: \continuous (at_left (Sup {(norm (f x)) / norm x | x. True})) ((*) (cmod r))\ using Elementary_Topology.continuous_at_imp_continuous_within by blast have \{(norm (f x)) / norm x | x. True} \ {}\ by blast moreover have \mono ((*) (cmod r))\ by (simp add: monoI ordered_comm_semiring_class.comm_mult_left_mono) ultimately have \Sup {((*) (cmod r)) ((norm (f x)) / norm x) | x. True} = ((*) (cmod r)) (Sup {(norm (f x)) / norm x | x. True})\ using t2 t3 by (simp add: continuous_at_Sup_mono full_SetCompr_eq image_image) hence \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (cmod r) * (Sup {(norm (f x)) / norm x | x. True})\ by blast moreover have \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (SUP x. cmod r * norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) moreover have \(Sup {(norm (f x)) / norm x | x. True}) = (SUP x. norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) ultimately have t1: "(SUP x. cmod r * norm (f x) / norm x) = cmod r * (SUP x. norm (f x) / norm x)" by simp have \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. norm ( (\ t. r *\<^sub>C (f t)) x) / norm x)\ by (simp add: onorm_def) hence \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. (cmod r) * (norm (f x)) / norm x)\ by simp also have \... = (cmod r) * (SUP x. (norm (f x)) / norm x)\ using t1. finally show ?thesis by (simp add: onorm_def) qed lemma onorm_scaleC_left_lemma: fixes f :: "'a::complex_normed_vector" assumes r: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" proof (rule onorm_bound) fix x have "norm (r x *\<^sub>C f) = norm (r x) * norm f" by simp also have "\ \ onorm r * norm x * norm f" by (simp add: bounded_clinear.bounded_linear mult.commute mult_left_mono onorm r) finally show "norm (r x *\<^sub>C f) \ onorm r * norm f * norm x" by (simp add: ac_simps) show "0 \ onorm r * norm f" by (simp add: bounded_clinear.bounded_linear onorm_pos_le r) qed lemma onorm_scaleC_left: fixes f :: "'a::complex_normed_vector" assumes f: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) = onorm r * norm f" proof (cases "f = 0") assume "f \ 0" show ?thesis proof (rule order_antisym) show "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" using f by (rule onorm_scaleC_left_lemma) next have bl1: "bounded_clinear (\x. r x *\<^sub>C f)" by (metis bounded_clinear_scaleC_const f) have x1:"bounded_clinear (\x. r x * norm f)" by (metis bounded_clinear_mult_const f) have "onorm r \ onorm (\x. r x * complex_of_real (norm f)) / norm f" if "onorm r \ onorm (\x. r x * complex_of_real (norm f)) * cmod (1 / complex_of_real (norm f))" and "f \ 0" using that by (metis complex_of_real_cmod complex_of_real_nn_iff field_class.field_divide_inverse inverse_eq_divide nice_ordered_field_class.zero_le_divide_1_iff norm_ge_zero of_real_1 of_real_divide of_real_eq_iff) hence "onorm r \ onorm (\x. r x * norm f) * inverse (norm f)" using \f \ 0\ onorm_scaleC_left_lemma[OF x1, of "inverse (norm f)"] by (simp add: inverse_eq_divide) also have "onorm (\x. r x * norm f) \ onorm (\x. r x *\<^sub>C f)" proof (rule onorm_bound) have "bounded_linear (\x. r x *\<^sub>C f)" using bl1 bounded_clinear.bounded_linear by auto thus "0 \ onorm (\x. r x *\<^sub>C f)" by (rule Operator_Norm.onorm_pos_le) show "cmod (r x * complex_of_real (norm f)) \ onorm (\x. r x *\<^sub>C f) * norm x" for x :: 'b by (smt \bounded_linear (\x. r x *\<^sub>C f)\ complex_of_real_cmod complex_of_real_nn_iff complex_scaleC_def norm_ge_zero norm_scaleC of_real_eq_iff onorm) qed finally show "onorm r * norm f \ onorm (\x. r x *\<^sub>C f)" using \f \ 0\ by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) qed qed (simp add: onorm_zero) subsection \Finite dimension and canonical basis\ lemma vector_finitely_spanned: assumes \z \ cspan T\ shows \\ S. finite S \ S \ T \ z \ cspan S\ proof- have \\ S r. finite S \ S \ T \ z = (\a\S. r a *\<^sub>C a)\ using complex_vector.span_explicit[where b = "T"] assms by auto then obtain S r where \finite S\ and \S \ T\ and \z = (\a\S. r a *\<^sub>C a)\ by blast thus ?thesis by (meson complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset subset_iff) qed setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a set \ 'a set\)\ class cfinite_dim = complex_vector + assumes cfinitely_spanned: "\S::'a set. finite S \ cspan S = UNIV" class basis_enum = complex_vector + fixes canonical_basis :: "'a list" assumes distinct_canonical_basis[simp]: "distinct canonical_basis" and is_cindependent_set[simp]: "cindependent (set canonical_basis)" and is_generator_set[simp]: "cspan (set canonical_basis) = UNIV" setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a::complex_vector set \ 'a set\)\ lemma cdim_UNIV_basis_enum[simp]: \cdim (UNIV::'a::basis_enum set) = length (canonical_basis::'a list)\ apply (subst is_generator_set[symmetric]) apply (subst complex_vector.dim_span_eq_card_independent) apply (rule is_cindependent_set) using distinct_canonical_basis distinct_card by blast lemma finite_basis: "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = UNIV" proof - from cfinitely_spanned obtain S :: \'a set\ where \finite S\ and \cspan S = UNIV\ by auto from complex_vector.maximal_independent_subset obtain B :: \'a set\ where \B \ S\ and \cindependent B\ and \S \ cspan B\ by metis moreover have \finite B\ using \B \ S\ \finite S\ by (meson finite_subset) moreover have \cspan B = UNIV\ using \cspan S = UNIV\ \S \ cspan B\ by (metis complex_vector.span_eq top_greatest) ultimately show ?thesis by auto qed instance basis_enum \ cfinite_dim apply intro_classes apply (rule exI[of _ \set canonical_basis\]) using is_cindependent_set is_generator_set by auto lemma cindependent_cfinite_dim_finite: assumes \cindependent (S::'a::cfinite_dim set)\ shows \finite S\ by (metis assms cfinitely_spanned complex_vector.independent_span_bound top_greatest) lemma cfinite_dim_finite_subspace_basis: assumes \csubspace X\ shows "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = X" by (meson assms cindependent_cfinite_dim_finite complex_vector.basis_exists complex_vector.span_subspace) text \The following auxiliary lemma (\finite_span_complete_aux\) shows more or less the same as \finite_span_representation_bounded\, \finite_span_complete\ below (see there for an intuition about the mathematical content of the lemmas). However, there is one difference: Here we additionally assume here that there is a bijection rep/abs between a finite type \<^typ>\'basis\ and the set $B$. This is needed to be able to use results about euclidean spaces that are formulated w.r.t. the type class \<^class>\finite\ Since we anyway assume that $B$ is finite, this added assumption does not make the lemma weaker. However, we cannot derive the existence of \<^typ>\'basis\ inside the proof (HOL does not support such reasoning). Therefore we have the type \<^typ>\'basis\ as an explicit assumption and remove it using @{attribute internalize_sort} after the proof.\ lemma finite_span_complete_aux: fixes b :: "'b::real_normed_vector" and B :: "'b set" and rep :: "'basis::finite \ 'b" and abs :: "'b \ 'basis" assumes t: "type_definition rep abs B" and t1: "finite B" and t2: "b\B" and t3: "independent B" shows "\D>0. \\. norm (representation B \ b) \ norm \ * D" and "complete (span B)" proof - define repr where "repr = real_vector.representation B" define repr' where "repr' \ = Abs_euclidean_space (repr \ o rep)" for \ define comb where "comb l = (\b\B. l b *\<^sub>R b)" for l define comb' where "comb' l = comb (Rep_euclidean_space l o abs)" for l have comb_cong: "comb x = comb y" if "\z. z\B \ x z = y z" for x y unfolding comb_def using that by auto have comb_repr[simp]: "comb (repr \) = \" if "\ \ real_vector.span B" for \ using \comb \ \l. \b\B. l b *\<^sub>R b\ local.repr_def real_vector.sum_representation_eq t1 t3 that by fastforce have w5:"(\b | (b \ B \ x b \ 0) \ b \ B. x b *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" for x using \finite B\ by (smt DiffD1 DiffD2 mem_Collect_eq real_vector.scale_eq_0_iff subset_eq sum.mono_neutral_left) have "representation B (\b\B. x b *\<^sub>R b) = (\b. if b \ B then x b else 0)" for x proof (rule real_vector.representation_eqI) show "independent B" by (simp add: t3) show "(\b\B. x b *\<^sub>R b) \ span B" by (meson real_vector.span_scale real_vector.span_sum real_vector.span_superset subset_iff) show "b \ B" if "(if b \ B then x b else 0) \ 0" for b :: 'b using that by meson show "finite {b. (if b \ B then x b else 0) \ 0}" using t1 by auto show "(\b | (if b \ B then x b else 0) \ 0. (if b \ B then x b else 0) *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" using w5 by simp qed hence repr_comb[simp]: "repr (comb x) = (\b. if b\B then x b else 0)" for x unfolding repr_def comb_def. have repr_bad[simp]: "repr \ = (\_. 0)" if "\ \ real_vector.span B" for \ unfolding repr_def using that by (simp add: real_vector.representation_def) have [simp]: "repr' \ = 0" if "\ \ real_vector.span B" for \ unfolding repr'_def repr_bad[OF that] apply transfer by auto have comb'_repr'[simp]: "comb' (repr' \) = \" if "\ \ real_vector.span B" for \ proof - have x1: "(repr \ \ rep \ abs) z = repr \ z" if "z \ B" for z unfolding o_def using t that type_definition.Abs_inverse by fastforce have "comb' (repr' \) = comb ((repr \ \ rep) \ abs)" unfolding comb'_def repr'_def by (subst Abs_euclidean_space_inverse; simp) also have "\ = comb (repr \)" using x1 comb_cong by blast also have "\ = \" using that by simp finally show ?thesis by - qed have t1: "Abs_euclidean_space (Rep_euclidean_space t) = t" if "\x. rep x \ B" for t::"'a euclidean_space" apply (subst Rep_euclidean_space_inverse) by simp have "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x y else 0) = x" for x using type_definition.Rep[OF t] apply simp using t1 by blast hence "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x (abs (rep y)) else 0) = x" for x apply (subst type_definition.Rep_inverse[OF t]) by simp hence repr'_comb'[simp]: "repr' (comb' x) = x" for x unfolding comb'_def repr'_def o_def by simp have sphere: "compact (sphere 0 d :: 'basis euclidean_space set)" for d using compact_sphere by blast have "complete (UNIV :: 'basis euclidean_space set)" by (simp add: complete_UNIV) have "(\b\B. (Rep_euclidean_space (x + y) \ abs) b *\<^sub>R b) = (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b) + (\b\B. (Rep_euclidean_space y \ abs) b *\<^sub>R b)" for x :: "'basis euclidean_space" and y :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: scaleR_add_left sum.distrib) moreover have "(\b\B. (Rep_euclidean_space (c *\<^sub>R x) \ abs) b *\<^sub>R b) = c *\<^sub>R (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b)" for c :: real and x :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: real_vector.scale_sum_right) ultimately have blin_comb': "bounded_linear comb'" unfolding comb_def comb'_def by (rule bounded_linearI') hence "continuous_on X comb'" for X by (simp add: linear_continuous_on) hence "compact (comb' ` sphere 0 d)" for d using sphere by (rule compact_continuous_image) hence compact_norm_comb': "compact (norm ` comb' ` sphere 0 1)" using compact_continuous_image continuous_on_norm_id by blast have not0: "0 \ norm ` comb' ` sphere 0 1" proof (rule ccontr, simp) assume "0 \ norm ` comb' ` sphere 0 1" then obtain x where nc0: "norm (comb' x) = 0" and x: "x \ sphere 0 1" by auto hence "comb' x = 0" by simp hence "repr' (comb' x) = 0" unfolding repr'_def o_def repr_def apply simp by (smt repr'_comb' blin_comb' dist_0_norm linear_simps(3) mem_sphere norm_zero x) hence "x = 0" by auto with x show False by simp qed have "closed (norm ` comb' ` sphere 0 1)" using compact_imp_closed compact_norm_comb' by blast moreover have "0 \ norm ` comb' ` sphere 0 1" by (simp add: not0) ultimately have "\d>0. \x\norm ` comb' ` sphere 0 1. d \ dist 0 x" by (meson separate_point_closed) then obtain d where d: "x\norm ` comb' ` sphere 0 1 \ d \ dist 0 x" and "d > 0" for x by metis define D where "D = 1/d" hence "D > 0" using \d>0\ unfolding D_def by auto have "x \ d" if "x\norm ` comb' ` sphere 0 1" for x using d that apply auto by fastforce hence *: "norm (comb' x) \ d" if "norm x = 1" for x using that by auto have norm_comb': "norm (comb' x) \ d * norm x" for x proof (cases "x=0") show "d * norm x \ norm (comb' x)" if "x = 0" using that by simp show "d * norm x \ norm (comb' x)" if "x \ 0" using that using *[of "(1/norm x) *\<^sub>R x"] unfolding linear_simps(5)[OF blin_comb'] apply auto by (simp add: le_divide_eq) qed have *: "norm (repr' \) \ norm \ * D" for \ proof (cases "\ \ real_vector.span B") show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that unfolding D_def using norm_comb'[of "repr' \"] \d>0\ by (simp_all add: linordered_field_class.mult_imp_le_div_pos mult.commute) show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that \0 < D\ by auto qed hence "norm (Rep_euclidean_space (repr' \) (abs b)) \ norm \ * D" for \ proof - have "(Rep_euclidean_space (repr' \) (abs b)) = repr' \ \ euclidean_space_basis_vector (abs b)" apply (transfer fixing: abs b) by auto also have "\\\ \ norm (repr' \)" apply (rule Basis_le_norm) unfolding Basis_euclidean_space_def by simp also have "\ \ norm \ * D" using * by auto finally show ?thesis by simp qed hence "norm (repr \ b) \ norm \ * D" for \ unfolding repr'_def by (smt \comb' \ \l. comb (Rep_euclidean_space l \ abs)\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' comp_apply norm_le_zero_iff repr_bad repr_comb) thus "\D>0. \\. norm (repr \ b) \ norm \ * D" using \D>0\ by auto from \d>0\ have complete_comb': "complete (comb' ` UNIV)" proof (rule complete_isometric_image) show "subspace (UNIV::'basis euclidean_space set)" by simp show "bounded_linear comb'" by (simp add: blin_comb') show "\x\UNIV. d * norm x \ norm (comb' x)" by (simp add: norm_comb') show "complete (UNIV::'basis euclidean_space set)" by (simp add: \complete UNIV\) qed have range_comb': "comb' ` UNIV = real_vector.span B" proof (auto simp: image_def) show "comb' x \ real_vector.span B" for x by (metis comb'_def comb_cong comb_repr local.repr_def repr_bad repr_comb real_vector.representation_zero real_vector.span_zero) next fix \ assume "\ \ real_vector.span B" then obtain f where f: "comb f = \" apply atomize_elim unfolding span_finite[OF \finite B\] comb_def by auto define f' where "f' b = (if b\B then f b else 0)" for b :: 'b have f': "comb f' = \" unfolding f[symmetric] apply (rule comb_cong) unfolding f'_def by simp define x :: "'basis euclidean_space" where "x = Abs_euclidean_space (f' o rep)" have "\ = comb' x" by (metis (no_types, lifting) \\ \ span B\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' f' fun.map_cong repr_comb t type_definition.Rep_range x_def) thus "\x. \ = comb' x" by auto qed from range_comb' complete_comb' show "complete (real_vector.span B)" by simp qed lemma finite_span_complete[simp]: fixes A :: "'a::real_normed_vector set" assumes "finite A" shows "complete (span A)" text \The span of a finite set is complete.\ proof (cases "A \ {} \ A \ {0}") case True obtain B where BT: "real_vector.span B = real_vector.span A" and "independent B" and "finite B" by (meson True assms finite_subset real_vector.maximal_independent_subset real_vector.span_eq real_vector.span_superset subset_trans) have "B\{}" apply (rule ccontr, simp) using BT True by (metis real_vector.span_superset real_vector.span_empty subset_singletonD) (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) { (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux, otherwise "internalize_sort" below fails *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) note finite_span_complete_aux(2)[internalize_sort "'basis::finite"] note this[OF basisT_finite t] } note this[cancel_type_definition, OF \B\{}\ \finite B\ _ \independent B\] hence "complete (real_vector.span B)" using \B\{}\ by auto thus "complete (real_vector.span A)" unfolding BT by simp next case False thus ?thesis using complete_singleton by auto qed lemma finite_span_representation_bounded: fixes B :: "'a::real_normed_vector set" assumes "finite B" and "independent B" shows "\D>0. \\ b. abs (representation B \ b) \ norm \ * D" text \ Assume $B$ is a finite linear independent set of vectors (in a real normed vector space). Let $\alpha^\psi_b$ be the coefficients of $\psi$ expressed as a linear combination over $B$. Then $\alpha$ is is uniformly cblinfun (i.e., $\lvert\alpha^\psi_b \leq D \lVert\psi\rVert\psi$ for some $D$ independent of $\psi,b$). (This also holds when $b$ is not in the span of $B$ because of the way \real_vector.representation\ is defined in this corner case.)\ proof (cases "B\{}") case True (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) define repr where "repr = real_vector.representation B" { (* Step 1: Create a fake type definition by introducing a new type variable 'basis and then assuming the existence of the morphisms Rep/Abs to B This is then roughly equivalent to "typedef 'basis = B" *) (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux (I.e., we cannot call it 'basis) *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto (* Step 2: We show that our fake typedef 'basisT could be instantiated as type class finite *) have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) (* Step 3: We take the finite_span_complete_aux and remove the requirement that 'basis::finite (instead, a precondition "class.finite TYPE('basisT)" is introduced) *) note finite_span_complete_aux(1)[internalize_sort "'basis::finite"] (* Step 4: We instantiate the premises *) note this[OF basisT_finite t] } (* Now we have the desired fact, except that it still assumes that B is isomorphic to some type 'basis together with the assumption that there are morphisms between 'basis and B. 'basis and that premise are removed using cancel_type_definition *) note this[cancel_type_definition, OF True \finite B\ _ \independent B\] hence d2:"\D. \\. D>0 \ norm (repr \ b) \ norm \ * D" if \b\B\ for b by (simp add: repr_def that True) have d1: " (\b. b \ B \ \D. \\. 0 < D \ norm (repr \ b) \ norm \ * D) \ \D. \b \. b \ B \ 0 < D b \ norm (repr \ b) \ norm \ * D b" apply (rule choice) by auto then obtain D where D: "D b > 0 \ norm (repr \ b) \ norm \ * D b" if "b\B" for b \ apply atomize_elim using d2 by blast hence Dpos: "D b > 0" and Dbound: "norm (repr \ b) \ norm \ * D b" if "b\B" for b \ using that by auto define Dall where "Dall = Max (D`B)" have "Dall > 0" unfolding Dall_def using \finite B\ \B\{}\ Dpos by (metis (mono_tags, lifting) Max_in finite_imageI image_iff image_is_empty) have "Dall \ D b" if "b\B" for b unfolding Dall_def using \finite B\ that by auto with Dbound have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ using that by (smt mult_left_mono norm_not_less_zero) moreover have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ unfolding repr_def using real_vector.representation_ne_zero True by (metis calculation empty_subsetI less_le_trans local.repr_def norm_ge_zero norm_zero not_less subsetI subset_antisym) ultimately show "\D>0. \\ b. abs (repr \ b) \ norm \ * D" using \Dall > 0\ real_norm_def by metis next case False thus ?thesis unfolding repr_def using real_vector.representation_ne_zero[of B] using nice_ordered_field_class.linordered_field_no_ub by fastforce qed hide_fact finite_span_complete_aux lemma finite_cspan_complete[simp]: fixes B :: "'a::complex_normed_vector set" assumes "finite B" shows "complete (cspan B)" by (simp add: assms cspan_as_span) lemma finite_span_closed[simp]: fixes B :: "'a::real_normed_vector set" assumes "finite B" shows "closed (real_vector.span B)" by (simp add: assms complete_imp_closed) lemma finite_cspan_closed[simp]: fixes S::\'a::complex_normed_vector set\ assumes a1: \finite S\ shows \closed (cspan S)\ by (simp add: assms complete_imp_closed) lemma closure_finite_cspan: fixes T::\'a::complex_normed_vector set\ assumes \finite T\ shows \closure (cspan T) = cspan T\ by (simp add: assms) lemma finite_cspan_crepresentation_bounded: fixes B :: "'a::complex_normed_vector set" assumes a1: "finite B" and a2: "cindependent B" shows "\D>0. \\ b. norm (crepresentation B \ b) \ norm \ * D" proof - define B' where "B' = (B \ scaleC \ ` B)" have independent_B': "independent B'" using B'_def \cindependent B\ by (simp add: real_independent_from_complex_independent a1) have "finite B'" unfolding B'_def using \finite B\ by simp obtain D' where "D' > 0" and D': "norm (real_vector.representation B' \ b) \ norm \ * D'" for \ b apply atomize_elim using independent_B' \finite B'\ by (simp add: finite_span_representation_bounded) define D where "D = 2*D'" from \D' > 0\ have \D > 0\ unfolding D_def by simp have "norm (crepresentation B \ b) \ norm \ * D" for \ b proof (cases "b\B") case True have d3: "norm \ = 1" by simp have "norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm \ * norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_scaleC by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d3 by simp finally have d2:"norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))". have "norm (crepresentation B \ b) = norm (complex_of_real (real_vector.representation B' \ b) + \ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" by (simp add: B'_def True a1 a2 crepresentation_from_representation) also have "\ \ norm (complex_of_real (real_vector.representation B' \ b)) + norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_triangle_ineq by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ b)) + norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d2 by simp also have "\ = norm (real_vector.representation B' \ b) + norm (real_vector.representation B' \ (\ *\<^sub>C b))" by simp also have "\ \ norm \ * D' + norm \ * D'" by (rule add_mono; rule D') also have "\ \ norm \ * D" unfolding D_def by linarith finally show ?thesis by auto next case False hence "crepresentation B \ b = 0" using complex_vector.representation_ne_zero by blast thus ?thesis by (smt \0 < D\ norm_ge_zero norm_zero split_mult_pos_le) qed with \D > 0\ show ?thesis by auto qed lemma bounded_clinear_finite_dim[simp]: fixes f :: \'a::{cfinite_dim,complex_normed_vector} \ 'b::complex_normed_vector\ assumes \clinear f\ shows \bounded_clinear f\ proof - include notation_norm obtain basis :: \'a set\ where b1: "complex_vector.span basis = UNIV" and b2: "cindependent basis" and b3:"finite basis" using finite_basis by auto have "\C>0. \\ b. cmod (crepresentation basis \ b) \ \\\ * C" using finite_cspan_crepresentation_bounded[where B = basis] b2 b3 by blast then obtain C where s1: "cmod (crepresentation basis \ b) \ \\\ * C" and s2: "C > 0" for \ b by blast define M where "M = C * (\a\basis. \f a\)" have "\f x\ \ \x\ * M" for x proof- define r where "r b = crepresentation basis x b" for b have x_span: "x \ complex_vector.span basis" by (simp add: b1) have f0: "v \ basis" if "r v \ 0" for v using complex_vector.representation_ne_zero r_def that by auto have w:"{a|a. r a \ 0} \ basis" using f0 by blast hence f1: "finite {a|a. r a \ 0}" using b3 rev_finite_subset by auto have f2: "(\a| r a \ 0. r a *\<^sub>C a) = x" unfolding r_def using b2 complex_vector.sum_nonzero_representation_eq x_span Collect_cong by fastforce have g1: "(\a\basis. crepresentation basis x a *\<^sub>C a) = x" by (simp add: b2 b3 complex_vector.sum_representation_eq x_span) have f3: "(\a\basis. r a *\<^sub>C a) = x" unfolding r_def by (simp add: g1) hence "f x = f (\a\basis. r a *\<^sub>C a)" by simp also have "\ = (\a\basis. r a *\<^sub>C f a)" by (smt (verit, ccfv_SIG) assms complex_vector.linear_scale complex_vector.linear_sum sum.cong) finally have "f x = (\a\basis. r a *\<^sub>C f a)". hence "\f x\ = \(\a\basis. r a *\<^sub>C f a)\" by simp also have "\ \ (\a\basis. \r a *\<^sub>C f a\)" by (simp add: sum_norm_le) also have "\ \ (\a\basis. \r a\ * \f a\)" by simp also have "\ \ (\a\basis. \x\ * C * \f a\)" using sum_mono s1 unfolding r_def by (simp add: sum_mono mult_right_mono) also have "\ \ \x\ * C * (\a\basis. \f a\)" using sum_distrib_left by (smt sum.cong) also have "\ = \x\ * M" unfolding M_def by linarith finally show ?thesis . qed thus ?thesis using assms bounded_clinear_def bounded_clinear_axioms_def by blast qed subsection \Closed subspaces\ lemma csubspace_INF[simp]: "(\x. x \ A \ csubspace x) \ csubspace (\A)" by (simp add: complex_vector.subspace_Inter) locale closed_csubspace = fixes A::"('a::{complex_vector,topological_space}) set" assumes subspace: "csubspace A" assumes closed: "closed A" declare closed_csubspace.subspace[simp] lemma closure_is_csubspace[simp]: fixes A::"('a::complex_normed_vector) set" assumes \csubspace A\ shows \csubspace (closure A)\ proof- have "x \ closure A \ y \ closure A \ x+y \ closure A" for x y proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast assume \y\(closure A)\ then obtain yy where \\ n::nat. yy n \ A\ and \yy \ y\ using closure_sequential by blast have \\ n::nat. (xx n) + (yy n) \ A\ using \\n. xx n \ A\ \\n. yy n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) hence \(\ n. (xx n) + (yy n)) \ x + y\ using \xx \ x\ \yy \ y\ by (simp add: tendsto_add) thus ?thesis using \\ n::nat. (xx n) + (yy n) \ A\ by (meson closure_sequential) qed moreover have "x\(closure A) \ c *\<^sub>C x \ (closure A)" for x c proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast have \\ n::nat. c *\<^sub>C (xx n) \ A\ using \\n. xx n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) have \isCont (\ t. c *\<^sub>C t) x\ using bounded_clinear.bounded_linear bounded_clinear_scaleC_right linear_continuous_at by auto hence \(\ n. c *\<^sub>C (xx n)) \ c *\<^sub>C x\ using \xx \ x\ by (simp add: isCont_tendsto_compose) thus ?thesis using \\ n::nat. c *\<^sub>C (xx n) \ A\ by (meson closure_sequential) qed moreover have "0 \ (closure A)" using assms closure_subset complex_vector.subspace_def by (metis in_mono) ultimately show ?thesis by (simp add: complex_vector.subspaceI) qed lemma csubspace_set_plus: assumes \csubspace A\ and \csubspace B\ shows \csubspace (A + B)\ proof - define C where \C = {\+\| \ \. \\A \ \\B}\ have "x\C \ y\C \ x+y\C" for x y using C_def assms(1) assms(2) complex_vector.subspace_add complex_vector.subspace_sums by blast moreover have "c *\<^sub>C x \ C" if \x\C\ for x c proof - have "csubspace C" by (simp add: C_def assms(1) assms(2) complex_vector.subspace_sums) then show ?thesis using that by (simp add: complex_vector.subspace_def) qed moreover have "0 \ C" using \C = {\ + \ |\ \. \ \ A \ \ \ B}\ add.inverse_neutral add_uminus_conv_diff assms(1) assms(2) diff_0 mem_Collect_eq add.right_inverse by (metis (mono_tags, lifting) complex_vector.subspace_0) ultimately show ?thesis unfolding C_def complex_vector.subspace_def by (smt mem_Collect_eq set_plus_elim set_plus_intro) qed lemma closed_csubspace_0[simp]: "closed_csubspace ({0} :: ('a::{complex_vector,t1_space}) set)" proof- have \csubspace {0}\ using add.right_neutral complex_vector.subspace_def scaleC_right.zero by blast moreover have "closed ({0} :: 'a set)" by simp ultimately show ?thesis by (simp add: closed_csubspace_def) qed lemma closed_csubspace_UNIV[simp]: "closed_csubspace (UNIV::('a::{complex_vector,topological_space}) set)" proof- have \csubspace UNIV\ by simp moreover have \closed UNIV\ by simp ultimately show ?thesis unfolding closed_csubspace_def by auto qed lemma closed_csubspace_inter[simp]: assumes "closed_csubspace A" and "closed_csubspace B" shows "closed_csubspace (A\B)" proof- obtain C where \C = A \ B\ by blast have \csubspace C\ proof- have "x\C \ y\C \ x+y\C" for x y by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "x\C \ c *\<^sub>C x \ C" for x c by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "0 \ C" using \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def by fastforce ultimately show ?thesis by (simp add: complex_vector.subspace_def) qed moreover have \closed C\ using \C = A \ B\ by (simp add: assms(1) assms(2) closed_Int closed_csubspace.closed) ultimately show ?thesis using \C = A \ B\ by (simp add: closed_csubspace_def) qed lemma closed_csubspace_INF[simp]: assumes a1: "\A\\. closed_csubspace A" shows "closed_csubspace (\\)" proof- have \csubspace (\\)\ by (simp add: assms closed_csubspace.subspace complex_vector.subspace_Inter) moreover have \closed (\\)\ by (simp add: assms closed_Inter closed_csubspace.closed) ultimately show ?thesis by (simp add: closed_csubspace.intro) qed typedef (overloaded) ('a::"{complex_vector,topological_space}") ccsubspace = \{S::'a set. closed_csubspace S}\ morphisms space_as_set Abs_clinear_space using Complex_Vector_Spaces.closed_csubspace_UNIV by blast setup_lifting type_definition_ccsubspace lemma csubspace_space_as_set[simp]: \csubspace (space_as_set S)\ by (metis closed_csubspace_def mem_Collect_eq space_as_set) instantiation ccsubspace :: (complex_normed_vector) scaleC begin lift_definition scaleC_ccsubspace :: "complex \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>C) c ` S" proof show "csubspace ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: closed_scaleC closed_csubspace.closed) qed lift_definition scaleR_ccsubspace :: "real \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>R) c ` S" proof show "csubspace ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that using bounded_clinear_def bounded_clinear_scaleC_right scaleR_scaleC by (simp add: scaleR_scaleC closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that by (simp add: closed_scaling closed_csubspace.closed) qed instance proof show "((*\<^sub>R) r::'a ccsubspace \ _) = (*\<^sub>C) (complex_of_real r)" for r :: real by (simp add: scaleR_scaleC scaleC_ccsubspace_def scaleR_ccsubspace_def) qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") bot begin lift_definition bot_ccsubspace :: \'a ccsubspace\ is \{0}\ by simp instance.. end lemma zero_cblinfun_image[simp]: "0 *\<^sub>C S = bot" for S :: "_ ccsubspace" proof transfer have "(0::'b) \ (\x. 0) ` S" if "closed_csubspace S" for S::"'b set" using that unfolding closed_csubspace_def by (simp add: complex_vector.linear_subspace_image complex_vector.module_hom_zero complex_vector.subspace_0) thus "(*\<^sub>C) 0 ` S = {0::'b}" if "closed_csubspace (S::'b set)" for S :: "'b set" using that by (auto intro !: exI [of _ 0]) qed lemma csubspace_scaleC_invariant: fixes a S assumes \a \ 0\ and \csubspace S\ shows \(*\<^sub>C) a ` S = S\ proof- have \x \ (*\<^sub>C) a ` S \ x \ S\ for x using assms(2) complex_vector.subspace_scale by blast moreover have \x \ S \ x \ (*\<^sub>C) a ` S\ for x proof - assume "x \ S" hence "\c aa. (c / a) *\<^sub>C aa \ S \ c *\<^sub>C aa = x" using assms(2) complex_vector.subspace_def scaleC_one by metis hence "\aa. aa \ S \ a *\<^sub>C aa = x" using assms(1) by auto thus ?thesis by (meson image_iff) qed ultimately show ?thesis by blast qed lemma ccsubspace_scaleC_invariant[simp]: "a \ 0 \ a *\<^sub>C S = S" for S :: "_ ccsubspace" apply transfer by (simp add: closed_csubspace.subspace csubspace_scaleC_invariant) instantiation ccsubspace :: ("{complex_vector,topological_space}") "top" begin lift_definition top_ccsubspace :: \'a ccsubspace\ is \UNIV\ by simp instance .. end lemma ccsubspace_top_not_bot[simp]: "(top::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ bot" (* The type class t1_space is needed because the definition of bot in ccsubspace needs it *) by (metis UNIV_not_singleton bot_ccsubspace.rep_eq top_ccsubspace.rep_eq) lemma ccsubspace_bot_not_top[simp]: "(bot::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ top" using ccsubspace_top_not_bot by metis instantiation ccsubspace :: ("{complex_vector,topological_space}") "Inf" begin lift_definition Inf_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\ S. \ S\ proof fix S :: "'a set set" assume closed: "closed_csubspace x" if \x \ S\ for x show "csubspace (\ S::'a set)" by (simp add: closed closed_csubspace.subspace) show "closed (\ S::'a set)" by (simp add: closed closed_csubspace.closed) qed instance .. end lift_definition ccspan :: "'a::complex_normed_vector set \ 'a ccsubspace" is "\G. closure (cspan G)" proof (rule closed_csubspace.intro) fix S :: "'a set" show "csubspace (closure (cspan S))" by (simp add: closure_is_csubspace) show "closed (closure (cspan S))" by simp qed lemma ccspan_canonical_basis[simp]: "ccspan (set canonical_basis) = top" using ccspan.rep_eq space_as_set_inject top_ccsubspace.rep_eq closure_UNIV is_generator_set by metis lemma ccspan_Inf_def: \ccspan A = Inf {S. A \ space_as_set S}\ for A::\('a::cbanach) set\ proof- have \x \ space_as_set (ccspan A) \ x \ space_as_set (Inf {S. A \ space_as_set S})\ for x::'a proof- assume \x \ space_as_set (ccspan A)\ hence "x \ closure (cspan A)" by (simp add: ccspan.rep_eq) hence \x \ closure (complex_vector.span A)\ unfolding ccspan_def by simp hence \\ y::nat \ 'a. (\ n. y n \ (complex_vector.span A)) \ y \ x\ by (simp add: closure_sequential) then obtain y where \\ n. y n \ (complex_vector.span A)\ and \y \ x\ by blast have \y n \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ for n using \\ n. y n \ (complex_vector.span A)\ by auto have \closed_csubspace S \ closed S\ for S::\'a set\ by (simp add: closed_csubspace.closed) hence \closed ( \ {S. (complex_vector.span A) \ S \ closed_csubspace S})\ by simp hence \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using \y \ x\ using \\n. y n \ \ {S. complex_vector.span A \ S \ closed_csubspace S}\ closed_sequentially by blast moreover have \{S. A \ S \ closed_csubspace S} \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using Collect_mono_iff by (simp add: Collect_mono_iff closed_csubspace.subspace complex_vector.span_minimal) ultimately have \x \ \ {S. A \ S \ closed_csubspace S}\ by blast moreover have "(x::'a) \ \ {x. A \ x \ closed_csubspace x}" if "(x::'a) \ \ {S. A \ S \ closed_csubspace S}" for x :: 'a and A :: "'a set" using that by simp ultimately show \x \ space_as_set (Inf {S. A \ space_as_set S})\ apply transfer. qed moreover have \x \ space_as_set (Inf {S. A \ space_as_set S}) \ x \ space_as_set (ccspan A)\ for x::'a proof- assume \x \ space_as_set (Inf {S. A \ space_as_set S})\ hence \x \ \ {S. A \ S \ closed_csubspace S}\ apply transfer by blast moreover have \{S. (complex_vector.span A) \ S \ closed_csubspace S} \ {S. A \ S \ closed_csubspace S}\ using Collect_mono_iff complex_vector.span_superset by fastforce ultimately have \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ by blast thus \x \ space_as_set (ccspan A)\ by (metis (no_types, lifting) Inter_iff space_as_set closure_subset mem_Collect_eq ccspan.rep_eq) qed ultimately have \space_as_set (ccspan A) = space_as_set (Inf {S. A \ space_as_set S})\ by blast thus ?thesis using space_as_set_inject by auto qed lemma cspan_singleton_scaleC[simp]: "(a::complex)\0 \ cspan { a *\<^sub>C \ } = cspan {\}" for \::"'a::complex_vector" by (smt complex_vector.dependent_single complex_vector.independent_insert complex_vector.scale_eq_0_iff complex_vector.span_base complex_vector.span_redundant complex_vector.span_scale doubleton_eq_iff insert_absorb insert_absorb2 insert_commute singletonI) lemma closure_is_closed_csubspace[simp]: fixes S::\'a::complex_normed_vector set\ assumes \csubspace S\ shows \closed_csubspace (closure S)\ proof- fix x y :: 'a and c :: complex have "x + y \ closure S" if "x \ closure S" and "y \ closure S" proof- have \\ r. (\ n::nat. r n \ S) \ r \ x\ using closure_sequential that(1) by auto then obtain r where \\ n::nat. r n \ S\ and \r \ x\ by blast have \\ s. (\ n::nat. s n \ S) \ s \ y\ using closure_sequential that(2) by auto then obtain s where \\ n::nat. s n \ S\ and \s \ y\ by blast have \\ n::nat. r n + s n \ S\ using \\n. r n \ S\ \\n. s n \ S\ assms complex_vector.subspace_add by blast moreover have \(\ n. r n + s n) \ x + y\ by (simp add: \r \ x\ \s \ y\ tendsto_add) ultimately show ?thesis using assms that(1) that(2) by (simp add: complex_vector.subspace_add) qed moreover have "c *\<^sub>C x \ closure S" if "x \ closure S" proof- have \\ y. (\ n::nat. y n \ S) \ y \ x\ using Elementary_Topology.closure_sequential that by auto then obtain y where \\ n::nat. y n \ S\ and \y \ x\ by blast have \isCont (scaleC c) x\ by simp hence \(\ n. scaleC c (y n)) \ scaleC c x\ using \y \ x\ by (simp add: isCont_tendsto_compose) from \\ n::nat. y n \ S\ have \\ n::nat. scaleC c (y n) \ S\ using assms complex_vector.subspace_scale by auto thus ?thesis using assms that by (simp add: complex_vector.subspace_scale) qed moreover have "0 \ closure S" by (simp add: assms complex_vector.subspace_0) moreover have "closed (closure S)" by auto ultimately show ?thesis by (simp add: assms closed_csubspace_def) qed lemma ccspan_singleton_scaleC[simp]: "(a::complex)\0 \ ccspan {a *\<^sub>C \} = ccspan {\}" apply transfer by simp lemma clinear_continuous_at: assumes \bounded_clinear f\ shows \isCont f x\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_at) lemma clinear_continuous_within: assumes \bounded_clinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_within) lemma antilinear_continuous_at: assumes \bounded_antilinear f\ shows \isCont f x\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_at) lemma antilinear_continuous_within: assumes \bounded_antilinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_within) lemma bounded_clinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_clinear A\ and \bounded_clinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - have eq': \A t = B t\ if \t \ cspan G\ for t using _ _ that eq apply (rule complex_vector.linear_eq_on) by (auto simp: assms bounded_clinear.clinear) have \A t - B t = 0\ using _ _ t apply (rule continuous_constant_on_closure) by (auto simp add: eq' assms(1) assms(2) clinear_continuous_at continuous_at_imp_continuous_on) then show ?thesis by auto qed instantiation ccsubspace :: ("{complex_vector,topological_space}") "order" begin lift_definition less_eq_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_eq_ccsubspace_def[code del] lift_definition less_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_ccsubspace_def[code del] instance proof fix x y z :: "'a ccsubspace" show "(x < y) = (x \ y \ \ y \ x)" by (simp add: less_eq_ccsubspace.rep_eq less_le_not_le less_ccsubspace.rep_eq) show "x \ x" by (simp add: less_eq_ccsubspace.rep_eq) show "x \ z" if "x \ y" and "y \ z" using that less_eq_ccsubspace.rep_eq by auto show "x = y" if "x \ y" and "y \ x" using that by (simp add: space_as_set_inject less_eq_ccsubspace.rep_eq) qed end lemma ccspan_leqI: assumes \M \ space_as_set S\ shows \ccspan M \ S\ using assms apply transfer by (simp add: closed_csubspace.closed closure_minimal complex_vector.span_minimal) lemma ccspan_mono: assumes \A \ B\ shows \ccspan A \ ccspan B\ apply (transfer fixing: A B) by (simp add: assms closure_mono complex_vector.span_mono) lemma bounded_sesquilinear_add: \bounded_sesquilinear (\ x y. A x y + B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "A (a + a') b + B (a + a') b = (A a b + B a b) + (A a' b + B a' b)" by (simp add: bounded_sesquilinear.add_left that(1) that(2)) show \A a (b + b') + B a (b + b') = (A a b + B a b) + (A a b' + B a b')\ by (simp add: bounded_sesquilinear.add_right that(1) that(2)) show \A (r *\<^sub>C a) b + B (r *\<^sub>C a) b = cnj r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_left scaleC_add_right that(1) that(2)) show \A a (r *\<^sub>C b) + B a (r *\<^sub>C b) = r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_right scaleC_add_right that(1) that(2)) show \\K. \a b. norm (A a b + B a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \\ KB. \ a b. norm (B a b) \ norm a * norm b * KB\ by (simp add: bounded_sesquilinear.bounded that(2)) then obtain KB where \\ a b. norm (B a b) \ norm a * norm b * KB\ by blast have \norm (A a b + B a b) \ norm a * norm b * (KA + KB)\ for a b proof- have \norm (A a b + B a b) \ norm (A a b) + norm (B a b)\ using norm_triangle_ineq by blast also have \\ \ norm a * norm b * KA + norm a * norm b * KB\ using \\ a b. norm (A a b) \ norm a * norm b * KA\ \\ a b. norm (B a b) \ norm a * norm b * KB\ using add_mono by blast also have \\= norm a * norm b * (KA + KB)\ by (simp add: mult.commute ring_class.ring_distribs(2)) finally show ?thesis by blast qed thus ?thesis by blast qed qed lemma bounded_sesquilinear_uminus: \bounded_sesquilinear (\ x y. - A x y)\ if \bounded_sesquilinear A\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "- A (a + a') b = (- A a b) + (- A a' b)" by (simp add: bounded_sesquilinear.add_left that) show \- A a (b + b') = (- A a b) + (- A a b')\ by (simp add: bounded_sesquilinear.add_right that) show \- A (r *\<^sub>C a) b = cnj r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_left that) show \- A a (r *\<^sub>C b) = r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_right that) show \\K. \a b. norm (- A a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \norm (- A a b) \ norm a * norm b * KA\ for a b by (simp add: \\a b. norm (A a b) \ norm a * norm b * KA\) thus ?thesis by blast qed qed lemma bounded_sesquilinear_diff: \bounded_sesquilinear (\ x y. A x y - B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof - have \bounded_sesquilinear (\ x y. - B x y)\ using that(2) by (rule bounded_sesquilinear_uminus) then have \bounded_sesquilinear (\ x y. A x y + (- B x y))\ using that(1) by (rule bounded_sesquilinear_add[rotated]) then show ?thesis by auto qed lemma ccsubspace_leI: assumes t1: "space_as_set A \ space_as_set B" shows "A \ B" using t1 apply transfer by - lemma ccspan_of_empty[simp]: "ccspan {} = bot" proof transfer show "closure (cspan {}) = {0::'a}" by simp qed instantiation ccsubspace :: ("{complex_vector,topological_space}") inf begin lift_definition inf_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" is "(\)" by simp instance .. end lemma space_as_set_inf[simp]: "space_as_set (A \ B) = space_as_set A \ space_as_set B" by (rule inf_ccsubspace.rep_eq) instantiation ccsubspace :: ("{complex_vector,topological_space}") order_top begin instance proof show "a \ \" for a :: "'a ccsubspace" apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") order_bot begin instance proof show "(\::'a ccsubspace) \ a" for a :: "'a ccsubspace" apply transfer apply auto using closed_csubspace.subspace complex_vector.subspace_0 by blast qed end instantiation ccsubspace :: ("{complex_vector,topological_space}") semilattice_inf begin instance proof fix x y z :: \'a ccsubspace\ show "x \ y \ x" apply transfer by simp show "x \ y \ y" apply transfer by simp show "x \ y \ z" if "x \ y" and "x \ z" using that apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") zero begin definition zero_ccsubspace :: "'a ccsubspace" where [simp]: "zero_ccsubspace = bot" lemma zero_ccsubspace_transfer[transfer_rule]: \pcr_ccsubspace (=) {0} 0\ unfolding zero_ccsubspace_def by transfer_prover instance .. end subsection \Closed sums\ definition closed_sum:: \'a::{semigroup_add,topological_space} set \ 'a set \ 'a set\ where \closed_sum A B = closure (A + B)\ notation closed_sum (infixl "+\<^sub>M" 65) lemma closed_sum_comm: \A +\<^sub>M B = B +\<^sub>M A\ for A B :: "_::ab_semigroup_add" by (simp add: add.commute closed_sum_def) lemma closed_sum_left_subset: \0 \ B \ A \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.right_neutral closed_sum_def closure_subset in_mono set_plus_intro subsetI) lemma closed_sum_right_subset: \0 \ A \ B \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.left_neutral closed_sum_def closure_subset set_plus_intro subset_iff) lemma finite_cspan_closed_csubspace: assumes "finite (S::'a::complex_normed_vector set)" shows "closed_csubspace (cspan S)" by (simp add: assms closed_csubspace.intro) lemma closed_sum_is_sup: fixes A B C:: \('a::{complex_vector,topological_space}) set\ assumes \closed_csubspace C\ assumes \A \ C\ and \B \ C\ shows \(A +\<^sub>M B) \ C\ proof - have \A + B \ C\ using assms unfolding set_plus_def using closed_csubspace.subspace complex_vector.subspace_add by blast then show \(A +\<^sub>M B) \ C\ unfolding closed_sum_def using \closed_csubspace C\ by (simp add: closed_csubspace.closed closure_minimal) qed lemma closed_subspace_closed_sum: fixes A B::"('a::complex_normed_vector) set" assumes a1: \csubspace A\ and a2: \csubspace B\ shows \closed_csubspace (A +\<^sub>M B)\ using a1 a2 closed_sum_def by (metis closure_is_closed_csubspace csubspace_set_plus) lemma closed_sum_assoc: fixes A B C::"'a::real_normed_vector set" shows \A +\<^sub>M (B +\<^sub>M C) = (A +\<^sub>M B) +\<^sub>M C\ proof - have \A + closure B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \A +\<^sub>M (B +\<^sub>M C) = closure (A + (B + C))\ unfolding closed_sum_def by (meson antisym_conv closed_closure closure_minimal closure_mono closure_subset equalityD1 set_plus_mono2) moreover have \closure A + B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \(A +\<^sub>M B) +\<^sub>M C = closure ((A + B) + C)\ unfolding closed_sum_def by (meson closed_closure closure_minimal closure_mono closure_subset eq_iff set_plus_mono2) ultimately show ?thesis by (simp add: ab_semigroup_add_class.add_ac(1)) qed lemma closed_sum_zero_left[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \{0} +\<^sub>M A = closure A\ unfolding closed_sum_def by (metis add.left_neutral set_zero) lemma closed_sum_zero_right[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \A +\<^sub>M {0} = closure A\ unfolding closed_sum_def by (metis add.right_neutral set_zero) lemma closed_sum_closure_right[simp]: fixes A B :: \'a::real_normed_vector set\ shows \A +\<^sub>M closure B = A +\<^sub>M B\ by (metis closed_sum_assoc closed_sum_def closed_sum_zero_right closure_closure) lemma closed_sum_closure_left[simp]: fixes A B :: \'a::real_normed_vector set\ shows \closure A +\<^sub>M B = A +\<^sub>M B\ by (simp add: closed_sum_comm) lemma closed_sum_mono_left: assumes \A \ B\ shows \A +\<^sub>M C \ B +\<^sub>M C\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) lemma closed_sum_mono_right: assumes \A \ B\ shows \C +\<^sub>M A \ C +\<^sub>M B\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) instantiation ccsubspace :: (complex_normed_vector) sup begin lift_definition sup_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" \ \Note that \<^term>\A+B\ would not be a closed subspace, we need the closure. See, e.g., \<^url>\https://math.stackexchange.com/a/1786792/403528\.\ is "\A B::'a set. A +\<^sub>M B" by (simp add: closed_subspace_closed_sum) instance .. end lemma closed_sum_cspan[simp]: shows \cspan X +\<^sub>M cspan Y = closure (cspan (X \ Y))\ by (smt (verit, best) Collect_cong closed_sum_def complex_vector.span_Un set_plus_def) lemma closure_image_closed_sum: assumes \bounded_linear U\ shows \closure (U ` (A +\<^sub>M B)) = closure (U ` A) +\<^sub>M closure (U ` B)\ proof - have \closure (U ` (A +\<^sub>M B)) = closure (U ` closure (closure A + closure B))\ unfolding closed_sum_def by (smt (verit, best) closed_closure closure_minimal closure_mono closure_subset closure_sum set_plus_mono2 subset_antisym) also have \\ = closure (U ` (closure A + closure B))\ using assms closure_bounded_linear_image_subset_eq by blast also have \\ = closure (U ` closure A + U ` closure B)\ apply (subst image_set_plus) by (simp_all add: assms bounded_linear.linear) also have \\ = closure (closure (U ` A) + closure (U ` B))\ by (smt (verit, ccfv_SIG) assms closed_closure closure_bounded_linear_image_subset closure_bounded_linear_image_subset_eq closure_minimal closure_mono closure_sum dual_order.eq_iff set_plus_mono2) also have \\ = closure (U ` A) +\<^sub>M closure (U ` B)\ using closed_sum_def by blast finally show ?thesis by - qed lemma ccspan_union: "ccspan A \ ccspan B = ccspan (A \ B)" apply transfer by simp instantiation ccsubspace :: (complex_normed_vector) "Sup" begin lift_definition Sup_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\S. closure (complex_vector.span (Union S))\ proof show "csubspace (closure (complex_vector.span (\ S::'a set)))" if "\x::'a set. x \ S \ closed_csubspace x" for S :: "'a set set" using that by (simp add: closure_is_closed_csubspace) show "closed (closure (complex_vector.span (\ S::'a set)))" if "\x. (x::'a set) \ S \ closed_csubspace x" for S :: "'a set set" using that by simp qed instance.. end instance ccsubspace :: ("{complex_normed_vector}") semilattice_sup proof fix x y z :: \'a ccsubspace\ show \x \ sup x y\ apply transfer by (simp add: closed_csubspace_def closed_sum_left_subset complex_vector.subspace_0) show "y \ sup x y" apply transfer by (simp add: closed_csubspace_def closed_sum_right_subset complex_vector.subspace_0) show "sup x y \ z" if "x \ z" and "y \ z" using that apply transfer apply (rule closed_sum_is_sup) by auto qed instance ccsubspace :: ("{complex_normed_vector}") complete_lattice proof show "Inf A \ x" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by auto have b1: "z \ \ A" if "Ball A closed_csubspace" and "closed_csubspace z" and "(\x. closed_csubspace x \ x \ A \ z \ x)" for z::"'a set" and A using that by auto show "z \ Inf A" if "\x::'a ccsubspace. x \ A \ z \ x" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer using b1 by blast show "x \ Sup A" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by (meson Union_upper closure_subset complex_vector.span_superset dual_order.trans) show "Sup A \ z" if "\x::'a ccsubspace. x \ A \ x \ z" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer proof - fix A :: "'a set set" and z :: "'a set" assume A_closed: "Ball A closed_csubspace" assume "closed_csubspace z" assume in_z: "\x. closed_csubspace x \ x \ A \ x \ z" from A_closed in_z have \V \ z\ if \V \ A\ for V by (simp add: that) then have \\ A \ z\ by (simp add: Sup_le_iff) with \closed_csubspace z\ show "closure (cspan (\ A)) \ z" by (simp add: closed_csubspace_def closure_minimal complex_vector.span_def subset_hull) qed show "Inf {} = (top::'a ccsubspace)" using \\z A. (\x. x \ A \ z \ x) \ z \ Inf A\ top.extremum_uniqueI by auto show "Sup {} = (bot::'a ccsubspace)" using \\z A. (\x. x \ A \ x \ z) \ Sup A \ z\ bot.extremum_uniqueI by auto qed instantiation ccsubspace :: (complex_normed_vector) comm_monoid_add begin definition plus_ccsubspace :: "'a ccsubspace \ _ \ _" where [simp]: "plus_ccsubspace = sup" instance proof fix a b c :: \'a ccsubspace\ show "a + b + c = a + (b + c)" using sup.assoc by auto show "a + b = b + a" by (simp add: sup.commute) show "0 + a = a" by (simp add: zero_ccsubspace_def) qed end lemma ccsubspace_plus_sup: "y \ x \ z \ x \ y + z \ x" for x y z :: "'a::complex_normed_vector ccsubspace" unfolding plus_ccsubspace_def by auto lemma ccsubspace_Sup_empty: "Sup {} = (0::_ ccsubspace)" unfolding zero_ccsubspace_def by auto lemma ccsubspace_add_right_incr[simp]: "a \ a + c" for a::"_ ccsubspace" by (simp add: add_increasing2) lemma ccsubspace_add_left_incr[simp]: "a \ c + a" for a::"_ ccsubspace" by (simp add: add_increasing) subsection \Conjugate space\ typedef 'a conjugate_space = "UNIV :: 'a set" morphisms from_conjugate_space to_conjugate_space .. setup_lifting type_definition_conjugate_space instantiation conjugate_space :: (complex_vector) complex_vector begin lift_definition scaleC_conjugate_space :: \complex \ 'a conjugate_space \ 'a conjugate_space\ is \\c x. cnj c *\<^sub>C x\. lift_definition scaleR_conjugate_space :: \real \ 'a conjugate_space \ 'a conjugate_space\ is \\r x. r *\<^sub>R x\. lift_definition plus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(+)". lift_definition uminus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is \\x. -x\. lift_definition zero_conjugate_space :: "'a conjugate_space" is 0. lift_definition minus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(-)". instance apply (intro_classes; transfer) by (simp_all add: scaleR_scaleC scaleC_add_right scaleC_left.add) end instantiation conjugate_space :: (complex_normed_vector) complex_normed_vector begin lift_definition sgn_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is "sgn". lift_definition norm_conjugate_space :: "'a conjugate_space \ real" is norm. lift_definition dist_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ real" is dist. lift_definition uniformity_conjugate_space :: "('a conjugate_space \ 'a conjugate_space) filter" is uniformity. lift_definition open_conjugate_space :: "'a conjugate_space set \ bool" is "open". instance apply (intro_classes; transfer) by (simp_all add: dist_norm sgn_div_norm open_uniformity uniformity_dist norm_triangle_ineq) end instantiation conjugate_space :: (cbanach) cbanach begin instance apply intro_classes unfolding Cauchy_def convergent_def LIMSEQ_def apply transfer using Cauchy_convergent unfolding Cauchy_def convergent_def LIMSEQ_def by metis end lemma bounded_antilinear_to_conjugate_space[simp]: \bounded_antilinear to_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma bounded_antilinear_from_conjugate_space[simp]: \bounded_antilinear from_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma antilinear_to_conjugate_space[simp]: \antilinear to_conjugate_space\ by (rule antilinearI; transfer, auto) lemma antilinear_from_conjugate_space[simp]: \antilinear from_conjugate_space\ by (rule antilinearI; transfer, auto) lemma cspan_to_conjugate_space[simp]: "cspan (to_conjugate_space ` X) = to_conjugate_space ` cspan X" unfolding complex_vector.span_def complex_vector.subspace_def hull_def apply transfer apply simp by (metis (no_types, opaque_lifting) complex_cnj_cnj) lemma surj_to_conjugate_space[simp]: "surj to_conjugate_space" by (meson surj_def to_conjugate_space_cases) lemmas has_derivative_scaleC[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] lemma norm_to_conjugate_space[simp]: \norm (to_conjugate_space x) = norm x\ by (fact norm_conjugate_space.abs_eq) lemma norm_from_conjugate_space[simp]: \norm (from_conjugate_space x) = norm x\ by (simp add: norm_conjugate_space.rep_eq) lemma closure_to_conjugate_space: \closure (to_conjugate_space ` X) = to_conjugate_space ` closure X\ proof - have 1: \to_conjugate_space ` closure X \ closure (to_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = to_conjugate_space ` from_conjugate_space ` closure (to_conjugate_space ` X)\ by (simp add: from_conjugate_space_inverse image_image) also have \\ \ to_conjugate_space ` closure (from_conjugate_space ` to_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = to_conjugate_space ` closure X\ by (simp add: to_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma closure_from_conjugate_space: \closure (from_conjugate_space ` X) = from_conjugate_space ` closure X\ proof - have 1: \from_conjugate_space ` closure X \ closure (from_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = from_conjugate_space ` to_conjugate_space ` closure (from_conjugate_space ` X)\ by (simp add: to_conjugate_space_inverse image_image) also have \\ \ from_conjugate_space ` closure (to_conjugate_space ` from_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = from_conjugate_space ` closure X\ by (simp add: from_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma bounded_antilinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_antilinear A\ and \bounded_antilinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - let ?A = \\x. A (from_conjugate_space x)\ and ?B = \\x. B (from_conjugate_space x)\ and ?G = \to_conjugate_space ` G\ and ?t = \to_conjugate_space t\ have \bounded_clinear ?A\ and \bounded_clinear ?B\ by (auto intro!: bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear A\] bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear B\]) moreover from eq have \\x. x \ ?G \ ?A x = ?B x\ by (metis image_iff iso_tuple_UNIV_I to_conjugate_space_inverse) moreover from t have \?t \ closure (cspan ?G)\ by (metis bounded_antilinear.bounded_linear bounded_antilinear_to_conjugate_space closure_bounded_linear_image_subset cspan_to_conjugate_space imageI subsetD) ultimately have \?A ?t = ?B ?t\ by (rule bounded_clinear_eq_on) then show \A t = B t\ by (simp add: to_conjugate_space_inverse) qed instantiation complex :: basis_enum begin definition "canonical_basis = [1::complex]" instance proof show "distinct (canonical_basis::complex list)" by (simp add: canonical_basis_complex_def) show "cindependent (set (canonical_basis::complex list))" unfolding canonical_basis_complex_def by auto show "cspan (set (canonical_basis::complex list)) = UNIV" unfolding canonical_basis_complex_def apply (auto simp add: cspan_raw_def vector_space_over_itself.span_Basis) by (metis complex_scaleC_def complex_vector.span_base complex_vector.span_scale cspan_raw_def insertI1 mult.right_neutral) qed end lemma csubspace_is_convex[simp]: assumes a1: "csubspace M" shows "convex M" proof- have \\x\M. \y\ M. \u. \v. u *\<^sub>C x + v *\<^sub>C y \ M\ using a1 by (simp add: complex_vector.subspace_def) hence \\x\M. \y\M. \u::real. \v::real. u *\<^sub>R x + v *\<^sub>R y \ M\ by (simp add: scaleR_scaleC) hence \\x\M. \y\M. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \M\ by blast thus ?thesis using convex_def by blast qed lemma kernel_is_csubspace[simp]: assumes a1: "clinear f" shows "csubspace (f -` {0})" proof- have w3: \t *\<^sub>C x \ {x. f x = 0}\ if b1: "x \ {x. f x = 0}" for x t by (metis assms complex_vector.linear_subspace_kernel complex_vector.subspace_def that) have \f 0 = 0\ by (simp add: assms complex_vector.linear_0) hence s2: \0 \ {x. f x = 0}\ by blast have w4: "x + y \ {x. f x = 0}" if c1: "x \ {x. f x = 0}" and c2: "y \ {x. f x = 0}" for x y using assms c1 c2 complex_vector.linear_add by fastforce have s4: \c *\<^sub>C t \ {x. f x = 0}\ if "t \ {x. f x = 0}" for t c using that w3 by auto have s5: "u + v \ {x. f x = 0}" if "u \ {x. f x = 0}" and "v \ {x. f x = 0}" for u v using w4 that(1) that(2) by auto have f3: "f -` {b. b = 0 \ b \ {}} = {a. f a = 0}" by blast have "csubspace {a. f a = 0}" by (metis complex_vector.subspace_def s2 s4 s5) thus ?thesis using f3 by auto qed lemma kernel_is_closed_csubspace[simp]: assumes a1: "bounded_clinear f" shows "closed_csubspace (f -` {0})" proof- have \csubspace (f -` {0})\ using assms bounded_clinear.clinear complex_vector.linear_subspace_vimage complex_vector.subspace_single_0 by blast have "L \ {x. f x = 0}" if "r \ L" and "\ n. r n \ {x. f x = 0}" for r and L proof- have d1: \\ n. f (r n) = 0\ using that(2) by auto have \(\ n. f (r n)) \ f L\ using assms clinear_continuous_at continuous_within_tendsto_compose' that(1) by fastforce hence \(\ n. 0) \ f L\ using d1 by simp hence \f L = 0\ using limI by fastforce thus ?thesis by blast qed then have s3: \closed (f -` {0})\ using closed_sequential_limits by force with \csubspace (f -` {0})\ show ?thesis using closed_csubspace.intro by blast qed lemma range_is_clinear[simp]: assumes a1: "clinear f" shows "csubspace (range f)" using assms complex_vector.linear_subspace_image complex_vector.subspace_UNIV by blast lemma ccspan_superset: \A \ space_as_set (ccspan A)\ for A :: \'a::complex_normed_vector set\ apply transfer by (meson closure_subset complex_vector.span_superset subset_trans) subsection \Product is a Complex Vector Space\ (* Follows closely Product_Vector.thy *) instantiation prod :: (complex_vector, complex_vector) complex_vector begin definition scaleC_prod_def: "scaleC r A = (scaleC r (fst A), scaleC r (snd A))" lemma fst_scaleC [simp]: "fst (scaleC r A) = scaleC r (fst A)" unfolding scaleC_prod_def by simp lemma snd_scaleC [simp]: "snd (scaleC r A) = scaleC r (snd A)" unfolding scaleC_prod_def by simp proposition scaleC_Pair [simp]: "scaleC r (a, b) = (scaleC r a, scaleC r b)" unfolding scaleC_prod_def by simp instance proof fix a b :: complex and x y :: "'a \ 'b" show "scaleC a (x + y) = scaleC a x + scaleC a y" by (simp add: scaleC_add_right scaleC_prod_def) show "scaleC (a + b) x = scaleC a x + scaleC b x" by (simp add: Complex_Vector_Spaces.scaleC_prod_def scaleC_left.add) show "scaleC a (scaleC b x) = scaleC (a * b) x" by (simp add: prod_eq_iff) show "scaleC 1 x = x" by (simp add: prod_eq_iff) show \(scaleR :: _ \ _ \ 'a*'b) r = (*\<^sub>C) (complex_of_real r)\ for r by (auto intro!: ext simp: scaleR_scaleC scaleC_prod_def scaleR_prod_def) qed end lemma module_prod_scale_eq_scaleC: "module_prod.scale (*\<^sub>C) (*\<^sub>C) = scaleC" apply (rule ext) apply (rule ext) apply (subst module_prod.scale_def) subgoal by unfold_locales by (simp add: scaleC_prod_def) interpretation complex_vector?: vector_space_prod "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" rewrites "scale = ((*\<^sub>C)::_\_\('a \ 'b))" and "module.dependent (*\<^sub>C) = cdependent" and "module.representation (*\<^sub>C) = crepresentation" and "module.subspace (*\<^sub>C) = csubspace" and "module.span (*\<^sub>C) = cspan" and "vector_space.extend_basis (*\<^sub>C) = cextend_basis" and "vector_space.dim (*\<^sub>C) = cdim" and "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" subgoal by unfold_locales subgoal by (fact module_prod_scale_eq_scaleC) unfolding cdependent_raw_def crepresentation_raw_def csubspace_raw_def cspan_raw_def cextend_basis_raw_def cdim_raw_def clinear_def by (rule refl)+ subsection \Copying existing theorems into sublocales\ context bounded_clinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = continuous lemmas uniform_limit = uniform_limit lemmas Cauchy = Cauchy end context bounded_antilinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = continuous lemmas uniform_limit = uniform_limit end context bounded_cbilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = tendsto lemmas isCont = isCont end context bounded_sesquilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = tendsto lemmas isCont = isCont end lemmas tendsto_scaleC [tendsto_intros] = bounded_cbilinear.tendsto [OF bounded_cbilinear_scaleC] +unbundle no_lattice_syntax + end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy b/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy @@ -1,271 +1,249 @@ section \\Extra_Lattice\ -- Additional results about lattices\ theory Extra_Lattice imports Main begin subsection\\Lattice_Missing\ -- Miscellaneous missing facts about lattices\ -text \Two bundles to activate and deactivate lattice specific notation (e.g., \\\ etc.). - Activate the notation locally via "@{theory_text \includes lattice_notation\}" in a lemma statement. - (Or sandwich a declaration using that notation between "@{theory_text \unbundle lattice_notation ... unbundle no_lattice_notation\}.)\ - -bundle lattice_notation begin -notation inf (infixl "\" 70) -notation sup (infixl "\" 65) -notation Inf ("\") -notation Sup ("\") -notation bot ("\") -notation top ("\") -end - -bundle no_lattice_notation begin -notation inf (infixl "\" 70) -notation sup (infixl "\" 65) -notation Inf ("\") -notation Sup ("\") -notation bot ("\") -notation top ("\") -end - -unbundle lattice_notation +unbundle lattice_syntax text \The following class \complemented_lattice\ describes complemented lattices (with \<^const>\uminus\ for the complement). The definition follows \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Definition_and_basic_properties\. Additionally, it adopts the convention from \<^class>\boolean_algebra\ of defining \<^const>\minus\ in terms of the complement.\ class complemented_lattice = bounded_lattice + uminus + minus + assumes inf_compl_bot[simp]: "inf x (-x) = bot" and sup_compl_top[simp]: "sup x (-x) = top" and diff_eq: "x - y = inf x (- y)" begin lemma dual_complemented_lattice: "class.complemented_lattice (\x y. x \ (- y)) uminus sup greater_eq greater inf \ \" proof (rule class.complemented_lattice.intro) show "class.bounded_lattice (\) (\x y. (y::'a) \ x) (\x y. y < x) (\) \ \" by (rule dual_bounded_lattice) show "class.complemented_lattice_axioms (\x y. (x::'a) \ - y) uminus (\) (\) \ \" by (unfold_locales, auto simp add: diff_eq) qed lemma compl_inf_bot [simp]: "inf (- x) x = bot" by (simp add: inf_commute) lemma compl_sup_top [simp]: "sup (- x) x = top" by (simp add: sup_commute) end class complete_complemented_lattice = complemented_lattice + complete_lattice text \The following class \complemented_lattice\ describes orthocomplemented lattices, following \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Orthocomplementation\.\ class orthocomplemented_lattice = complemented_lattice + assumes ortho_involution[simp]: "- (- x) = x" and ortho_antimono: "x \ y \ -x \ -y" begin lemma dual_orthocomplemented_lattice: "class.orthocomplemented_lattice (\x y. x \ - y) uminus sup greater_eq greater inf \ \" proof (rule class.orthocomplemented_lattice.intro) show "class.complemented_lattice (\x y. (x::'a) \ - y) uminus (\) (\x y. y \ x) (\x y. y < x) (\) \ \" by (rule dual_complemented_lattice) show "class.orthocomplemented_lattice_axioms uminus (\x y. (y::'a) \ x)" by (unfold_locales, auto simp add: diff_eq intro: ortho_antimono) qed lemma compl_eq_compl_iff [simp]: "- x = - y \ x = y" by (metis ortho_involution) lemma compl_bot_eq [simp]: "- bot = top" by (metis inf_compl_bot inf_top_left ortho_involution) lemma compl_top_eq [simp]: "- top = bot" using compl_bot_eq ortho_involution by blast text \De Morgan's law\ (* Proof from: https://planetmath.org/orthocomplementedlattice *) lemma compl_sup [simp]: "- (x \ y) = - x \ - y" proof - have "- (x \ y) \ - x" by (simp add: ortho_antimono) moreover have "- (x \ y) \ - y" by (simp add: ortho_antimono) ultimately have 1: "- (x \ y) \ - x \ - y" by (simp add: sup.coboundedI1) have \x \ - (-x \ -y)\ by (metis inf.cobounded1 ortho_antimono ortho_involution) moreover have \y \ - (-x \ -y)\ by (metis inf.cobounded2 ortho_antimono ortho_involution) ultimately have \x \ y \ - (-x \ -y)\ by auto hence 2: \-x \ -y \ - (x \ y)\ using ortho_antimono by fastforce from 1 2 show ?thesis using dual_order.antisym by blast qed text \De Morgan's law\ lemma compl_inf [simp]: "- (x \ y) = - x \ - y" using compl_sup by (metis ortho_involution) lemma compl_mono: assumes "x \ y" shows "- y \ - x" by (simp add: assms local.ortho_antimono) lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" by (auto dest: compl_mono) lemma compl_le_swap1: assumes "y \ - x" shows "x \ -y" using assms ortho_antimono by fastforce lemma compl_le_swap2: assumes "- y \ x" shows "- x \ y" using assms local.ortho_antimono by fastforce lemma compl_less_compl_iff[simp]: "- x < - y \ y < x" by (auto simp add: less_le) lemma compl_less_swap1: assumes "y < - x" shows "x < - y" using assms compl_less_compl_iff by fastforce lemma compl_less_swap2: assumes "- y < x" shows "- x < y" using assms compl_le_swap1 compl_le_swap2 less_le_not_le by auto lemma sup_cancel_left1: "sup (sup x a) (sup (- x) b) = top" by (simp add: sup_commute sup_left_commute) lemma sup_cancel_left2: "sup (sup (- x) a) (sup x b) = top" by (simp add: sup.commute sup_left_commute) lemma inf_cancel_left1: "inf (inf x a) (inf (- x) b) = bot" by (simp add: inf.left_commute inf_commute) lemma inf_cancel_left2: "inf (inf (- x) a) (inf x b) = bot" using inf.left_commute inf_commute by auto lemma sup_compl_top_left1 [simp]: "sup (- x) (sup x y) = top" by (simp add: sup_assoc[symmetric]) lemma sup_compl_top_left2 [simp]: "sup x (sup (- x) y) = top" using sup_compl_top_left1[of "- x" y] by simp lemma inf_compl_bot_left1 [simp]: "inf (- x) (inf x y) = bot" by (simp add: inf_assoc[symmetric]) lemma inf_compl_bot_left2 [simp]: "inf x (inf (- x) y) = bot" using inf_compl_bot_left1[of "- x" y] by simp lemma inf_compl_bot_right [simp]: "inf x (inf y (- x)) = bot" by (subst inf_left_commute) simp end class complete_orthocomplemented_lattice = orthocomplemented_lattice + complete_lattice instance complete_orthocomplemented_lattice \ complete_complemented_lattice by intro_classes text \The following class \orthomodular_lattice\ describes orthomodular lattices, following \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Orthomodular_lattices\.\ class orthomodular_lattice = orthocomplemented_lattice + assumes orthomodular: "x \ y \ sup x (inf (-x) y) = y" begin lemma dual_orthomodular_lattice: "class.orthomodular_lattice (\x y. x \ - y) uminus sup greater_eq greater inf \ \" proof (rule class.orthomodular_lattice.intro) show "class.orthocomplemented_lattice (\x y. (x::'a) \ - y) uminus (\) (\x y. y \ x) (\x y. y < x) (\) \ \" by (rule dual_orthocomplemented_lattice) show "class.orthomodular_lattice_axioms uminus (\) (\x y. (y::'a) \ x) (\)" proof (unfold_locales) show "(x::'a) \ (- x \ y) = y" if "(y::'a) \ x" for x :: 'a and y :: 'a using that local.compl_eq_compl_iff local.ortho_antimono local.orthomodular by fastforce qed qed end class complete_orthomodular_lattice = orthomodular_lattice + complete_lattice begin end instance complete_orthomodular_lattice \ complete_orthocomplemented_lattice by intro_classes instance boolean_algebra \ orthomodular_lattice proof fix x y :: 'a show "sup (x::'a) (inf (- x) y) = y" if "(x::'a) \ y" using that by (simp add: sup.absorb_iff2 sup_inf_distrib1) show "x - y = inf x (- y)" by (simp add: boolean_algebra_class.diff_eq) qed auto instance complete_boolean_algebra \ complete_orthomodular_lattice by intro_classes lemma image_of_maximum: fixes f::"'a::order \ 'b::conditionally_complete_lattice" assumes "mono f" and "\x. x:M \ x\m" and "m:M" shows "(SUP x\M. f x) = f m" by (smt (verit, ccfv_threshold) assms(1) assms(2) assms(3) cSup_eq_maximum imageE imageI monoD) lemma cSup_eq_cSup: fixes A B :: \'a::conditionally_complete_lattice set\ assumes bdd: \bdd_above A\ assumes B: \\a. a\A \ \b\B. b \ a\ assumes A: \\b. b\B \ \a\A. a \ b\ shows \Sup A = Sup B\ proof (cases \B = {}\) case True with A B have \A = {}\ by auto with True show ?thesis by simp next case False have \bdd_above B\ by (meson A bdd bdd_above_def order_trans) have \A \ {}\ using A False by blast moreover have \a \ Sup B\ if \a \ A\ for a proof - obtain b where \b \ B\ and \b \ a\ using B \a \ A\ by auto then show ?thesis apply (rule cSup_upper2) using \bdd_above B\ by simp qed moreover have \Sup B \ c\ if \\a. a \ A \ a \ c\ for c using False apply (rule cSup_least) using A that by fastforce ultimately show ?thesis by (rule cSup_eq_non_empty) qed -unbundle no_lattice_notation +unbundle no_lattice_syntax end diff --git a/thys/Registers/Quantum_Extra.thy b/thys/Registers/Quantum_Extra.thy --- a/thys/Registers/Quantum_Extra.thy +++ b/thys/Registers/Quantum_Extra.thy @@ -1,162 +1,163 @@ section \Derived facts about quantum registers\ theory Quantum_Extra imports Laws_Quantum Quantum begin no_notation meet (infixl "\\" 70) no_notation Group.mult (infixl "\\" 70) no_notation Order.top ("\\") +unbundle lattice_syntax unbundle register_notation unbundle cblinfun_notation lemma zero_not_register[simp]: \~ register (\_. 0)\ unfolding register_def by simp lemma register_pair_is_register_converse: \register (F;G) \ register F\ \register (F;G) \ register G\ using [[simproc del: Laws_Quantum.compatibility_warn]] apply (cases \register F\) apply (auto simp: register_pair_def)[2] apply (cases \register G\) by (auto simp: register_pair_def)[2] lemma register_id'[simp]: \register (\x. x)\ using register_id by (simp add: id_def) lemma register_projector: assumes "register F" assumes "is_Proj a" shows "is_Proj (F a)" using assms unfolding register_def is_Proj_algebraic by metis lemma register_unitary: assumes "register F" assumes "unitary a" shows "unitary (F a)" using assms by (smt (verit, best) register_def unitary_def) lemma compatible_proj_intersect: (* I think this also holds without is_Proj premises, but my proof ideas use the Penrose-Moore pseudoinverse or simultaneous diagonalization and we do not have an existence theorem for either. *) assumes "compatible R S" and "is_Proj a" and "is_Proj b" shows "(R a *\<^sub>S \) \ (S b *\<^sub>S \) = ((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \)" proof (rule antisym) have "((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \) \ (S b *\<^sub>S \)" apply (subst swap_registers[OF assms(1)]) by (simp add: cblinfun_compose_image cblinfun_image_mono) moreover have "((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \) \ (R a *\<^sub>S \)" by (simp add: cblinfun_compose_image cblinfun_image_mono) ultimately show \((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \) \ (R a *\<^sub>S \) \ (S b *\<^sub>S \)\ by auto have "is_Proj (R a)" using assms(1) assms(2) compatible_register1 register_projector by blast have "is_Proj (S b)" using assms(1) assms(3) compatible_register2 register_projector by blast show \(R a *\<^sub>S \) \ (S b *\<^sub>S \) \ (R a o\<^sub>C\<^sub>L S b) *\<^sub>S \\ proof (unfold less_eq_ccsubspace.rep_eq, rule) fix \ assume asm: \\ \ space_as_set ((R a *\<^sub>S \) \ (S b *\<^sub>S \))\ then have \\ \ space_as_set (R a *\<^sub>S \)\ by auto then have R: \R a *\<^sub>V \ = \\ using \is_Proj (R a)\ cblinfun_fixes_range is_Proj_algebraic by blast from asm have \\ \ space_as_set (S b *\<^sub>S \)\ by auto then have S: \S b *\<^sub>V \ = \\ using \is_Proj (S b)\ cblinfun_fixes_range is_Proj_algebraic by blast from R S have \\ = (R a o\<^sub>C\<^sub>L S b) *\<^sub>V \\ by (simp add: cblinfun_apply_cblinfun_compose) also have \\ \ space_as_set ((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \)\ apply simp by (metis R S calculation cblinfun_apply_in_image) finally show \\ \ space_as_set ((R a o\<^sub>C\<^sub>L S b) *\<^sub>S \)\ by - qed qed lemma compatible_proj_mult: assumes "compatible R S" and "is_Proj a" and "is_Proj b" shows "is_Proj (R a o\<^sub>C\<^sub>L S b)" using [[simproc del: Laws_Quantum.compatibility_warn]] using assms unfolding is_Proj_algebraic compatible_def apply auto apply (metis (no_types, lifting) cblinfun_compose_assoc register_mult) by (simp add: assms(2) assms(3) is_proj_selfadj register_projector) lemma unitary_sandwich_register: \unitary a \ register (sandwich a)\ unfolding register_def apply (auto simp: sandwich_def) apply (metis (no_types, lifting) cblinfun_assoc_left(1) cblinfun_compose_id_right unitaryD1) by (simp add: lift_cblinfun_comp(2)) lemma sandwich_tensor: fixes a :: \'a::finite ell2 \\<^sub>C\<^sub>L 'a ell2\ and b :: \'b::finite ell2 \\<^sub>C\<^sub>L 'b ell2\ assumes \unitary a\ \unitary b\ shows "sandwich (a \\<^sub>o b) = sandwich a \\<^sub>r sandwich b" apply (rule tensor_extensionality) by (auto simp: unitary_sandwich_register assms sandwich_def register_tensor_is_register comp_tensor_op tensor_op_adjoint) lemma sandwich_grow_left: fixes a :: \'a::finite ell2 \\<^sub>C\<^sub>L 'a ell2\ assumes "unitary a" shows "sandwich a \\<^sub>r id = sandwich (a \\<^sub>o id_cblinfun)" by (simp add: unitary_sandwich_register assms sandwich_tensor sandwich_id) lemma register_sandwich: \register F \ F (sandwich a b) = sandwich (F a) (F b)\ by (smt (verit, del_insts) register_def sandwich_def) lemma assoc_ell2_sandwich: \assoc = sandwich assoc_ell2\ apply (rule tensor_extensionality3') apply (simp_all add: unitary_sandwich_register)[2] apply (rule equal_ket) apply (case_tac x) by (simp add: sandwich_def assoc_apply cblinfun_apply_cblinfun_compose tensor_op_ell2 assoc_ell2_tensor assoc_ell2'_tensor flip: tensor_ell2_ket) lemma assoc_ell2'_sandwich: \assoc' = sandwich assoc_ell2'\ apply (rule tensor_extensionality3) apply (simp_all add: unitary_sandwich_register)[2] apply (rule equal_ket) apply (case_tac x) by (simp add: sandwich_def assoc'_apply cblinfun_apply_cblinfun_compose tensor_op_ell2 assoc_ell2_tensor assoc_ell2'_tensor flip: tensor_ell2_ket) lemma swap_sandwich: "swap = sandwich Uswap" apply (rule tensor_extensionality) apply (auto simp: sandwich_def)[2] apply (rule tensor_ell2_extensionality) by (simp add: sandwich_def cblinfun_apply_cblinfun_compose tensor_op_ell2) lemma id_tensor_sandwich: fixes a :: "'a::finite ell2 \\<^sub>C\<^sub>L 'b::finite ell2" assumes "unitary a" shows "id \\<^sub>r sandwich a = sandwich (id_cblinfun \\<^sub>o a)" apply (rule tensor_extensionality) using assms by (auto simp: register_tensor_is_register comp_tensor_op sandwich_def tensor_op_adjoint unitary_sandwich_register) lemma compatible_selfbutter_join: assumes [register]: "compatible R S" shows "R (selfbutter \) o\<^sub>C\<^sub>L S (selfbutter \) = (R; S) (selfbutter (\ \\<^sub>s \))" apply (subst register_pair_apply[symmetric, where F=R and G=S]) using assms by auto lemma register_mult': assumes \register F\ shows \F a *\<^sub>V F b *\<^sub>V c = F (a o\<^sub>C\<^sub>L b) *\<^sub>V c\ by (simp add: assms lift_cblinfun_comp(4) register_mult) lemma register_scaleC: assumes \register F\ shows \F (c *\<^sub>C a) = c *\<^sub>C F a\ by (simp add: assms complex_vector.linear_scale) lemma register_bounded_clinear: \register F \ bounded_clinear F\ using bounded_clinear_finite_dim register_def by blast lemma register_adjoint: "F (a*) = (F a)*" if \register F\ using register_def that by blast end