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,666 +1,670 @@ 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 @{thm [source] mat_of_cblinfun_plus}.) See \<^cite>\"code-generation-tutorial"\ for more information on @{theory_text \[code abstype]\}.\ declare mat_of_cblinfun_inverse [code abstype] 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_explicit_cblinfun[code] + \ \Code equation for computing the \<^const>\explicit_cblinfun\\ + 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_carrier_vec[simp]: \vec_of_ell2 v \ carrier_vec CARD('a)\ for v :: \'a::enum ell2\ + using vec_of_basis_enum_carrier_vec[of v] + apply (simp only: canonical_basis_length canonical_basis_length_ell2) by (simp add: vec_of_ell2_def) 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_scaleC[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 \ = norm_vec (vec_of_ell2 \)" - by (simp add: norm_ell2_vec_of_basis_enum vec_of_ell2_def) + by (simp add: norm_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 additional 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_ell2[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_ell2\ 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_ell2\ (where possible), and we add code generation for \cblinfun_apply_ell2\ instead of \<^term>\(*\<^sub>V)\. \ definition cblinfun_apply_ell2 :: "'a ell2 \\<^sub>C\<^sub>L 'b ell2 \ 'a ell2 \ 'b ell2" where [code del, code_abbrev]: "cblinfun_apply_ell2 = (*\<^sub>V)" \ \@{attribute code_abbrev} instructs the code generation to replace the rhs \<^term>\(*\<^sub>V)\ by the lhs \<^term>\cblinfun_apply_ell2\ before starting the actual code generation.\ lemma cblinfun_apply_ell2[code]: \ \Code equation for \<^term>\cblinfun_apply_ell2\, i.e., for applying an operator to an \<^type>\ell2\ vector\ "vec_of_ell2 (cblinfun_apply_ell2 M x) = (mult_mat_vec (mat_of_cblinfun M) (vec_of_ell2 x))" by (simp add: cblinfun_apply_ell2_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) definition butterfly_code :: \'a ell2 \ 'b ell2 \ 'b ell2 \\<^sub>C\<^sub>L 'a ell2\ where [code del, code_abbrev]: \butterfly_code = butterfly\ lemma butterfly_code[code]: \mat_of_cblinfun (butterfly_code s t) = mat_of_cols (CARD('a)) [vec_of_ell2 s] * mat_of_rows (CARD('b)) [map_vec cnj (vec_of_ell2 t)]\ for s :: \'a::enum ell2\ and t :: \'b::enum ell2\ by (simp add: butterfly_code_def butterfly_def vector_to_cblinfun_code mat_of_cblinfun_compose mat_of_cblinfun_adj mat_adjoint_def map_map_vec_cols flip: vector_to_cblinfun_code_def map_vec_conjugate[abs_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\).\ + (done by \<^term>\mk_projector\).\ "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)))" + (let d = length (canonical_basis :: 'a list) in mk_projector 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 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_ell2\, 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]: +lemma cblinfun_image_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) + apply (subst cblinfun_image_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,120 +1,122 @@ 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 \explicit_cblinfun (\x y :: bool. of_bool (x \ y))\ + 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/Cblinfun_Matrix.thy b/thys/Complex_Bounded_Operators/Cblinfun_Matrix.thy --- a/thys/Complex_Bounded_Operators/Cblinfun_Matrix.thy +++ b/thys/Complex_Bounded_Operators/Cblinfun_Matrix.thy @@ -1,1554 +1,1552 @@ section \\Cblinfun_Matrix\ -- Matrix representation of bounded operators\ theory Cblinfun_Matrix imports Complex_L2 "Jordan_Normal_Form.Gram_Schmidt" "HOL-Analysis.Starlike" "Complex_Bounded_Operators.Extra_Jordan_Normal_Form" begin hide_const (open) Order.bottom Order.top hide_type (open) Finite_Cartesian_Product.vec hide_const (open) Finite_Cartesian_Product.mat hide_fact (open) Finite_Cartesian_Product.mat_def hide_const (open) Finite_Cartesian_Product.vec hide_fact (open) Finite_Cartesian_Product.vec_def hide_const (open) Finite_Cartesian_Product.row hide_fact (open) Finite_Cartesian_Product.row_def no_notation Finite_Cartesian_Product.vec_nth (infixl "$" 90) unbundle jnf_notation unbundle cblinfun_notation subsection \Isomorphism between vectors\ text \We define the canonical isomorphism between vectors in some complex vector space \<^typ>\'a::basis_enum\ and the complex \<^term>\n\-dimensional vectors (where \<^term>\n\ is the dimension of \<^typ>\'a\). This is possible if \<^typ>\'a\, \<^typ>\'b\ are of class \<^class>\basis_enum\ since that class fixes a finite canonical basis. Vector are represented using the \<^typ>\complex vec\ type from \<^session>\Jordan_Normal_Form\. (The isomorphism will be called \<^term>\vec_of_onb_basis\ below.)\ definition vec_of_basis_enum :: \'a::basis_enum \ complex vec\ where \ \Maps \<^term>\v\ to a \<^typ>\'a vec\ represented in basis \<^const>\canonical_basis\\ \vec_of_basis_enum v = vec_of_list (map (crepresentation (set canonical_basis) v) canonical_basis)\ lemma dim_vec_of_basis_enum'[simp]: \dim_vec (vec_of_basis_enum (v::'a)) = length (canonical_basis::'a::basis_enum list)\ unfolding vec_of_basis_enum_def by simp definition basis_enum_of_vec :: \complex vec \ 'a::basis_enum\ where \basis_enum_of_vec v = (if dim_vec v = length (canonical_basis :: 'a list) - then sum_list (map2 (*\<^sub>C) (list_of_vec v) (canonical_basis::'a list)) + then sum_list (map2 (*\<^sub>C) (list_of_vec v) (canonical_basis :: 'a list)) else 0)\ lemma vec_of_basis_enum_inverse[simp]: - fixes w::"'a::basis_enum" - shows "basis_enum_of_vec (vec_of_basis_enum w) = w" + fixes \ :: "'a::basis_enum" + shows "basis_enum_of_vec (vec_of_basis_enum \) = \" unfolding vec_of_basis_enum_def basis_enum_of_vec_def unfolding list_vec zip_map1 zip_same_conv_map map_map apply (simp add: o_def) apply (subst sum.distinct_set_conv_list[symmetric], simp) apply (rule complex_vector.sum_representation_eq) using is_generator_set by auto lemma basis_enum_of_vec_inverse[simp]: - fixes v::"complex vec" + fixes v :: "complex vec" defines "n \ length (canonical_basis :: 'a::basis_enum list)" assumes f1: "dim_vec v = n" shows "vec_of_basis_enum ((basis_enum_of_vec v)::'a) = v" proof (rule eq_vecI) show \dim_vec (vec_of_basis_enum (basis_enum_of_vec v :: 'a)) = dim_vec v\ by (auto simp: vec_of_basis_enum_def f1 n_def) next fix j assume j_v: \j < dim_vec v\ define w where "w = list_of_vec v" define basis where "basis = (canonical_basis::'a list)" have [simp]: "length w = n" "length basis = n" \dim_vec v = n\ \length (canonical_basis::'a list) = n\ \j < n\ using j_v by (auto simp: f1 basis_def w_def n_def) have [simp]: \cindependent (set basis)\ \cspan (set basis) = UNIV\ by (auto simp: basis_def is_cindependent_set is_generator_set) have \vec_of_basis_enum ((basis_enum_of_vec v)::'a) $ j = map (crepresentation (set basis) (sum_list (map2 (*\<^sub>C) w basis))) basis ! j\ by (auto simp: vec_of_list_index vec_of_basis_enum_def basis_enum_of_vec_def simp flip: w_def basis_def) also have \\ = crepresentation (set basis) (sum_list (map2 (*\<^sub>C) w basis)) (basis!j)\ by simp also have \\ = crepresentation (set basis) (\iC (basis!i)) (basis!j)\ by (auto simp: sum_list_sum_nth atLeast0LessThan) also have \\ = (\iC crepresentation (set basis) (basis!i) (basis!j))\ by (auto simp: complex_vector.representation_sum complex_vector.representation_scale) also have \\ = w!j\ apply (subst sum_single[where i=j]) apply (auto simp: complex_vector.representation_basis) using \j < n\ \length basis = n\ basis_def distinct_canonical_basis nth_eq_iff_index_eq by blast also have \\ = v $ j\ by (simp add: w_def) finally show \vec_of_basis_enum (basis_enum_of_vec v :: 'a) $ j = v $ j\ by - qed lemma basis_enum_eq_vec_of_basis_enumI: fixes a b :: "_::basis_enum" assumes "vec_of_basis_enum a = vec_of_basis_enum b" shows "a = b" by (metis assms vec_of_basis_enum_inverse) -lemma vec_of_basis_enum_carrier_vec[simp]: \vec_of_basis_enum v \ carrier_vec CARD('a)\ for v :: \'a::enum ell2\ +lemma vec_of_basis_enum_carrier_vec[simp]: \vec_of_basis_enum v \ carrier_vec (canonical_basis_length TYPE('a))\ for v :: \'a::basis_enum\ apply (simp only: dim_vec_of_basis_enum' carrier_vec_def vec_of_basis_enum_def) - by simp + by (simp add: canonical_basis_length) + +lemma vec_of_basis_enum_inj: "inj vec_of_basis_enum" + by (simp add: basis_enum_eq_vec_of_basis_enumI injI) + +lemma basis_enum_of_vec_inj: "inj_on (basis_enum_of_vec :: complex vec \ 'a) + (carrier_vec (length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)))" + by (metis basis_enum_of_vec_inverse carrier_dim_vec inj_on_inverseI) subsection \Operations on vectors\ lemma basis_enum_of_vec_add: assumes [simp]: \dim_vec v1 = length (canonical_basis :: 'a::basis_enum list)\ \dim_vec v2 = length (canonical_basis :: 'a list)\ shows \((basis_enum_of_vec (v1 + v2)) :: 'a) = basis_enum_of_vec v1 + basis_enum_of_vec v2\ proof - have \length (list_of_vec v1) = length (list_of_vec v2)\ and \length (list_of_vec v2) = length (canonical_basis :: 'a list)\ by simp_all then have \sum_list (map2 (*\<^sub>C) (map2 (+) (list_of_vec v1) (list_of_vec v2)) (canonical_basis::'a list)) = sum_list (map2 (*\<^sub>C) (list_of_vec v1) canonical_basis) + sum_list (map2 (*\<^sub>C) (list_of_vec v2) canonical_basis)\ apply (induction rule: list_induct3) by (auto simp: scaleC_add_left) then show ?thesis using assms by (auto simp: basis_enum_of_vec_def list_of_vec_plus) qed lemma basis_enum_of_vec_mult: assumes [simp]: \dim_vec v = length (canonical_basis :: 'a::basis_enum list)\ shows \((basis_enum_of_vec (c \\<^sub>v v)) :: 'a) = c *\<^sub>C basis_enum_of_vec v\ proof - have *: \monoid_add_hom ((*\<^sub>C) c :: 'a \ _)\ by (simp add: monoid_add_hom_def plus_hom.intro scaleC_add_right semigroup_add_hom.intro zero_hom.intro) show ?thesis apply (auto simp: basis_enum_of_vec_def list_of_vec_mult map_zip_map monoid_add_hom.hom_sum_list[OF *]) by (metis case_prod_unfold comp_apply scaleC_scaleC) qed lemma vec_of_basis_enum_add: - "vec_of_basis_enum (b1 + b2) = vec_of_basis_enum b1 + vec_of_basis_enum b2" + \vec_of_basis_enum (a + b) = vec_of_basis_enum a + vec_of_basis_enum b\ by (auto simp: vec_of_basis_enum_def complex_vector.representation_add) lemma vec_of_basis_enum_scaleC: "vec_of_basis_enum (c *\<^sub>C b) = c \\<^sub>v (vec_of_basis_enum b)" by (auto simp: vec_of_basis_enum_def complex_vector.representation_scale) lemma vec_of_basis_enum_scaleR: "vec_of_basis_enum (r *\<^sub>R b) = complex_of_real r \\<^sub>v (vec_of_basis_enum b)" by (simp add: scaleR_scaleC vec_of_basis_enum_scaleC) lemma vec_of_basis_enum_uminus: "vec_of_basis_enum (- b2) = - vec_of_basis_enum b2" unfolding scaleC_minus1_left[symmetric, of b2] unfolding scaleC_minus1_left_vec[symmetric] by (rule vec_of_basis_enum_scaleC) lemma vec_of_basis_enum_minus: "vec_of_basis_enum (b1 - b2) = vec_of_basis_enum b1 - vec_of_basis_enum b2" by (metis (mono_tags, opaque_lifting) carrier_vec_dim_vec diff_conv_add_uminus diff_zero index_add_vec(2) minus_add_uminus_vec vec_of_basis_enum_add vec_of_basis_enum_uminus) lemma cinner_basis_enum_of_vec: defines "n \ length (canonical_basis :: 'a::onb_enum list)" assumes [simp]: "dim_vec x = n" "dim_vec y = n" shows "(basis_enum_of_vec x :: 'a) \\<^sub>C basis_enum_of_vec y = y \c x" proof - have \(basis_enum_of_vec x :: 'a) \\<^sub>C basis_enum_of_vec y = (\iC canonical_basis ! i :: 'a) \\<^sub>C (\iC canonical_basis ! i)\ by (auto simp: basis_enum_of_vec_def sum_list_sum_nth atLeast0LessThan simp flip: n_def) also have \\ = (\ijC y$j *\<^sub>C ((canonical_basis ! i :: 'a) \\<^sub>C (canonical_basis ! j)))\ apply (subst cinner_sum_left) apply (subst cinner_sum_right) by (auto simp: mult_ac) also have \\ = (\ijC y$j *\<^sub>C (if i=j then 1 else 0))\ apply (rule sum.cong[OF refl]) apply (rule sum.cong[OF refl]) by (auto simp: cinner_canonical_basis n_def) also have \\ = (\iC y$i)\ apply (rule sum.cong[OF refl]) apply (subst sum_single) by auto also have \\ = y \c x\ by (smt (z3) assms(2) complex_scaleC_def conjugate_complex_def dim_vec_conjugate lessThan_atLeast0 lessThan_iff mult.commute scalar_prod_def sum.cong vec_index_conjugate) finally show ?thesis by - qed lemma cscalar_prod_vec_of_basis_enum: "cscalar_prod (vec_of_basis_enum \) (vec_of_basis_enum \) = cinner \ \" for \ :: "'a::onb_enum" apply (subst cinner_basis_enum_of_vec[symmetric, where 'a='a]) by simp_all definition \norm_vec \ = sqrt (\ i \ {0 ..< dim_vec \}. let z = vec_index \ i in (Re z)\<^sup>2 + (Im z)\<^sup>2)\ -lemma norm_ell2_vec_of_basis_enum: \norm \ = norm_vec (vec_of_basis_enum \)\ for \ :: "'a::onb_enum" +lemma norm_vec_of_basis_enum: \norm \ = norm_vec (vec_of_basis_enum \)\ for \ :: "'a::onb_enum" proof - have "norm \ = sqrt (cmod (\i = 0..). vec_of_basis_enum \ $ i * conjugate (vec_of_basis_enum \) $ i))" unfolding norm_eq_sqrt_cinner[where 'a='a] cscalar_prod_vec_of_basis_enum[symmetric] scalar_prod_def dim_vec_conjugate by rule also have "\ = sqrt (cmod (\x = 0..). let z = vec_of_basis_enum \ $ x in (Re z)\<^sup>2 + (Im z)\<^sup>2))" apply (subst sum.cong, rule refl) apply (subst vec_index_conjugate) by (auto simp: Let_def complex_mult_cnj) also have "\ = norm_vec (vec_of_basis_enum \)" unfolding Let_def norm_of_real norm_vec_def apply (subst abs_of_nonneg) apply (rule sum_nonneg) by auto finally show ?thesis by - qed lemma basis_enum_of_vec_unit_vec: defines "basis \ (canonical_basis::'a::basis_enum list)" and "n \ length (canonical_basis :: 'a list)" assumes a3: "i < n" shows "basis_enum_of_vec (unit_vec n i) = basis!i" proof- define L::"complex list" where "L = list_of_vec (unit_vec n i)" define I::"nat list" where "I = [0..C) L basis = map (\j. L!j *\<^sub>C basis!j) I" by (simp add: I_def list_eq_iff_nth_eq) hence "sum_list (map2 (*\<^sub>C) L basis) = sum_list (map (\j. L!j *\<^sub>C basis!j) I)" by simp also have "\ = sum (\j. L!j *\<^sub>C basis!j) {0..n-1}" proof- have "set I = {0..n-1}" using I_def a3 by auto thus ?thesis using Groups_List.sum_code[where xs = I and g = "(\j. L!j *\<^sub>C basis!j)"] by (simp add: I_def) qed also have "\ = sum (\j. (list_of_vec (unit_vec n i))!j *\<^sub>C basis!j) {0..n-1}" unfolding L_def by blast finally have "sum_list (map2 (*\<^sub>C) (list_of_vec (unit_vec n i)) basis) = sum (\j. (list_of_vec (unit_vec n i))!j *\<^sub>C basis!j) {0..n-1}" using L_def by blast also have "\ = basis ! i" proof- have "(\j = 0..n - 1. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j) = (\j \ {0..n - 1}. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j)" by simp also have "\ = list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i + (\j \ {0..n - 1}-{i}. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j)" proof- define a where "a j = list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j" for j define S where "S = {0..n - 1}" have "finite S" by (simp add: S_def) hence "(\j \ insert i S. a j) = a i + (\j\S-{i}. a j)" using Groups_Big.comm_monoid_add_class.sum.insert_remove by auto moreover have "S-{i} = {0..n-1}-{i}" unfolding S_def by blast moreover have "insert i S = {0..n-1}" using S_def Suc_diff_1 a3 atLeastAtMost_iff diff_is_0_eq' le_SucE le_numeral_extra(4) less_imp_le not_gr_zero by fastforce ultimately show ?thesis using \a \ \j. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j\ by simp qed also have "\ = list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i" proof- have "j \ {0..n - 1}-{i} \ list_of_vec (unit_vec n i) ! j = 0" for j using a3 atMost_atLeast0 atMost_iff diff_Suc_less index_unit_vec(1) le_less_trans list_of_vec_index member_remove zero_le by fastforce hence "j \ {0..n - 1}-{i} \ list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j = 0" for j by auto hence "(\j \ {0..n - 1}-{i}. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j) = 0" by (simp add: \\j. j \ {0..n - 1} - {i} \ list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j = 0\) thus ?thesis by simp qed also have "\ = basis ! i" by (simp add: a3) finally show ?thesis using \(\j = 0..n - 1. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j) = list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i + (\j\{0..n - 1} - {i}. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j)\ \list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i + (\j\{0..n - 1} - {i}. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j) = list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i\ \list_of_vec (unit_vec n i) ! i *\<^sub>C basis ! i = basis ! i\ by auto qed finally have "sum_list (map2 (*\<^sub>C) (list_of_vec (unit_vec n i)) basis) = basis ! i" by (simp add: assms) hence "sum_list (map2 scaleC (list_of_vec (unit_vec n i)) (canonical_basis::'a list)) = (canonical_basis::'a list) ! i" by (simp add: assms) thus ?thesis unfolding basis_enum_of_vec_def by (simp add: assms) qed lemma vec_of_basis_enum_ket: "vec_of_basis_enum (ket i) = unit_vec (CARD('a)) (enum_idx i)" for i::"'a::enum" proof- have "dim_vec (vec_of_basis_enum (ket i)) = dim_vec (unit_vec (CARD('a)) (enum_idx i))" proof- have "dim_vec (unit_vec (CARD('a)) (enum_idx i)) = CARD('a)" by simp moreover have "dim_vec (vec_of_basis_enum (ket i)) = CARD('a)" unfolding vec_of_basis_enum_def vec_of_basis_enum_def by auto ultimately show ?thesis by simp qed moreover have "vec_of_basis_enum (ket i) $ j = (unit_vec (CARD('a)) (enum_idx i)) $ j" if "j < dim_vec (vec_of_basis_enum (ket i))" for j proof- have j_bound: "j < length (canonical_basis::'a ell2 list)" by (metis dim_vec_of_basis_enum' that) have y1: "cindependent (set (canonical_basis::'a ell2 list))" using is_cindependent_set by blast have y2: "canonical_basis ! j \ set (canonical_basis::'a ell2 list)" using j_bound by auto have p1: "enum_class.enum ! enum_idx i = i" using enum_idx_correct by blast moreover have p2: "(canonical_basis::'a ell2 list) ! t = ket ((enum_class.enum::'a list) ! t)" if "t < length (enum_class.enum::'a list)" for t unfolding canonical_basis_ell2_def using that by auto moreover have p3: "enum_idx i < length (enum_class.enum::'a list)" proof- have "set (enum_class.enum::'a list) = UNIV" using UNIV_enum by blast hence "i \ set (enum_class.enum::'a list)" by blast thus ?thesis unfolding enum_idx_def by (metis index_of_bound length_greater_0_conv length_pos_if_in_set) qed ultimately have p4: "(canonical_basis::'a ell2 list) ! (enum_idx i) = ket i" by auto have "enum_idx i < length (enum_class.enum::'a list)" using p3 by auto moreover have "length (enum_class.enum::'a list) = dim_vec (vec_of_basis_enum (ket i))" unfolding vec_of_basis_enum_def canonical_basis_ell2_def using dim_vec_of_basis_enum'[where v = "ket i"] unfolding canonical_basis_ell2_def by simp ultimately have enum_i_dim_vec: "enum_idx i < dim_vec (unit_vec (CARD('a)) (enum_idx i))" using \dim_vec (vec_of_basis_enum (ket i)) = dim_vec (unit_vec (CARD('a)) (enum_idx i))\ by auto hence r1: "(unit_vec (CARD('a)) (enum_idx i)) $ j = (if enum_idx i = j then 1 else 0)" using \dim_vec (vec_of_basis_enum (ket i)) = dim_vec (unit_vec (CARD('a)) (enum_idx i))\ that by auto moreover have "vec_of_basis_enum (ket i) $ j = (if enum_idx i = j then 1 else 0)" proof(cases "enum_idx i = j") case True have "crepresentation (set (canonical_basis::'a ell2 list)) ((canonical_basis::'a ell2 list) ! j) ((canonical_basis::'a ell2 list) ! j) = 1" using y1 y2 complex_vector.representation_basis[where basis = "set (canonical_basis::'a ell2 list)" and b = "(canonical_basis::'a ell2 list) ! j"] by smt hence "vec_of_basis_enum ((canonical_basis::'a ell2 list) ! j) $ j = 1" unfolding vec_of_basis_enum_def by (metis j_bound nth_map vec_of_list_index) hence "vec_of_basis_enum ((canonical_basis::'a ell2 list) ! (enum_idx i)) $ enum_idx i = 1" using True by simp hence "vec_of_basis_enum (ket i) $ enum_idx i = 1" using p4 by simp thus ?thesis using True unfolding vec_of_basis_enum_def by auto next case False have "crepresentation (set (canonical_basis::'a ell2 list)) ((canonical_basis::'a ell2 list) ! (enum_idx i)) ((canonical_basis::'a ell2 list) ! j) = 0" using y1 y2 complex_vector.representation_basis[where basis = "set (canonical_basis::'a ell2 list)" and b = "(canonical_basis::'a ell2 list) ! j"] by (metis (mono_tags, opaque_lifting) False enum_i_dim_vec basis_enum_of_vec_inverse basis_enum_of_vec_unit_vec canonical_basis_length_ell2 index_unit_vec(3) j_bound list_of_vec_index list_vec nth_map r1 vec_of_basis_enum_def) hence "vec_of_basis_enum ((canonical_basis::'a ell2 list) ! (enum_idx i)) $ j = 0" unfolding vec_of_basis_enum_def by (smt j_bound nth_map vec_of_list_index) hence "vec_of_basis_enum ((canonical_basis::'a ell2 list) ! (enum_idx i)) $ j = 0" by auto hence "vec_of_basis_enum (ket i) $ j = 0" using p4 by simp thus ?thesis using False unfolding vec_of_basis_enum_def by simp qed ultimately show ?thesis by auto qed ultimately show ?thesis using Matrix.eq_vecI by auto qed lemma vec_of_basis_enum_zero: defines "nA \ length (canonical_basis :: 'a::basis_enum list)" shows "vec_of_basis_enum (0::'a) = 0\<^sub>v nA" by (metis assms carrier_vecI dim_vec_of_basis_enum' minus_cancel_vec right_minus_eq vec_of_basis_enum_minus) lemma (in complex_vec_space) vec_of_basis_enum_cspan: fixes X :: "'a::basis_enum set" assumes "length (canonical_basis :: 'a list) = n" shows "vec_of_basis_enum ` cspan X = span (vec_of_basis_enum ` X)" proof - have carrier: "vec_of_basis_enum ` X \ carrier_vec n" by (metis assms carrier_vecI dim_vec_of_basis_enum' image_subsetI) have lincomb_sum: "lincomb c A = vec_of_basis_enum (\b\B. c' b *\<^sub>C b)" if B_def: "B = basis_enum_of_vec ` A" and \finite A\ and AX: "A \ vec_of_basis_enum ` X" and c'_def: "\b. c' b = c (vec_of_basis_enum b)" for c c' A and B::"'a set" unfolding B_def using \finite A\ AX proof induction case empty then show ?case by (simp add: assms vec_of_basis_enum_zero) next case (insert x F) then have IH: "lincomb c F = vec_of_basis_enum (\b\basis_enum_of_vec ` F. c' b *\<^sub>C b)" (is "_ = ?sum") by simp have xx: "vec_of_basis_enum (basis_enum_of_vec x :: 'a) = x" apply (rule basis_enum_of_vec_inverse) using assms carrier carrier_vecD insert.prems by auto have "lincomb c (insert x F) = c x \\<^sub>v x + lincomb c F" apply (rule lincomb_insert2) using insert.hyps insert.prems carrier by auto also have "c x \\<^sub>v x = vec_of_basis_enum (c' (basis_enum_of_vec x) *\<^sub>C (basis_enum_of_vec x :: 'a))" by (simp add: xx vec_of_basis_enum_scaleC c'_def) also note IH also have "vec_of_basis_enum (c' (basis_enum_of_vec x) *\<^sub>C (basis_enum_of_vec x :: 'a)) + ?sum = vec_of_basis_enum (\b\basis_enum_of_vec ` insert x F. c' b *\<^sub>C b)" apply simp apply (rule sym) apply (subst sum.insert) using \finite F\ \x \ F\ dim_vec_of_basis_enum' insert.prems vec_of_basis_enum_add c'_def by auto finally show ?case by - qed show ?thesis proof auto fix x assume "x \ local.span (vec_of_basis_enum ` X)" then obtain c A where xA: "x = lincomb c A" and "finite A" and AX: "A \ vec_of_basis_enum ` X" unfolding span_def by auto define B::"'a set" and c' where "B = basis_enum_of_vec ` A" and "c' b = c (vec_of_basis_enum b)" for b::'a note xA also have "lincomb c A = vec_of_basis_enum (\b\B. c' b *\<^sub>C b)" using B_def \finite A\ AX c'_def by (rule lincomb_sum) also have "\ \ vec_of_basis_enum ` cspan X" unfolding complex_vector.span_explicit apply (rule imageI) apply (rule CollectI) apply (rule exI) apply (rule exI) using \finite A\ AX by (auto simp: B_def) finally show "x \ vec_of_basis_enum ` cspan X" by - next fix x assume "x \ cspan X" then obtain c' B where x: "x = (\b\B. c' b *\<^sub>C b)" and "finite B" and BX: "B \ X" unfolding complex_vector.span_explicit by auto define A and c where "A = vec_of_basis_enum ` B" and "c b = c' (basis_enum_of_vec b)" for b have "vec_of_basis_enum x = vec_of_basis_enum (\b\B. c' b *\<^sub>C b)" using x by simp also have "\ = lincomb c A" apply (rule lincomb_sum[symmetric]) unfolding A_def c_def using BX \finite B\ by (auto simp: image_image) also have "\ \ span (vec_of_basis_enum ` X)" unfolding span_def apply (rule CollectI) apply (rule exI, rule exI) unfolding A_def using \finite B\ BX by auto finally show "vec_of_basis_enum x \ local.span (vec_of_basis_enum ` X)" by - qed qed lemma (in complex_vec_space) basis_enum_of_vec_span: assumes "length (canonical_basis :: 'a list) = n" assumes "Y \ carrier_vec n" shows "basis_enum_of_vec ` local.span Y = cspan (basis_enum_of_vec ` Y :: 'a::basis_enum set)" proof - define X::"'a set" where "X = basis_enum_of_vec ` Y" then have "cspan (basis_enum_of_vec ` Y :: 'a set) = basis_enum_of_vec ` vec_of_basis_enum ` cspan X" by (simp add: image_image) also have "\ = basis_enum_of_vec ` local.span (vec_of_basis_enum ` X)" apply (subst vec_of_basis_enum_cspan) using assms by simp_all also have "vec_of_basis_enum ` X = Y" unfolding X_def image_image apply (rule image_cong[where g=id and M=Y and N=Y, simplified]) using assms(1) assms(2) by auto finally show ?thesis by simp qed lemma vec_of_basis_enum_canonical_basis: assumes "i < length (canonical_basis :: 'a list)" shows "vec_of_basis_enum (canonical_basis!i :: 'a) = unit_vec (length (canonical_basis :: 'a::basis_enum list)) i" by (metis assms basis_enum_of_vec_inverse basis_enum_of_vec_unit_vec index_unit_vec(3)) lemma vec_of_basis_enum_times: fixes \ \ :: "'a::one_dim" shows "vec_of_basis_enum (\ * \) = vec_of_list [vec_index (vec_of_basis_enum \) 0 * vec_index (vec_of_basis_enum \) 0]" proof - have [simp]: \crepresentation {1} x 1 = one_dim_iso x\ for x :: 'a apply (subst one_dim_scaleC_1[where x=x, symmetric]) apply (subst complex_vector.representation_scale) by (auto simp add: complex_vector.span_base complex_vector.representation_basis) show ?thesis apply (rule eq_vecI) by (auto simp: vec_of_basis_enum_def) qed lemma vec_of_basis_enum_to_inverse: fixes \ :: "'a::one_dim" shows "vec_of_basis_enum (inverse \) = vec_of_list [inverse (vec_index (vec_of_basis_enum \) 0)]" proof - have [simp]: \crepresentation {1} x 1 = one_dim_iso x\ for x :: 'a apply (subst one_dim_scaleC_1[where x=x, symmetric]) apply (subst complex_vector.representation_scale) by (auto simp add: complex_vector.span_base complex_vector.representation_basis) show ?thesis apply (rule eq_vecI) apply (auto simp: vec_of_basis_enum_def) by (metis complex_vector.scale_cancel_right one_dim_inverse one_dim_scaleC_1 zero_neq_one) qed lemma vec_of_basis_enum_divide: fixes \ \ :: "'a::one_dim" shows "vec_of_basis_enum (\ / \) = vec_of_list [vec_index (vec_of_basis_enum \) 0 / vec_index (vec_of_basis_enum \) 0]" by (simp add: divide_inverse vec_of_basis_enum_to_inverse vec_of_basis_enum_times) lemma vec_of_basis_enum_1: "vec_of_basis_enum (1 :: 'a::one_dim) = vec_of_list [1]" proof - have [simp]: \crepresentation {1} x 1 = one_dim_iso x\ for x :: 'a apply (subst one_dim_scaleC_1[where x=x, symmetric]) apply (subst complex_vector.representation_scale) by (auto simp add: complex_vector.span_base complex_vector.representation_basis) show ?thesis apply (rule eq_vecI) by (auto simp: vec_of_basis_enum_def) qed lemma vec_of_basis_enum_ell2_component: fixes \ :: \'a::enum ell2\ assumes [simp]: \i < CARD('a)\ shows \vec_of_basis_enum \ $ i = Rep_ell2 \ (Enum.enum ! i)\ proof - let ?i = \Enum.enum ! i\ have \Rep_ell2 \ (Enum.enum ! i) = ket ?i \\<^sub>C \\ by (simp add: cinner_ket_left) also have \\ = vec_of_basis_enum \ \c vec_of_basis_enum (ket ?i :: 'a ell2)\ by (rule cscalar_prod_vec_of_basis_enum[symmetric]) also have \\ = vec_of_basis_enum \ \c unit_vec (CARD('a)) i\ by (simp add: vec_of_basis_enum_ket enum_idx_enum) also have \\ = vec_of_basis_enum \ \ unit_vec (CARD('a)) i\ by (smt (verit, best) assms carrier_vecI conjugate_conjugate_sprod conjugate_id conjugate_vec_sprod_comm dim_vec_conjugate eq_vecI index_unit_vec(3) scalar_prod_left_unit vec_index_conjugate) also have \\ = vec_of_basis_enum \ $ i\ by simp finally show ?thesis by simp qed lemma corthogonal_vec_of_basis_enum: fixes S :: "'a::onb_enum list" shows "corthogonal (map vec_of_basis_enum S) \ is_ortho_set (set S) \ distinct S" proof auto assume assm: \corthogonal (map vec_of_basis_enum S)\ then show \is_ortho_set (set S)\ by (smt (verit, ccfv_SIG) cinner_eq_zero_iff corthogonal_def cscalar_prod_vec_of_basis_enum in_set_conv_nth is_ortho_set_def length_map nth_map) show \distinct S\ using assm corthogonal_distinct distinct_map by blast next assume \is_ortho_set (set S)\ and \distinct S\ then show \corthogonal (map vec_of_basis_enum S)\ by (smt (verit, ccfv_threshold) cinner_eq_zero_iff corthogonalI cscalar_prod_vec_of_basis_enum is_ortho_set_def length_map length_map nth_eq_iff_index_eq nth_map nth_map nth_mem nth_mem) qed subsection \Isomorphism between bounded linear functions and matrices\ text \We define the canonical isomorphism between \<^typ>\'a::basis_enum \\<^sub>C\<^sub>L'b::basis_enum\ and the complex \<^term>\n*m\-matrices (where n,m are the dimensions of \<^typ>\'a\, \<^typ>\'b\, respectively). This is possible if \<^typ>\'a\, \<^typ>\'b\ are of class \<^class>\basis_enum\ since that class fixes a finite canonical basis. Matrices are represented using the \<^typ>\complex mat\ type from \<^session>\Jordan_Normal_Form\. (The isomorphism will be called \<^term>\mat_of_cblinfun\ below.)\ definition mat_of_cblinfun :: \'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L'b::{basis_enum,complex_normed_vector} \ complex mat\ where \mat_of_cblinfun f = mat (length (canonical_basis :: 'b list)) (length (canonical_basis :: 'a list)) ( \ (i, j). crepresentation (set (canonical_basis::'b list)) (f *\<^sub>V ((canonical_basis::'a list)!j)) ((canonical_basis::'b list)!i))\ for f lift_definition cblinfun_of_mat :: \complex mat \ 'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L'b::{basis_enum,complex_normed_vector}\ is - \\M. \v. (if M\carrier_mat (length (canonical_basis :: 'b list)) (length (canonical_basis :: 'a list)) - then basis_enum_of_vec (M *\<^sub>v vec_of_basis_enum v) - else 0)\ -proof + \\M. if M\carrier_mat (length (canonical_basis :: 'b list)) (length (canonical_basis :: 'a list)) + then \v. basis_enum_of_vec (M *\<^sub>v vec_of_basis_enum v) + else (\v. 0)\ +proof (intro bounded_clinear_finite_dim clinearI) fix M :: "complex mat" define m where "m = length (canonical_basis :: 'b list)" define n where "n = length (canonical_basis :: 'a list)" define f::"complex mat \ 'a \ 'b" - where "f M v = (if M\carrier_mat m n - then basis_enum_of_vec (M *\<^sub>v vec_of_basis_enum (v::'a)) - else (0::'b))" - for M::"complex mat" and v::'a + where "f M = (if M\carrier_mat m n + then \v. basis_enum_of_vec (M *\<^sub>v vec_of_basis_enum (v::'a)) + else (\v. 0))" + for M::"complex mat" show add: \f M (b1 + b2) = f M b1 + f M b2\ for b1 b2 apply (auto simp: f_def) by (metis (mono_tags, lifting) carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec dim_vec_of_basis_enum' m_def mult_add_distrib_mat_vec n_def basis_enum_of_vec_add vec_of_basis_enum_add) show scale: \f M (c *\<^sub>C b) = c *\<^sub>C f M b\ for c b apply (auto simp: f_def) by (metis carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec dim_vec_of_basis_enum' m_def mult_mat_vec n_def basis_enum_of_vec_mult vec_of_basis_enum_scaleC) - - from add scale have \clinear (f M)\ - by (simp add: clinear_iff) - - show \\K. \b. norm (f M b) \ norm b * K\ - proof (cases "M\carrier_mat m n") - case True - have f_def': "f M v = basis_enum_of_vec (M *\<^sub>v (vec_of_basis_enum v))" for v - using True f_def - m_def n_def by auto - have M_carrier_mat: - "M \ carrier_mat m n" - by (simp add: True) - have "bounded_clinear (f M)" - apply (rule bounded_clinear_finite_dim) using \clinear (f M)\ by auto - thus ?thesis - by (simp add: bounded_clinear.bounded) - next - case False - thus ?thesis - unfolding f_def m_def n_def - by (metis (full_types) order_refl mult_eq_0_iff norm_eq_zero) - qed qed lemma cblinfun_of_mat_invalid: assumes \M \ carrier_mat (canonical_basis_length TYPE('b::{basis_enum,complex_normed_vector})) (canonical_basis_length TYPE('a::{basis_enum,complex_normed_vector}))\ shows \(cblinfun_of_mat M :: 'a \\<^sub>C\<^sub>L 'b) = 0\ apply (transfer fixing: M) using assms by (simp add: canonical_basis_length) -lemma mat_of_cblinfun_ell2_carrier[simp]: \mat_of_cblinfun (a::'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2) \ carrier_mat CARD('b) CARD('a)\ - by (simp add: mat_of_cblinfun_def) +lemma dim_row_mat_of_cblinfun[simp]: \dim_row (mat_of_cblinfun (a::'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector})) = canonical_basis_length TYPE('b)\ + by (simp add: mat_of_cblinfun_def canonical_basis_length) -lemma dim_row_mat_of_cblinfun[simp]: \dim_row (mat_of_cblinfun (a::'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2)) = CARD('b)\ - by (simp add: mat_of_cblinfun_def) +lemma dim_col_mat_of_cblinfun[simp]: \dim_col (mat_of_cblinfun (a::'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector})) = canonical_basis_length TYPE('a)\ + by (simp add: mat_of_cblinfun_def canonical_basis_length) -lemma dim_col_mat_of_cblinfun[simp]: \dim_col (mat_of_cblinfun (a::'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2)) = CARD('a)\ - by (simp add: mat_of_cblinfun_def) +lemma mat_of_cblinfun_ell2_carrier[simp]: \mat_of_cblinfun (a::'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}) \ carrier_mat (canonical_basis_length TYPE('b)) (canonical_basis_length TYPE('a))\ + by (auto intro!: carrier_matI) + +lemma basis_enum_of_vec_cblinfun_apply: + fixes M :: "complex mat" + defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" + and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" + assumes "M \ carrier_mat nB nA" and "dim_vec x = nA" + shows "basis_enum_of_vec (M *\<^sub>v x) = (cblinfun_of_mat M :: 'a \\<^sub>C\<^sub>L 'b) *\<^sub>V basis_enum_of_vec x" + by (metis assms basis_enum_of_vec_inverse cblinfun_of_mat.rep_eq) lemma mat_of_cblinfun_cblinfun_apply: - "vec_of_basis_enum (F *\<^sub>V u) = mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u" - for F::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" and u::'a + \vec_of_basis_enum (F *\<^sub>V u) = mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u\ + for F::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" and u::'a proof (rule eq_vecI) show \dim_vec (vec_of_basis_enum (F *\<^sub>V u)) = dim_vec (mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u)\ by (simp add: dim_vec_of_basis_enum' mat_of_cblinfun_def) next fix i define BasisA where "BasisA = (canonical_basis::'a list)" define BasisB where "BasisB = (canonical_basis::'b list)" define nA where "nA = length (canonical_basis :: 'a list)" define nB where "nB = length (canonical_basis :: 'b list)" assume \i < dim_vec (mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u)\ then have [simp]: \i < nB\ by (simp add: mat_of_cblinfun_def nB_def) define v where \v = BasisB ! i\ - have [simp]: \dim_row (mat_of_cblinfun F) = nB\ + have dim_row_F[simp]: \dim_row (mat_of_cblinfun F) = nB\ by (simp add: mat_of_cblinfun_def nB_def) have [simp]: \length BasisB = nB\ by (simp add: nB_def BasisB_def) have [simp]: \length BasisA = nA\ using BasisA_def nA_def by auto have [simp]: \cindependent (set BasisA)\ using BasisA_def is_cindependent_set by auto have [simp]: \cindependent (set BasisB)\ using BasisB_def is_cindependent_set by blast have [simp]: \cspan (set BasisB) = UNIV\ by (simp add: BasisB_def is_generator_set) have [simp]: \cspan (set BasisA) = UNIV\ by (simp add: BasisA_def is_generator_set) have \(mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u) $ i = (\j = 0.. - by (auto simp: vec_of_basis_enum_def scalar_prod_def simp flip: BasisA_def) + using dim_row_F by (auto simp: mult_mat_vec_def vec_of_basis_enum_def scalar_prod_def simp flip: BasisA_def) also have \\ = (\j = 0..V BasisA ! j) v * crepresentation (set BasisA) u (BasisA ! j))\ apply (rule sum.cong[OF refl]) by (auto simp: vec_of_list_index mat_of_cblinfun_def scalar_prod_def v_def simp flip: BasisA_def BasisB_def) also have \\ = crepresentation (set BasisB) (F *\<^sub>V u) v\ (is \(\j=_..<_. ?lhs v j) = _\) proof (rule complex_vector.representation_eqI[symmetric, THEN fun_cong]) show \cindependent (set BasisB)\ \F *\<^sub>V u \ cspan (set BasisB)\ by simp_all show only_basis: \(\j = 0.. 0 \ b \ set BasisB\ for b by (metis (mono_tags, lifting) complex_vector.representation_ne_zero mult_not_zero sum.not_neutral_contains_not_neutral) then show \finite {b. (\j = 0.. 0}\ by (smt (z3) List.finite_set finite_subset mem_Collect_eq subsetI) have \(\b | (\j = 0.. 0. (\j = 0..C b) = (\b\set BasisB. (\j = 0..C b)\ apply (rule sum.mono_neutral_left) using only_basis by auto also have \\ = (\b\set BasisB. (\a\set BasisA. crepresentation (set BasisB) (F *\<^sub>V a) b * crepresentation (set BasisA) u a) *\<^sub>C b)\ apply (subst sum.reindex_bij_betw[where h=\nth BasisA\ and T=\set BasisA\]) apply (metis BasisA_def \length BasisA = nA\ atLeast0LessThan bij_betw_nth distinct_canonical_basis) by simp also have \\ = (\a\set BasisA. crepresentation (set BasisA) u a *\<^sub>C (\b\set BasisB. crepresentation (set BasisB) (F *\<^sub>V a) b *\<^sub>C b))\ apply (simp add: scaleC_sum_left scaleC_sum_right) apply (subst sum.swap) by (metis (no_types, lifting) mult.commute sum.cong) also have \\ = (\a\set BasisA. crepresentation (set BasisA) u a *\<^sub>C (F *\<^sub>V a))\ apply (subst complex_vector.sum_representation_eq) by auto also have \\ = F *\<^sub>V (\a\set BasisA. crepresentation (set BasisA) u a *\<^sub>C a)\ by (simp flip: cblinfun.scaleC_right cblinfun.sum_right) also have \\ = F *\<^sub>V u\ apply (subst complex_vector.sum_representation_eq) by auto finally show \(\b | (\j = 0.. 0. (\j = 0..C b) = F *\<^sub>V u\ by auto qed also have \crepresentation (set BasisB) (F *\<^sub>V u) v = vec_of_basis_enum (F *\<^sub>V u) $ i\ by (auto simp: vec_of_list_index vec_of_basis_enum_def v_def simp flip: BasisB_def) finally show \vec_of_basis_enum (F *\<^sub>V u) $ i = (mat_of_cblinfun F *\<^sub>v vec_of_basis_enum u) $ i\ by simp qed -lemma basis_enum_of_vec_cblinfun_apply: - fixes M :: "complex mat" - defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" - and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" - assumes "M \ carrier_mat nB nA" and "dim_vec x = nA" - shows "basis_enum_of_vec (M *\<^sub>v x) = (cblinfun_of_mat M :: 'a \\<^sub>C\<^sub>L 'b) *\<^sub>V basis_enum_of_vec x" - by (metis assms basis_enum_of_vec_inverse cblinfun_of_mat.rep_eq) - lemma mat_of_cblinfun_inverse: "cblinfun_of_mat (mat_of_cblinfun B) = B" for B::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" proof (rule cblinfun_eqI) fix x :: 'a define y where \y = vec_of_basis_enum x\ then have \cblinfun_of_mat (mat_of_cblinfun B) *\<^sub>V x = ((cblinfun_of_mat (mat_of_cblinfun B) :: 'a\\<^sub>C\<^sub>L'b) *\<^sub>V basis_enum_of_vec y)\ by simp also have \\ = basis_enum_of_vec (mat_of_cblinfun B *\<^sub>v vec_of_basis_enum (basis_enum_of_vec y :: 'a))\ apply (transfer fixing: B) by (simp add: mat_of_cblinfun_def) also have \\ = basis_enum_of_vec (vec_of_basis_enum (B *\<^sub>V x))\ by (simp add: mat_of_cblinfun_cblinfun_apply y_def) also have \\ = B *\<^sub>V x\ by simp finally show \cblinfun_of_mat (mat_of_cblinfun B) *\<^sub>V x = B *\<^sub>V x\ by - qed lemma mat_of_cblinfun_inj: "inj mat_of_cblinfun" by (metis inj_on_def mat_of_cblinfun_inverse) lemma cblinfun_of_mat_inverse: fixes M::"complex mat" defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" assumes "M \ carrier_mat nB nA" shows "mat_of_cblinfun (cblinfun_of_mat M :: 'a \\<^sub>C\<^sub>L 'b) = M" by (smt (verit) assms(3) basis_enum_of_vec_inverse carrier_matD(1) carrier_vecD cblinfun_of_mat.rep_eq dim_mult_mat_vec eq_mat_on_vecI mat_carrier mat_of_cblinfun_def mat_of_cblinfun_cblinfun_apply nA_def nB_def) lemma cblinfun_of_mat_inj: "inj_on (cblinfun_of_mat::complex mat \ 'a \\<^sub>C\<^sub>L 'b) (carrier_mat (length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)) (length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)))" using cblinfun_of_mat_inverse by (metis inj_onI) lemma cblinfun_eq_mat_of_cblinfunI: assumes "mat_of_cblinfun a = mat_of_cblinfun b" shows "a = b" by (metis assms mat_of_cblinfun_inverse) subsection \Operations on matrices\ lemma cblinfun_of_mat_plus: defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" assumes [simp,intro]: "M \ carrier_mat nB nA" and [simp,intro]: "N \ carrier_mat nB nA" shows "(cblinfun_of_mat (M + N) :: 'a \\<^sub>C\<^sub>L 'b) = ((cblinfun_of_mat M + cblinfun_of_mat N))" proof - have [simp]: \vec_of_basis_enum (v :: 'a) \ carrier_vec nA\ for v by (auto simp add: carrier_dim_vec dim_vec_of_basis_enum' nA_def) have [simp]: \dim_row M = nB\ \dim_row N = nB\ using carrier_matD(1) by auto show ?thesis apply (transfer fixing: M N) by (auto intro!: ext simp add: add_mult_distrib_mat_vec nA_def[symmetric] nB_def[symmetric] add_mult_distrib_mat_vec[where nr=nB and nc=nA] basis_enum_of_vec_add) qed lemma mat_of_cblinfun_zero: "mat_of_cblinfun (0 :: ('a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector})) = 0\<^sub>m (length (canonical_basis :: 'b list)) (length (canonical_basis :: 'a list))" unfolding mat_of_cblinfun_def by (auto simp: complex_vector.representation_zero) lemma mat_of_cblinfun_plus: "mat_of_cblinfun (F + G) = mat_of_cblinfun F + mat_of_cblinfun G" for F G::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L'b::{basis_enum,complex_normed_vector}" by (auto simp add: mat_of_cblinfun_def cblinfun.add_left complex_vector.representation_add) lemma mat_of_cblinfun_id: "mat_of_cblinfun (id_cblinfun :: ('a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L'a)) = 1\<^sub>m (length (canonical_basis :: 'a list))" apply (rule eq_matI) by (auto simp: mat_of_cblinfun_def complex_vector.representation_basis is_cindependent_set nth_eq_iff_index_eq) lemma mat_of_cblinfun_1: "mat_of_cblinfun (1 :: ('a::one_dim \\<^sub>C\<^sub>L'b::one_dim)) = 1\<^sub>m 1" apply (rule eq_matI) by (auto simp: mat_of_cblinfun_def complex_vector.representation_basis nth_eq_iff_index_eq) lemma mat_of_cblinfun_uminus: "mat_of_cblinfun (- M) = - mat_of_cblinfun M" for M::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L'b::{basis_enum,complex_normed_vector}" proof- define nA where "nA = length (canonical_basis :: 'a list)" define nB where "nB = length (canonical_basis :: 'b list)" have M1: "mat_of_cblinfun M \ carrier_mat nB nA" unfolding nB_def nA_def by (metis add.right_neutral add_carrier_mat mat_of_cblinfun_plus mat_of_cblinfun_zero nA_def nB_def zero_carrier_mat) have M2: "mat_of_cblinfun (-M) \ carrier_mat nB nA" by (metis add_carrier_mat mat_of_cblinfun_plus mat_of_cblinfun_zero diff_0 nA_def nB_def uminus_add_conv_diff zero_carrier_mat) have "mat_of_cblinfun (M - M) = 0\<^sub>m nB nA" unfolding nA_def nB_def by (simp add: mat_of_cblinfun_zero) moreover have "mat_of_cblinfun (M - M) = mat_of_cblinfun M + mat_of_cblinfun (- M)" by (metis mat_of_cblinfun_plus pth_2) ultimately have "mat_of_cblinfun M + mat_of_cblinfun (- M) = 0\<^sub>m nB nA" by simp thus ?thesis using M1 M2 by (smt add_uminus_minus_mat assoc_add_mat comm_add_mat left_add_zero_mat minus_r_inv_mat uminus_carrier_mat) qed lemma mat_of_cblinfun_minus: "mat_of_cblinfun (M - N) = mat_of_cblinfun M - mat_of_cblinfun N" for M::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" and N::"'a \\<^sub>C\<^sub>L'b" by (smt (z3) add_uminus_minus_mat mat_of_cblinfun_uminus mat_carrier mat_of_cblinfun_def mat_of_cblinfun_plus pth_2) lemma cblinfun_of_mat_uminus: defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" assumes "M \ carrier_mat nB nA" shows "(cblinfun_of_mat (-M) :: 'a \\<^sub>C\<^sub>L 'b) = - cblinfun_of_mat M" by (smt assms add.group_axioms add.right_neutral add_minus_cancel add_uminus_minus_mat cblinfun_of_mat_plus group.inverse_inverse mat_of_cblinfun_inverse mat_of_cblinfun_zero minus_r_inv_mat uminus_carrier_mat) lemma cblinfun_of_mat_minus: fixes M::"complex mat" defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" assumes "M \ carrier_mat nB nA" and "N \ carrier_mat nB nA" shows "(cblinfun_of_mat (M - N) :: 'a \\<^sub>C\<^sub>L 'b) = cblinfun_of_mat M - cblinfun_of_mat N" by (metis assms add_uminus_minus_mat cblinfun_of_mat_plus cblinfun_of_mat_uminus pth_2 uminus_carrier_mat) lemma cblinfun_of_mat_times: fixes M N ::"complex mat" defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" and "nB \ length (canonical_basis :: 'b::{basis_enum,complex_normed_vector} list)" and "nC \ length (canonical_basis :: 'c::{basis_enum,complex_normed_vector} list)" assumes a1: "M \ carrier_mat nC nB" and a2: "N \ carrier_mat nB nA" shows "cblinfun_of_mat (M * N) = ((cblinfun_of_mat M)::'b \\<^sub>C\<^sub>L'c) o\<^sub>C\<^sub>L ((cblinfun_of_mat N)::'a \\<^sub>C\<^sub>L'b)" proof - have b1: "((cblinfun_of_mat M)::'b \\<^sub>C\<^sub>L'c) v = basis_enum_of_vec (M *\<^sub>v vec_of_basis_enum v)" for v by (metis assms(4) cblinfun_of_mat.rep_eq nB_def nC_def) have b2: "((cblinfun_of_mat N)::'a \\<^sub>C\<^sub>L'b) v = basis_enum_of_vec (N *\<^sub>v vec_of_basis_enum v)" for v by (metis assms(5) cblinfun_of_mat.rep_eq nA_def nB_def) have b3: "((cblinfun_of_mat (M * N))::'a \\<^sub>C\<^sub>L'c) v = basis_enum_of_vec ((M * N) *\<^sub>v vec_of_basis_enum v)" for v by (metis assms(4) assms(5) cblinfun_of_mat.rep_eq mult_carrier_mat nA_def nC_def) have "(basis_enum_of_vec ((M * N) *\<^sub>v vec_of_basis_enum v)::'c) = (basis_enum_of_vec (M *\<^sub>v ( vec_of_basis_enum ( (basis_enum_of_vec (N *\<^sub>v vec_of_basis_enum v))::'b ))))" for v::'a proof- have c1: "vec_of_basis_enum (basis_enum_of_vec x :: 'b) = x" if "dim_vec x = nB" for x::"complex vec" using that unfolding nB_def by simp have c2: "vec_of_basis_enum v \ carrier_vec nA" by (metis (mono_tags, opaque_lifting) add.commute carrier_vec_dim_vec index_add_vec(2) index_zero_vec(2) nA_def vec_of_basis_enum_add basis_enum_of_vec_inverse) have "(M * N) *\<^sub>v vec_of_basis_enum v = M *\<^sub>v (N *\<^sub>v vec_of_basis_enum v)" using Matrix.assoc_mult_mat_vec a1 a2 c2 by blast hence "(basis_enum_of_vec ((M * N) *\<^sub>v vec_of_basis_enum v)::'c) = (basis_enum_of_vec (M *\<^sub>v (N *\<^sub>v vec_of_basis_enum v))::'c)" by simp also have "\ = (basis_enum_of_vec (M *\<^sub>v ( vec_of_basis_enum ( (basis_enum_of_vec (N *\<^sub>v vec_of_basis_enum v))::'b ))))" using c1 a2 by auto finally show ?thesis by simp qed thus ?thesis using b1 b2 b3 by (simp add: cblinfun_eqI scaleC_cblinfun.rep_eq) qed lemma cblinfun_of_mat_adjoint: defines "nA \ length (canonical_basis :: 'a::onb_enum list)" and "nB \ length (canonical_basis :: 'b::onb_enum list)" fixes M:: "complex mat" assumes "M \ carrier_mat nB nA" shows "((cblinfun_of_mat (mat_adjoint M)) :: 'b \\<^sub>C\<^sub>L 'a) = (cblinfun_of_mat M)*" proof (rule adjoint_eqI) show "(cblinfun_of_mat (mat_adjoint M) *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (cblinfun_of_mat M *\<^sub>V y)" for x::'b and y::'a proof- define u where "u = vec_of_basis_enum x" define v where "v = vec_of_basis_enum y" have c1: "vec_of_basis_enum ((cblinfun_of_mat (mat_adjoint M) *\<^sub>V x)::'a) = (mat_adjoint M) *\<^sub>v u" unfolding u_def by (metis (mono_tags, lifting) assms(3) cblinfun_of_mat_inverse map_carrier_mat mat_adjoint_def' mat_of_cblinfun_cblinfun_apply nA_def nB_def transpose_carrier_mat) have c2: "(vec_of_basis_enum ((cblinfun_of_mat M *\<^sub>V y)::'b)) = M *\<^sub>v v" by (metis assms(3) cblinfun_of_mat_inverse mat_of_cblinfun_cblinfun_apply nA_def nB_def v_def) have c3: "dim_vec v = nA" unfolding v_def nA_def vec_of_basis_enum_def by (simp add:) have c4: "dim_vec u = nB" unfolding u_def nB_def vec_of_basis_enum_def by (simp add:) have "v \c ((mat_adjoint M) *\<^sub>v u) = (M *\<^sub>v v) \c u" using c3 c4 cscalar_prod_adjoint assms(3) by blast hence "v \c (vec_of_basis_enum ((cblinfun_of_mat (mat_adjoint M) *\<^sub>V x)::'a)) = (vec_of_basis_enum ((cblinfun_of_mat M *\<^sub>V y)::'b)) \c u" using c1 c2 by simp thus "(cblinfun_of_mat (mat_adjoint M) *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (cblinfun_of_mat M *\<^sub>V y)" unfolding u_def v_def by (simp add: cscalar_prod_vec_of_basis_enum) qed qed +lemma mat_of_cblinfun_compose: + "mat_of_cblinfun (F o\<^sub>C\<^sub>L G) = mat_of_cblinfun F * mat_of_cblinfun G" + for F::"'b::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'c::{basis_enum,complex_normed_vector}" + and G::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b" + by (smt (verit, del_insts) cblinfun_of_mat_inverse mat_carrier mat_of_cblinfun_def mat_of_cblinfun_inverse cblinfun_of_mat_times mult_carrier_mat) + +lemma mat_of_cblinfun_scaleC: + "mat_of_cblinfun ((a::complex) *\<^sub>C F) = a \\<^sub>m (mat_of_cblinfun F)" + for F :: "'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" + by (auto simp add: complex_vector.representation_scale mat_of_cblinfun_def) + +lemma mat_of_cblinfun_scaleR: + "mat_of_cblinfun ((a::real) *\<^sub>R F) = (complex_of_real a) \\<^sub>m (mat_of_cblinfun F)" + unfolding scaleR_scaleC by (rule mat_of_cblinfun_scaleC) + +lemma mat_of_cblinfun_adj: + "mat_of_cblinfun (F*) = mat_adjoint (mat_of_cblinfun F)" + for F :: "'a::onb_enum \\<^sub>C\<^sub>L 'b::onb_enum" + by (metis (no_types, lifting) cblinfun_of_mat_inverse map_carrier_mat mat_adjoint_def' mat_carrier cblinfun_of_mat_adjoint mat_of_cblinfun_def mat_of_cblinfun_inverse transpose_carrier_mat) + +lemma mat_of_cblinfun_vector_to_cblinfun: + "mat_of_cblinfun (vector_to_cblinfun \) + = mat_of_cols (length (canonical_basis :: 'a list)) [vec_of_basis_enum \]" + for \::"'a::{basis_enum,complex_normed_vector}" + by (auto simp: mat_of_cols_Cons_index_0 mat_of_cblinfun_def vec_of_basis_enum_def vec_of_list_index) + +lemma mat_of_cblinfun_proj: + fixes a::"'a::onb_enum" + defines "d \ length (canonical_basis :: 'a list)" + and "norm2 \ (vec_of_basis_enum a) \c (vec_of_basis_enum a)" + shows "mat_of_cblinfun (proj a) = + 1 / norm2 \\<^sub>m (mat_of_cols d [vec_of_basis_enum a] + * mat_of_rows d [conjugate (vec_of_basis_enum a)])" (is \_ = ?rhs\) +proof (cases "a = 0") + case False + have norm2: \norm2 = (norm a)\<^sup>2\ + by (simp add: cscalar_prod_vec_of_basis_enum norm2_def cdot_square_norm[symmetric, simplified]) + have [simp]: \vec_of_basis_enum a \ carrier_vec d\ + by (simp add: carrier_vecI d_def) + + have \mat_of_cblinfun (proj a) = mat_of_cblinfun (proj (a /\<^sub>R norm a))\ + by (metis (mono_tags, opaque_lifting) ccspan_singleton_scaleC complex_vector.scale_eq_0_iff + nonzero_imp_inverse_nonzero norm_eq_zero scaleR_scaleC scale_left_imp_eq) + also have \\ = mat_of_cblinfun (selfbutter (a /\<^sub>R norm a))\ + apply (subst butterfly_eq_proj) + by (auto simp add: False) + also have \\ = 1/norm2 \\<^sub>m mat_of_cblinfun (selfbutter a)\ + apply (simp add: mat_of_cblinfun_scaleC norm2) + by (simp add: inverse_eq_divide power2_eq_square) + also have \\ = 1 / norm2 \\<^sub>m (mat_of_cblinfun (vector_to_cblinfun a :: complex \\<^sub>C\<^sub>L 'a) * mat_adjoint (mat_of_cblinfun (vector_to_cblinfun a :: complex \\<^sub>C\<^sub>L 'a)))\ + by (simp add: butterfly_def mat_of_cblinfun_compose mat_of_cblinfun_adj) + also have \\ = ?rhs\ + by (simp add: mat_of_cblinfun_vector_to_cblinfun mat_adjoint_def flip: d_def) + finally show ?thesis + by - +next + case True + show ?thesis + apply (rule eq_matI) + by (auto simp: True mat_of_cblinfun_zero vec_of_basis_enum_zero scalar_prod_def mat_of_rows_index + simp flip: d_def) +qed + +lemma mat_of_cblinfun_ell2_component: + fixes a :: \'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2\ + assumes [simp]: \i < CARD('b)\ \j < CARD('a)\ + shows \mat_of_cblinfun a $$ (i,j) = Rep_ell2 (a *\<^sub>V ket (Enum.enum ! j)) (Enum.enum ! i)\ +proof - + let ?i = \Enum.enum ! i\ and ?j = \Enum.enum ! j\ and ?aj = \a *\<^sub>V ket (Enum.enum ! j)\ + have \Rep_ell2 ?aj (Enum.enum ! i) = vec_of_basis_enum ?aj $ i\ + by (rule vec_of_basis_enum_ell2_component[symmetric], simp) + also have \\ = (mat_of_cblinfun a *\<^sub>v vec_of_basis_enum (ket (enum_class.enum ! j) :: 'a ell2)) $ i\ + by (simp add: mat_of_cblinfun_cblinfun_apply) + also have \\ = (mat_of_cblinfun a *\<^sub>v unit_vec CARD('a) j) $ i\ + by (simp add: vec_of_basis_enum_ket enum_idx_enum) + also have \\ = mat_of_cblinfun a $$ (i, j)\ + apply (subst mat_entry_explicit[where m=\CARD('b)\]) + by (auto intro!: simp: canonical_basis_length) + finally show ?thesis + by auto +qed + +lemma cblinfun_of_mat_mat: + shows \cblinfun_of_mat (mat (CARD('b)) (CARD('a)) f) = explicit_cblinfun (\(r::'b::enum) (c::'a::enum). f (enum_idx r, enum_idx c))\ + apply (intro equal_ket basis_enum_eq_vec_of_basis_enumI) + by (auto intro!: eq_vecI simp: enum_idx_correct enum_idx_enum vec_of_basis_enum_ket + mat_of_cblinfun_ell2_component canonical_basis_length cblinfun_of_mat_inverse index_row + mat_of_cblinfun_cblinfun_apply) + +lemma mat_of_cblinfun_explicit_cblinfun: + fixes f :: \'a::enum \ 'b::enum \ complex\ + shows \mat_of_cblinfun (explicit_cblinfun f) = mat (CARD('a)) (CARD('b)) (\(r,c). f (Enum.enum!r) (Enum.enum!c))\ + by (auto intro!: cblinfun_of_mat_inj[where 'a=\'b ell2\ and 'b=\'a ell2\, THEN inj_onD] + simp: cblinfun_of_mat_mat canonical_basis_length enum_idx_correct mat_of_cblinfun_inverse) + lemma mat_of_cblinfun_classical_operator: fixes f::"'a::enum \ 'b::enum option" shows "mat_of_cblinfun (classical_operator f) = mat (CARD('b)) (CARD('a)) (\(r,c). if f (Enum.enum!c) = Some (Enum.enum!r) then 1 else 0)" proof - define nA where "nA = CARD('a)" define nB where "nB = CARD('b)" define BasisA where "BasisA = (canonical_basis::'a ell2 list)" define BasisB where "BasisB = (canonical_basis::'b ell2 list)" have "mat_of_cblinfun (classical_operator f) \ carrier_mat nB nA" - unfolding nA_def nB_def by simp + by (auto simp: canonical_basis_length nA_def nB_def) moreover have "nA = CARD ('a)" unfolding nA_def by (simp add:) moreover have "nB = CARD ('b)" unfolding nB_def by (simp add:) ultimately have "mat_of_cblinfun (classical_operator f) \ carrier_mat (CARD('b)) (CARD('a))" unfolding nA_def nB_def by simp moreover have "(mat_of_cblinfun (classical_operator f))$$(r,c) = (mat (CARD('b)) (CARD('a)) (\(r,c). if f (Enum.enum!c) = Some (Enum.enum!r) then 1 else 0))$$(r,c)" if a1: "r < CARD('b)" and a2: "c < CARD('a)" for r c proof- have "CARD('a) = length (enum_class.enum::'a list)" using card_UNIV_length_enum[where 'a = 'a] . hence x1: "BasisA!c = ket ((Enum.enum::'a list)!c)" unfolding BasisA_def using a2 canonical_basis_ell2_def nth_map[where n = c and xs = "Enum.enum::'a list" and f = ket] by metis have cardb: "CARD('b) = length (enum_class.enum::'b list)" using card_UNIV_length_enum[where 'a = 'b] . hence x2: "BasisB!r = ket ((Enum.enum::'b list)!r)" unfolding BasisB_def using a1 canonical_basis_ell2_def nth_map[where n = r and xs = "Enum.enum::'b list" and f = ket] by metis have "inj (map (ket::'b \_))" by (meson injI ket_injective list.inj_map) hence "length (Enum.enum::'b list) = length (map (ket::'b \_) (Enum.enum::'b list))" by simp hence lengthBasisB: "CARD('b) = length BasisB" unfolding BasisB_def canonical_basis_ell2_def using cardb by smt have v1: "(mat_of_cblinfun (classical_operator f))$$(r,c) = 0" if c1: "f (Enum.enum!c) = None" proof- have "(classical_operator f) *\<^sub>V ket (Enum.enum!c) = (case f (Enum.enum!c) of None \ 0 | Some i \ ket i)" using classical_operator_ket_finite . also have "\ = 0" using c1 by simp finally have "(classical_operator f) *\<^sub>V ket (Enum.enum!c) = 0" . hence *: "(classical_operator f) *\<^sub>V BasisA!c = 0" using x1 by simp hence "is_orthogonal (BasisB!r) (classical_operator f *\<^sub>V BasisA!c)" by simp thus ?thesis unfolding mat_of_cblinfun_def BasisA_def BasisB_def by (smt (verit, del_insts) BasisA_def * a1 a2 canonical_basis_length_ell2 complex_vector.representation_zero index_mat(1) old.prod.case) qed have v2: "(mat_of_cblinfun (classical_operator f))$$(r,c) = 0" if c1: "f (Enum.enum!c) = Some (Enum.enum!r')" and c2: "r\r'" and c3: "r' < CARD('b)" for r' proof- have x3: "BasisB!r' = ket ((Enum.enum::'b list)!r')" unfolding BasisB_def using cardb c3 canonical_basis_ell2_def nth_map[where n = r' and xs = "Enum.enum::'b list" and f = ket] by smt have "distinct BasisB" unfolding BasisB_def by simp moreover have "r < length BasisB" using a1 lengthBasisB by simp moreover have "r' < length BasisB" using c3 lengthBasisB by simp ultimately have h1: "BasisB!r \ BasisB!r'" using nth_eq_iff_index_eq[where xs = BasisB and i = r and j = r'] c2 by blast have "(classical_operator f) *\<^sub>V ket (Enum.enum!c) = (case f (Enum.enum!c) of None \ 0 | Some i \ ket i)" using classical_operator_ket_finite . also have "\ = ket (Enum.enum!r')" using c1 by simp finally have "(classical_operator f) *\<^sub>V ket (Enum.enum!c) = ket (Enum.enum!r')" . hence *: "(classical_operator f) *\<^sub>V BasisA!c = BasisB!r'" using x1 x3 by simp moreover have "is_orthogonal (BasisB!r) (BasisB!r')" using h1 using BasisB_def \r < length BasisB\ \r' < length BasisB\ is_ortho_set_def is_orthonormal nth_mem by metis ultimately have "is_orthogonal (BasisB!r) (classical_operator f *\<^sub>V BasisA!c)" by simp thus ?thesis unfolding mat_of_cblinfun_def BasisA_def BasisB_def by (smt (z3) BasisA_def BasisB_def * \r < length BasisB\ \r' < length BasisB\ a2 canonical_basis_length_ell2 case_prod_conv complex_vector.representation_basis h1 index_mat(1) is_cindependent_set nth_mem) qed have "(mat_of_cblinfun (classical_operator f))$$(r,c) = 0" if b1: "f (Enum.enum!c) \ Some (Enum.enum!r)" proof (cases "f (Enum.enum!c) = None") case True thus ?thesis using v1 by blast next case False hence "\R. f (Enum.enum!c) = Some R" apply (induction "f (Enum.enum!c)") apply simp by simp then obtain R where R0: "f (Enum.enum!c) = Some R" by blast have "R \ set (Enum.enum::'b list)" using UNIV_enum by blast hence "\r'. R = (Enum.enum::'b list)!r' \ r' < length (Enum.enum::'b list)" by (metis in_set_conv_nth) then obtain r' where u1: "R = (Enum.enum::'b list)!r'" and u2: "r' < length (Enum.enum::'b list)" by blast have R1: "f (Enum.enum!c) = Some (Enum.enum!r')" using R0 u1 by blast have "Some ((Enum.enum::'b list)!r') \ Some ((Enum.enum::'b list)!r)" proof(rule classical) assume "\(Some ((Enum.enum::'b list)!r') \ Some ((Enum.enum::'b list)!r))" hence "Some ((Enum.enum::'b list)!r') = Some ((Enum.enum::'b list)!r)" by blast hence "f (Enum.enum!c) = Some ((Enum.enum::'b list)!r)" using R1 by auto thus ?thesis using b1 by blast qed hence "((Enum.enum::'b list)!r') \ ((Enum.enum::'b list)!r)" by simp hence "r' \ r" by blast moreover have "r' < CARD('b)" using u2 cardb by simp ultimately show ?thesis using R1 v2[where r' = r'] by blast qed moreover have "(mat_of_cblinfun (classical_operator f))$$(r,c) = 1" if b1: "f (Enum.enum!c) = Some (Enum.enum!r)" proof- have "CARD('b) = length (enum_class.enum::'b list)" using card_UNIV_length_enum[where 'a = 'b]. hence "length (map (ket::'b\_) enum_class.enum) = CARD('b)" by simp hence w0: "map (ket::'b\_) enum_class.enum ! r = ket (enum_class.enum ! r)" by (simp add: a1) have "CARD('a) = length (enum_class.enum::'a list)" using card_UNIV_length_enum[where 'a = 'a]. hence "length (map (ket::'a\_) enum_class.enum) = CARD('a)" by simp hence "map (ket::'a\_) enum_class.enum ! c = ket (enum_class.enum ! c)" by (simp add: a2) hence "(classical_operator f) *\<^sub>V (BasisA!c) = (classical_operator f) *\<^sub>V (ket (Enum.enum!c))" unfolding BasisA_def canonical_basis_ell2_def by simp also have "... = (case f (enum_class.enum ! c) of None \ 0 | Some x \ ket x)" by (rule classical_operator_ket_finite) also have "\ = BasisB!r" using w0 b1 by (simp add: BasisB_def canonical_basis_ell2_def) finally have w1: "(classical_operator f) *\<^sub>V (BasisA!c) = BasisB!r" by simp have "(mat_of_cblinfun (classical_operator f))$$(r,c) = (BasisB!r) \\<^sub>C (classical_operator f *\<^sub>V (BasisA!c))" unfolding BasisB_def BasisA_def mat_of_cblinfun_def using \nA = CARD('a)\ \nB = CARD('b)\ a1 a2 nA_def nB_def apply auto by (metis BasisA_def BasisB_def canonical_basis_length_ell2 cinner_canonical_basis complex_vector.representation_basis is_cindependent_set nth_mem w1) also have "\ = (BasisB!r) \\<^sub>C (BasisB!r)" using w1 by simp also have "\ = 1" unfolding BasisB_def using \nB = CARD('b)\ a1 nB_def by (simp add: cinner_canonical_basis) finally show ?thesis by blast qed ultimately show ?thesis by (simp add: a1 a2) qed ultimately show ?thesis apply (rule_tac eq_matI) by auto qed -lemma mat_of_cblinfun_compose: - "mat_of_cblinfun (F o\<^sub>C\<^sub>L G) = mat_of_cblinfun F * mat_of_cblinfun G" - for F::"'b::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'c::{basis_enum,complex_normed_vector}" - and G::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b" - by (smt (verit, del_insts) cblinfun_of_mat_inverse mat_carrier mat_of_cblinfun_def mat_of_cblinfun_inverse cblinfun_of_mat_times mult_carrier_mat) - -lemma mat_of_cblinfun_scaleC: - "mat_of_cblinfun ((a::complex) *\<^sub>C F) = a \\<^sub>m (mat_of_cblinfun F)" - for F :: "'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::{basis_enum,complex_normed_vector}" - by (auto simp add: complex_vector.representation_scale mat_of_cblinfun_def) - -lemma mat_of_cblinfun_scaleR: - "mat_of_cblinfun ((a::real) *\<^sub>R F) = (complex_of_real a) \\<^sub>m (mat_of_cblinfun F)" - unfolding scaleR_scaleC by (rule mat_of_cblinfun_scaleC) - -lemma mat_of_cblinfun_adj: - "mat_of_cblinfun (F*) = mat_adjoint (mat_of_cblinfun F)" - for F :: "'a::onb_enum \\<^sub>C\<^sub>L 'b::onb_enum" - by (metis (no_types, lifting) cblinfun_of_mat_inverse map_carrier_mat mat_adjoint_def' mat_carrier cblinfun_of_mat_adjoint mat_of_cblinfun_def mat_of_cblinfun_inverse transpose_carrier_mat) - -lemma mat_of_cblinfun_vector_to_cblinfun: - "mat_of_cblinfun (vector_to_cblinfun \) - = mat_of_cols (length (canonical_basis :: 'a list)) [vec_of_basis_enum \]" - for \::"'a::{basis_enum,complex_normed_vector}" - by (auto simp: mat_of_cols_Cons_index_0 mat_of_cblinfun_def vec_of_basis_enum_def vec_of_list_index) - -lemma mat_of_cblinfun_proj: - fixes a::"'a::onb_enum" - defines "d \ length (canonical_basis :: 'a list)" - and "norm2 \ (vec_of_basis_enum a) \c (vec_of_basis_enum a)" - shows "mat_of_cblinfun (proj a) = - 1 / norm2 \\<^sub>m (mat_of_cols d [vec_of_basis_enum a] - * mat_of_rows d [conjugate (vec_of_basis_enum a)])" (is \_ = ?rhs\) -proof (cases "a = 0") - case False - have norm2: \norm2 = (norm a)\<^sup>2\ - by (simp add: cscalar_prod_vec_of_basis_enum norm2_def cdot_square_norm[symmetric, simplified]) - have [simp]: \vec_of_basis_enum a \ carrier_vec d\ - by (simp add: carrier_vecI d_def) - - have \mat_of_cblinfun (proj a) = mat_of_cblinfun (proj (a /\<^sub>R norm a))\ - by (metis (mono_tags, opaque_lifting) ccspan_singleton_scaleC complex_vector.scale_eq_0_iff - nonzero_imp_inverse_nonzero norm_eq_zero scaleR_scaleC scale_left_imp_eq) - also have \\ = mat_of_cblinfun (selfbutter (a /\<^sub>R norm a))\ - apply (subst butterfly_eq_proj) - by (auto simp add: False) - also have \\ = 1/norm2 \\<^sub>m mat_of_cblinfun (selfbutter a)\ - apply (simp add: mat_of_cblinfun_scaleC norm2) - by (simp add: inverse_eq_divide power2_eq_square) - also have \\ = 1 / norm2 \\<^sub>m (mat_of_cblinfun (vector_to_cblinfun a :: complex \\<^sub>C\<^sub>L 'a) * mat_adjoint (mat_of_cblinfun (vector_to_cblinfun a :: complex \\<^sub>C\<^sub>L 'a)))\ - by (simp add: butterfly_def mat_of_cblinfun_compose mat_of_cblinfun_adj) - also have \\ = ?rhs\ - by (simp add: mat_of_cblinfun_vector_to_cblinfun mat_adjoint_def flip: d_def) - finally show ?thesis - by - -next - case True - show ?thesis - apply (rule eq_matI) - by (auto simp: True mat_of_cblinfun_zero vec_of_basis_enum_zero scalar_prod_def mat_of_rows_index - simp flip: d_def) -qed - -lemma mat_of_cblinfun_ell2_component: - fixes a :: \'a::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2\ - assumes [simp]: \i < CARD('b)\ \j < CARD('a)\ - shows \mat_of_cblinfun a $$ (i,j) = Rep_ell2 (a *\<^sub>V ket (Enum.enum ! j)) (Enum.enum ! i)\ -proof - - let ?i = \Enum.enum ! i\ and ?j = \Enum.enum ! j\ and ?aj = \a *\<^sub>V ket (Enum.enum ! j)\ - have \Rep_ell2 ?aj (Enum.enum ! i) = vec_of_basis_enum ?aj $ i\ - by (rule vec_of_basis_enum_ell2_component[symmetric], simp) - also have \\ = (mat_of_cblinfun a *\<^sub>v vec_of_basis_enum (ket (enum_class.enum ! j) :: 'a ell2)) $ i\ - by (simp add: mat_of_cblinfun_cblinfun_apply) - also have \\ = (mat_of_cblinfun a *\<^sub>v unit_vec CARD('a) j) $ i\ - by (simp add: vec_of_basis_enum_ket enum_idx_enum) - also have \\ = mat_of_cblinfun a $$ (i, j)\ - apply (subst mat_entry_explicit[where m=\CARD('b)\]) - by auto - finally show ?thesis - by auto -qed lemma mat_of_cblinfun_sandwich: fixes a :: "(_::onb_enum, _::onb_enum) cblinfun" shows \mat_of_cblinfun (sandwich a *\<^sub>V b) = (let a' = mat_of_cblinfun a in a' * mat_of_cblinfun b * mat_adjoint a')\ by (simp add: mat_of_cblinfun_compose sandwich_apply Let_def mat_of_cblinfun_adj) subsection \Operations on subspaces\ lemma ccspan_gram_schmidt0_invariant: defines "basis_enum \ (basis_enum_of_vec :: _ \ 'a::{basis_enum,complex_normed_vector})" defines "n \ length (canonical_basis :: 'a list)" assumes "set ws \ carrier_vec n" shows "ccspan (set (map basis_enum (gram_schmidt0 n ws))) = ccspan (set (map basis_enum ws))" proof (transfer fixing: n ws basis_enum) interpret complex_vec_space. define gs where "gs = gram_schmidt0 n ws" have "closure (cspan (set (map basis_enum gs))) = cspan (set (map basis_enum gs))" apply (rule closure_finite_cspan) by simp also have "\ = cspan (basis_enum ` set gs)" by simp also have "\ = basis_enum ` span (set gs)" unfolding basis_enum_def apply (rule basis_enum_of_vec_span[symmetric]) using n_def apply simp by (simp add: assms(3) cof_vec_space.gram_schmidt0_result(1) gs_def) also have "\ = basis_enum ` span (set ws)" unfolding gs_def apply (subst gram_schmidt0_result(4)[where ws=ws, symmetric]) using assms by auto also have "\ = cspan (basis_enum ` set ws)" unfolding basis_enum_def apply (rule basis_enum_of_vec_span) using n_def apply simp by (simp add: assms(3)) also have "\ = cspan (set (map basis_enum ws))" by simp also have "\ = closure (cspan (set (map basis_enum ws)))" apply (rule closure_finite_cspan[symmetric]) by simp finally show "closure (cspan (set (map basis_enum gs))) = closure (cspan (set (map basis_enum ws)))". qed definition "is_subspace_of_vec_list n vs ws = (let ws' = gram_schmidt0 n ws in \v\set vs. adjuster n v ws' = - v)" lemma ccspan_leq_using_vec: - fixes A B :: "'a::{basis_enum,complex_normed_vector} list" - shows "(ccspan (set A) \ ccspan (set B)) \ + fixes A B :: \'a::{basis_enum,complex_normed_vector} list\ + shows \(ccspan (set A) \ ccspan (set B)) \ is_subspace_of_vec_list (length (canonical_basis :: 'a list)) - (map vec_of_basis_enum A) (map vec_of_basis_enum B)" + (map vec_of_basis_enum A) (map vec_of_basis_enum B)\ proof - define d Av Bv Bo where "d = length (canonical_basis :: 'a list)" and "Av = map vec_of_basis_enum A" and "Bv = map vec_of_basis_enum B" and "Bo = gram_schmidt0 d Bv" interpret complex_vec_space d. have Av_carrier: "set Av \ carrier_vec d" unfolding Av_def apply auto by (simp add: carrier_vecI d_def dim_vec_of_basis_enum') have Bv_carrier: "set Bv \ carrier_vec d" unfolding Bv_def apply auto by (simp add: carrier_vecI d_def dim_vec_of_basis_enum') have Bo_carrier: "set Bo \ carrier_vec d" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(1)) have orth_Bo: "corthogonal Bo" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(3)) have distinct_Bo: "distinct Bo" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(2)) have "ccspan (set A) \ ccspan (set B) \ cspan (set A) \ cspan (set B)" apply (transfer fixing: A B) apply (subst closure_finite_cspan, simp) by (subst closure_finite_cspan, simp_all) also have "\ \ span (set Av) \ span (set Bv)" apply (simp add: Av_def Bv_def) apply (subst vec_of_basis_enum_cspan[symmetric], simp add: d_def) apply (subst vec_of_basis_enum_cspan[symmetric], simp add: d_def) by (metis inj_image_subset_iff inj_on_def vec_of_basis_enum_inverse) also have "\ \ span (set Av) \ span (set Bo)" unfolding Bo_def Av_def Bv_def apply (subst gram_schmidt0_result(4)[symmetric]) by (simp_all add: carrier_dim_vec d_def dim_vec_of_basis_enum' image_subset_iff) also have "\ \ (\v\set Av. adjuster d v Bo = - v)" proof (intro iffI ballI) fix v assume "v \ set Av" and "span (set Av) \ span (set Bo)" then have "v \ span (set Bo)" using Av_carrier span_mem by auto have "adjuster d v Bo + v = 0\<^sub>v d" apply (rule adjuster_already_in_span) using Av_carrier \v \ set Av\ Bo_carrier orth_Bo \v \ span (set Bo)\ by simp_all then show "adjuster d v Bo = - v" using Av_carrier Bo_carrier \v \ set Av\ by (metis (no_types, lifting) add.inv_equality adjuster_carrier' local.vec_neg subsetD) next fix v assume adj_minusv: "\v\set Av. adjuster d v Bo = - v" have *: "adjuster d v Bo \ span (set Bo)" if "v \ set Av" for v apply (rule adjuster_in_span) using Bo_carrier that Av_carrier distinct_Bo by auto have "v \ span (set Bo)" if "v \ set Av" for v using *[OF that] adj_minusv[rule_format, OF that] apply auto by (metis (no_types, lifting) Av_carrier Bo_carrier adjust_nonzero distinct_Bo subsetD that uminus_l_inv_vec) then show "span (set Av) \ span (set Bo)" apply auto by (meson Bo_carrier in_mono span_subsetI subsetI) qed also have "\ = is_subspace_of_vec_list d Av Bv" by (simp add: is_subspace_of_vec_list_def d_def Bo_def) finally show "ccspan (set A) \ ccspan (set B) \ is_subspace_of_vec_list d Av Bv" by simp qed -lemma cblinfun_apply_ccspan_using_vec: +lemma cblinfun_image_ccspan_using_vec: "A *\<^sub>S ccspan (set S) = ccspan (basis_enum_of_vec ` set (map ((*\<^sub>v) (mat_of_cblinfun A)) (map vec_of_basis_enum S)))" apply (auto simp: cblinfun_image_ccspan image_image) by (metis mat_of_cblinfun_cblinfun_apply vec_of_basis_enum_inverse) text \\<^term>\mk_projector_orthog d L\ takes a list L of d-dimensional vectors and returns the projector onto the span of L. (Assuming that all vectors in L are orthogonal and nonzero.)\ fun mk_projector_orthog :: "nat \ complex vec list \ complex mat" where "mk_projector_orthog d [] = zero_mat d d" | "mk_projector_orthog d [v] = (let norm2 = cscalar_prod v v in smult_mat (1/norm2) (mat_of_cols d [v] * mat_of_rows d [conjugate v]))" | "mk_projector_orthog d (v#vs) = (let norm2 = cscalar_prod v v in smult_mat (1/norm2) (mat_of_cols d [v] * mat_of_rows d [conjugate v]) + mk_projector_orthog d vs)" lemma mk_projector_orthog_correct: fixes S :: "'a::onb_enum list" defines "d \ length (canonical_basis :: 'a list)" assumes ortho: "is_ortho_set (set S)" and distinct: "distinct S" shows "mk_projector_orthog d (map vec_of_basis_enum S) = mat_of_cblinfun (Proj (ccspan (set S)))" proof - define Snorm where "Snorm = map (\s. s /\<^sub>R norm s) S" have "distinct Snorm" proof (insert ortho distinct, unfold Snorm_def, induction S) case Nil show ?case by simp next case (Cons s S) then have "is_ortho_set (set S)" and "distinct S" unfolding is_ortho_set_def by auto note IH = Cons.IH[OF this] have "s /\<^sub>R norm s \ (\s. s /\<^sub>R norm s) ` set S" proof auto fix s' assume "s' \ set S" and same: "s /\<^sub>R norm s = s' /\<^sub>R norm s'" with Cons.prems have "s \ s'" by auto have "s \ 0" by (metis Cons.prems(1) is_ortho_set_def list.set_intros(1)) then have "0 \ (s /\<^sub>R norm s) \\<^sub>C (s /\<^sub>R norm s)" by simp also have \(s /\<^sub>R norm s) \\<^sub>C (s /\<^sub>R norm s) = (s /\<^sub>R norm s) \\<^sub>C (s' /\<^sub>R norm s')\ by (simp add: same) also have \(s /\<^sub>R norm s) \\<^sub>C (s' /\<^sub>R norm s') = (s \\<^sub>C s') / (norm s * norm s')\ by (simp add: scaleR_scaleC divide_inverse_commute) also from \s' \ set S\ \s \ s'\ have "\ = 0" using Cons.prems unfolding is_ortho_set_def by simp finally show False by simp qed then show ?case using IH by simp qed have norm_Snorm: "norm s = 1" if "s \ set Snorm" for s using that ortho unfolding Snorm_def is_ortho_set_def apply auto by (metis left_inverse norm_eq_zero) have ortho_Snorm: "is_ortho_set (set Snorm)" unfolding is_ortho_set_def proof (intro conjI ballI impI) fix x y show "0 \ set Snorm" using norm_Snorm[of 0] by auto assume "x \ set Snorm" and "y \ set Snorm" and "x \ y" from \x \ set Snorm\ obtain x' where x: "x = x' /\<^sub>R norm x'" and x': "x' \ set S" unfolding Snorm_def by auto from \y \ set Snorm\ obtain y' where y: "y = y' /\<^sub>R norm y'" and y': "y' \ set S" unfolding Snorm_def by auto from \x \ y\ x y have \x' \ y'\ by auto with x' y' ortho have "cinner x' y' = 0" unfolding is_ortho_set_def by auto then show "cinner x y = 0" unfolding x y scaleR_scaleC by auto qed have inj_butter: "inj_on selfbutter (set Snorm)" proof (rule inj_onI) fix x y assume "x \ set Snorm" and "y \ set Snorm" assume "selfbutter x = selfbutter y" then obtain c where xcy: "x = c *\<^sub>C y" and "cmod c = 1" using inj_selfbutter_upto_phase by auto have "0 \ cmod (cinner x x)" using \x \ set Snorm\ norm_Snorm by force also have "cmod (cinner x x) = cmod (c * (x \\<^sub>C y))" apply (subst (2) xcy) by simp also have "\ = cmod (x \\<^sub>C y)" using \cmod c = 1\ by (simp add: norm_mult) finally have "(x \\<^sub>C y) \ 0" by simp then show "x = y" using ortho_Snorm \x \ set Snorm\ \y \ set Snorm\ unfolding is_ortho_set_def by auto qed from \distinct Snorm\ inj_butter have distinct': "distinct (map selfbutter Snorm)" unfolding distinct_map by simp have Span_Snorm: "ccspan (set Snorm) = ccspan (set S)" apply (transfer fixing: Snorm S) apply (simp add: scaleR_scaleC Snorm_def) apply (subst complex_vector.span_image_scale) using is_ortho_set_def ortho by fastforce+ have "mk_projector_orthog d (map vec_of_basis_enum S) = mat_of_cblinfun (sum_list (map selfbutter Snorm))" unfolding Snorm_def proof (induction S) case Nil show ?case by (simp add: d_def mat_of_cblinfun_zero) next case (Cons a S) define sumS where "sumS = sum_list (map selfbutter (map (\s. s /\<^sub>R norm s) S))" with Cons have IH: "mk_projector_orthog d (map vec_of_basis_enum S) = mat_of_cblinfun sumS" by simp define factor where "factor = inverse ((complex_of_real (norm a))\<^sup>2)" have factor': "factor = 1 / (vec_of_basis_enum a \c vec_of_basis_enum a)" unfolding factor_def cscalar_prod_vec_of_basis_enum by (simp add: inverse_eq_divide power2_norm_eq_cinner) have "mk_projector_orthog d (map vec_of_basis_enum (a # S)) = factor \\<^sub>m (mat_of_cols d [vec_of_basis_enum a] * mat_of_rows d [conjugate (vec_of_basis_enum a)]) + mat_of_cblinfun sumS" apply (cases S) apply (auto simp add: factor' sumS_def d_def mat_of_cblinfun_zero)[1] by (auto simp add: IH[symmetric] factor' d_def) also have "\ = factor \\<^sub>m (mat_of_cols d [vec_of_basis_enum a] * mat_adjoint (mat_of_cols d [vec_of_basis_enum a])) + mat_of_cblinfun sumS" apply (rule arg_cong[where f="\x. _ \\<^sub>m (_ * x) + _"]) apply (rule mat_eq_iff[THEN iffD2]) apply (auto simp add: mat_adjoint_def) apply (subst mat_of_rows_index) apply auto apply (subst mat_of_rows_index) apply auto apply (subst mat_of_cols_index) apply auto by (simp add: assms(1) dim_vec_of_basis_enum') also have "\ = mat_of_cblinfun (selfbutter (a /\<^sub>R norm a)) + mat_of_cblinfun sumS" apply (simp add: butterfly_scaleR_left butterfly_scaleR_right power_inverse mat_of_cblinfun_scaleR factor_def) apply (simp add: butterfly_def mat_of_cblinfun_compose mat_of_cblinfun_adj mat_of_cblinfun_vector_to_cblinfun d_def) by (simp add: mat_of_cblinfun_compose mat_of_cblinfun_adj mat_of_cblinfun_vector_to_cblinfun mat_of_cblinfun_scaleC power2_eq_square) finally show ?case by (simp add: mat_of_cblinfun_plus sumS_def) qed also have "\ = mat_of_cblinfun (\s\set Snorm. selfbutter s)" by (metis distinct' distinct_map sum.distinct_set_conv_list) also have "\ = mat_of_cblinfun (\s\set Snorm. proj s)" apply (rule arg_cong[where f="mat_of_cblinfun"]) apply (rule sum.cong[OF refl]) apply (rule butterfly_eq_proj) using norm_Snorm by simp also have "\ = mat_of_cblinfun (Proj (ccspan (set Snorm)))" apply (rule arg_cong[where f=mat_of_cblinfun]) using ortho_Snorm \distinct Snorm\ apply (induction Snorm) apply auto apply (subst Proj_orthog_ccspan_insert[where Y=\set _\]) by (auto simp: is_ortho_set_def) also have "\ = mat_of_cblinfun (Proj (ccspan (set S)))" unfolding Span_Snorm by simp finally show ?thesis by - qed +definition \mk_projector d vs = mk_projector_orthog d (gram_schmidt0 d vs)\ + lemma mat_of_cblinfun_Proj_ccspan: - fixes S :: "'a::onb_enum list" - shows "mat_of_cblinfun (Proj (ccspan (set S))) = - (let d = length (canonical_basis :: 'a list) in - mk_projector_orthog d (gram_schmidt0 d (map vec_of_basis_enum S)))" + fixes S :: \'a::onb_enum list\ + shows \mat_of_cblinfun (Proj (ccspan (set S))) = mk_projector (length (canonical_basis :: 'a list)) (map vec_of_basis_enum S)\ proof- define d gs where "d = length (canonical_basis :: 'a list)" and "gs = gram_schmidt0 d (map vec_of_basis_enum S)" interpret complex_vec_space d. have gs_dim: "x \ set gs \ dim_vec x = d" for x by (smt carrier_vecD carrier_vec_dim_vec d_def dim_vec_of_basis_enum' ex_map_conv gram_schmidt0_result(1) gs_def subset_code(1)) have ortho_gs: "is_ortho_set (set (map basis_enum_of_vec gs :: 'a list))" apply (subst corthogonal_vec_of_basis_enum[THEN iffD1], auto) by (smt carrier_dim_vec cof_vec_space.gram_schmidt0_result(1) d_def dim_vec_of_basis_enum' gram_schmidt0_result(3) gs_def imageE map_idI map_map o_apply set_map subset_code(1) basis_enum_of_vec_inverse) have distinct_gs: "distinct (map basis_enum_of_vec gs :: 'a list)" by (metis (mono_tags, opaque_lifting) carrier_vec_dim_vec cof_vec_space.gram_schmidt0_result(2) d_def dim_vec_of_basis_enum' distinct_map gs_def gs_dim image_iff inj_on_inverseI set_map subsetI basis_enum_of_vec_inverse) have "mk_projector_orthog d gs = mk_projector_orthog d (map vec_of_basis_enum (map basis_enum_of_vec gs :: 'a list))" apply simp apply (subst map_cong[where ys=gs and g=id], simp) using gs_dim by (auto intro!: vec_of_basis_enum_inverse simp: d_def) also have "\ = mat_of_cblinfun (Proj (ccspan (set (map basis_enum_of_vec gs :: 'a list))))" unfolding d_def apply (subst mk_projector_orthog_correct) using ortho_gs distinct_gs by auto also have "\ = mat_of_cblinfun (Proj (ccspan (set S)))" apply (rule arg_cong[where f="\x. mat_of_cblinfun (Proj x)"]) unfolding gs_def d_def apply (subst ccspan_gram_schmidt0_invariant) by (auto simp add: carrier_vecI dim_vec_of_basis_enum') finally show ?thesis - unfolding d_def gs_def by auto + by (simp add: d_def gs_def mk_projector_def) qed unbundle no_jnf_notation unbundle no_cblinfun_notation 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,5062 +1,5053 @@ 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 "HOL-Types_To_Sets.Types_To_Sets" Banach_Steinhaus.Banach_Steinhaus Complex_Inner_Product One_Dimensional_Spaces Complex_Bounded_Linear_Function0 "HOL-Library.Function_Algebras" 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 \ = \" by simp lemma apply_id_cblinfun[simp]: \(*\<^sub>V) id_cblinfun = id\ by auto lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" by transfer (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 by transfer (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\ using bounded_clinear.axioms(1) cblinfun_apply assms complex_vector.linear_eq_0_on_span by blast 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\ using bounded_clinear.axioms(1) cblinfun_apply assms complex_vector.linear_eq_on_span by blast 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\ by (metis (no_types) assms cblinfun_eqI cblinfun_eq_on_span iso_tuple_UNIV_I) 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" using assms cblinfun_eq_on_UNIV_span is_generator_set by blast 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 \ is_orthogonal v (F *\<^sub>V u)" shows "F = 0" proof- have "F *\<^sub>V u = 0" if "u\basisA" for u proof- have "\v. v\basisB \ is_orthogonal v (F *\<^sub>V u)" by (simp add: assms(3) that) moreover have "(\v. v\basisB \ is_orthogonal v x) \ x = 0" for x proof- assume r1: "\v. v\basisB \ is_orthogonal v x" have "is_orthogonal v x" 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 \\<^sub>C x = (\a\t. s a *\<^sub>C a) \\<^sub>C x" by (simp add: b1) also have "\ = (\a\t. (s a *\<^sub>C a) \\<^sub>C x)" using cinner_sum_left by blast also have "\ = (\a\t. cnj (s a) * (a \\<^sub>C 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 \\<^sub>C 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 \\<^sub>C (F *\<^sub>V u) = v \\<^sub>C (G *\<^sub>V u)" shows "F = G" proof- define H where "H = F - G" have "\u v. u\basisA \ v\basisB \ is_orthogonal v (H *\<^sub>V u)" 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) \\<^sub>C v = (G *\<^sub>V u) \\<^sub>C v" shows "F = G" using cinner_canonical_basis_eq assms by (metis cinner_commute') lemma not_not_singleton_cblinfun_zero: \x = 0\ if \\ class.not_singleton TYPE('a)\ for x :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ apply (rule cblinfun_eqI) apply (subst not_not_singleton_zero[OF that]) by simp 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: 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 by auto (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_norm_approx_witness': fixes A :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ > 0\ shows \\\. norm (A *\<^sub>V \) / norm \ \ norm A - \\ proof (cases \class.not_singleton TYPE('a)\) case True obtain \ where \norm (A *\<^sub>V \) \ norm A - \\ and \norm \ = 1\ apply atomize_elim using complex_normed_vector_axioms True assms by (rule cblinfun_norm_approx_witness[internalize_sort' 'a]) then have \norm (A *\<^sub>V \) / norm \ \ norm A - \\ by simp then show ?thesis by auto next case False show ?thesis apply (subst not_not_singleton_cblinfun_zero[OF False]) apply simp apply (subst not_not_singleton_cblinfun_zero[OF False]) using \\ > 0\ by simp qed lemma cblinfun_to_CARD_1_0[simp]: \(A :: _ \\<^sub>C\<^sub>L _::CARD_1) = 0\ by (simp add: cblinfun_eqI) lemma cblinfun_from_CARD_1_0[simp]: \(A :: _::CARD_1 \\<^sub>C\<^sub>L _) = 0\ using cblinfun_eq_on_UNIV_span by force 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 by atomize_elim 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) complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset mem_Collect_eq subsetD) moreover have \G *\<^sub>V a' = (if a'=a then b else 0)\ if \a'\basisA'\ for 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\ by (auto simp: H'_def cblinfun_eq_on_UNIV_span cblinfun.sum_left 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})\ by (smt (verit) F_basis UN_subset_iff complex_vector.span_base complex_vector.span_minimal complex_vector.subspace_span mem_Collect_eq subsetI) 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 "\a b. \a \ basisA; b \ basisB\ \ \F\basis. \a'\basisA. F *\<^sub>V a' = (if a' = a then b else 0)" by (smt (verit, del_insts) F_apply basis_def f_a f_not_a mem_Collect_eq) then have \cspan basis = UNIV\ by (metis \cspan basisA = UNIV\ \cspan basisB = UNIV\ cblinfun_cspan_UNIV) moreover have \finite basis\ unfolding basis_def by (auto intro: finite_image_set2) ultimately show \\S :: ('a \\<^sub>C\<^sub>L 'b) set. finite S \ cspan S = UNIV\ by auto qed lemma norm_cblinfun_bound_dense: assumes \0 \ b\ assumes S: \closure S = UNIV\ assumes bound: \\x. x\S \ norm (cblinfun_apply f x) \ b * norm x\ shows \norm f \ b\ proof - have 1: \continuous_on UNIV (\a. norm (f *\<^sub>V a))\ by (simp add: continuous_on_eq_continuous_within) have 2: \continuous_on UNIV (\a. b * norm a)\ using continuous_on_mult_left continuous_on_norm_id by blast have \norm (cblinfun_apply f x) \ b * norm x\ for x by (metis (mono_tags, lifting) UNIV_I S bound 1 2 on_closure_leI) then show \norm f \ b\ using \0 \ b\ norm_cblinfun_bound by blast qed lemma infsum_cblinfun_apply: assumes \g summable_on S\ shows \infsum (\x. A *\<^sub>V g x) S = A *\<^sub>V (infsum g S)\ using infsum_bounded_linear[unfolded o_def] assms cblinfun.real.bounded_linear_right by blast lemma has_sum_cblinfun_apply: assumes \(g has_sum x) S\ shows \((\x. A *\<^sub>V g x) has_sum (A *\<^sub>V x)) S\ using assms has_sum_bounded_linear[unfolded o_def] using cblinfun.real.bounded_linear_right by blast lemma abs_summable_on_cblinfun_apply: assumes \g abs_summable_on S\ shows \(\x. A *\<^sub>V g x) abs_summable_on S\ using bounded_clinear.bounded_linear[OF cblinfun.bounded_clinear_right] assms by (rule abs_summable_on_bounded_linear[unfolded o_def]) lemma summable_on_cblinfun_apply: assumes \g summable_on S\ shows \(\x. A *\<^sub>V g x) summable_on S\ using bounded_clinear.bounded_linear[OF cblinfun.bounded_clinear_right] assms by (rule summable_on_bounded_linear[unfolded o_def]) lemma summable_on_cblinfun_apply_left: assumes \A summable_on S\ shows \(\x. A x *\<^sub>V g) summable_on S\ using bounded_clinear.bounded_linear[OF cblinfun.bounded_clinear_left] assms by (rule summable_on_bounded_linear[unfolded o_def]) lemma abs_summable_on_cblinfun_apply_left: assumes \A abs_summable_on S\ shows \(\x. A x *\<^sub>V g) abs_summable_on S\ using bounded_clinear.bounded_linear[OF cblinfun.bounded_clinear_left] assms by (rule abs_summable_on_bounded_linear[unfolded o_def]) lemma infsum_cblinfun_apply_left: assumes \A summable_on S\ shows \infsum (\x. A x *\<^sub>V g) S = (infsum A S) *\<^sub>V g\ apply (rule infsum_bounded_linear[unfolded o_def, of \\A. cblinfun_apply A g\]) using assms by (auto simp add: bounded_clinear.bounded_linear bounded_cbilinear_cblinfun_apply) lemma has_sum_cblinfun_apply_left: assumes \(A has_sum x) S\ shows \((\x. A x *\<^sub>V g) has_sum (x *\<^sub>V g)) S\ apply (rule has_sum_bounded_linear[unfolded o_def, of \\A. cblinfun_apply A g\]) using assms by (auto simp add: bounded_clinear.bounded_linear cblinfun.bounded_clinear_left) text \The next eight lemmas logically belong in \<^theory>\Complex_Bounded_Operators.Complex_Inner_Product\ but the proofs use facts from this theory.\ lemma has_sum_cinner_left: assumes \(f has_sum x) I\ shows \((\i. cinner a (f i)) has_sum cinner a x) I\ by (metis assms cblinfun_cinner_right.rep_eq has_sum_cblinfun_apply) lemma summable_on_cinner_left: assumes \f summable_on I\ shows \(\i. cinner a (f i)) summable_on I\ by (metis assms has_sum_cinner_left summable_on_def) lemma infsum_cinner_left: assumes \\ summable_on I\ shows \cinner \ (\\<^sub>\i\I. \ i) = (\\<^sub>\i\I. cinner \ (\ i))\ by (metis assms has_sum_cinner_left has_sum_infsum infsumI) lemma has_sum_cinner_right: assumes \(f has_sum x) I\ shows \((\i. f i \\<^sub>C a) has_sum (x \\<^sub>C a)) I\ using assms has_sum_bounded_linear[unfolded o_def] bounded_antilinear.bounded_linear bounded_antilinear_cinner_left by blast lemma summable_on_cinner_right: assumes \f summable_on I\ shows \(\i. f i \\<^sub>C a) summable_on I\ by (metis assms has_sum_cinner_right summable_on_def) lemma infsum_cinner_right: assumes \\ summable_on I\ shows \(\\<^sub>\i\I. \ i) \\<^sub>C \ = (\\<^sub>\i\I. \ i \\<^sub>C \)\ by (metis assms has_sum_cinner_right has_sum_infsum infsumI) lemma Cauchy_cinner_product_summable: assumes asum: \a summable_on UNIV\ assumes bsum: \b summable_on UNIV\ assumes \finite X\ \finite Y\ assumes pos: \\x y. x \ X \ y \ Y \ cinner (a x) (b y) \ 0\ shows absum: \(\(x, y). cinner (a x) (b y)) summable_on UNIV\ proof - have \(\(x,y)\F. norm (cinner (a x) (b y))) \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) + norm (infsum b (-Y)) + 1\ if \finite F\ and \F \ (-X)\(-Y)\ for F proof - from asum \finite X\ have \a summable_on (-X)\ by (simp add: Compl_eq_Diff_UNIV summable_on_cofin_subset) then obtain MA where \finite MA\ and \MA \ -X\ and MA: \G \ MA \ G \ -X \ finite G \ norm (sum a G - infsum a (-X)) \ 1\ for G apply (simp add: summable_iff_has_sum_infsum has_sum_metric dist_norm) by (meson less_eq_real_def zero_less_one) from bsum \finite Y\ have \b summable_on (-Y)\ by (simp add: Compl_eq_Diff_UNIV summable_on_cofin_subset) then obtain MB where \finite MB\ and \MB \ -Y\ and MB: \G \ MB \ G \ -Y \ finite G \ norm (sum b G - infsum b (-Y)) \ 1\ for G apply (simp add: summable_iff_has_sum_infsum has_sum_metric dist_norm) by (meson less_eq_real_def zero_less_one) define F1 F2 where \F1 = fst ` F \ MA\ and \F2 = snd ` F \ MB\ define t1 t2 where \t1 = sum a F1 - infsum a (-X)\ and \t2 = sum b F2 - infsum b (-Y)\ have [simp]: \finite F1\ \finite F2\ using F1_def F2_def \finite MA\ \finite MB\ that by auto have [simp]: \F1 \ -X\ \F2 \ -Y\ using \F \ (-X)\(-Y)\ \MA \ -X\ \MB \ -Y\ by (auto simp: F1_def F2_def) from MA[OF _ \F1 \ -X\ \finite F1\] have \norm t1 \ 1\ by (auto simp: t1_def F1_def) from MB[OF _ \F2 \ -Y\ \finite F2\] have \norm t2 \ 1\ by (auto simp: t2_def F2_def) have [simp]: \F \ F1 \ F2\ by (force simp: F1_def F2_def image_def) have \(\(x,y)\F. norm (cinner (a x) (b y))) \ (\(x,y)\F1\F2. norm (cinner (a x) (b y)))\ by (intro sum_mono2) auto also have "... = (\x\F1 \ F2. norm (a (fst x) \\<^sub>C b (snd x)))" by (auto simp: case_prod_beta) also have "... = norm (\x\F1 \ F2. a (fst x) \\<^sub>C b (snd x))" proof - have "(\x\F1 \ F2. \a (fst x) \\<^sub>C b (snd x)\) = \\x\F1 \ F2. a (fst x) \\<^sub>C b (snd x)\" by (smt (verit, best) pos sum.cong sum_nonneg ComplD SigmaE \F1 \ - X\ \F2 \ - Y\ abs_pos prod.sel subset_iff) then show ?thesis by (smt (verit) abs_complex_def norm_ge_zero norm_of_real o_def of_real_sum sum.cong sum_norm_le) qed also from pos have \\ = norm (\(x,y)\F1\F2. cinner (a x) (b y))\ by (auto simp: case_prod_beta) also have \\ = norm (cinner (sum a F1) (sum b F2))\ by (simp add: sum.cartesian_product sum_cinner) also have \\ = norm (cinner (infsum a (-X) + t1) (infsum b (-Y) + t2))\ by (simp add: t1_def t2_def) also have \\ \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) * norm t2 + norm t1 * norm (infsum b (-Y)) + norm t1 * norm t2\ apply (simp add: cinner_add_right cinner_add_left) by (smt (verit, del_insts) complex_inner_class.Cauchy_Schwarz_ineq2 norm_triangle_ineq) also from \norm t1 \ 1\ \norm t2 \ 1\ have \\ \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) + norm (infsum b (-Y)) + 1\ by (auto intro!: add_mono mult_left_le mult_left_le_one_le mult_le_one) finally show ?thesis by - qed then have \(\(x, y). cinner (a x) (b y)) abs_summable_on (-X)\(-Y)\ apply (rule_tac nonneg_bdd_above_summable_on) by (auto intro!: bdd_aboveI2 simp: case_prod_unfold) then have 1: \(\(x, y). cinner (a x) (b y)) summable_on (-X)\(-Y)\ using abs_summable_summable by blast from bsum have \(\y. b y) summable_on (-Y)\ by (simp add: Compl_eq_Diff_UNIV assms(4) summable_on_cofin_subset) then have \(\y. cinner (a x) (b y)) summable_on (-Y)\ for x using summable_on_cinner_left by blast with \finite X\ have 2: \(\(x, y). cinner (a x) (b y)) summable_on X\(-Y)\ by (force intro: summable_on_product_finite_left) from asum have \(\x. a x) summable_on (-X)\ by (simp add: Compl_eq_Diff_UNIV assms(3) summable_on_cofin_subset) then have \(\x. cinner (a x) (b y)) summable_on (-X)\ for y using summable_on_cinner_right by blast with \finite Y\ have 3: \(\(x, y). cinner (a x) (b y)) summable_on (-X)\Y\ by (force intro: summable_on_product_finite_right) have 4: \(\(x, y). cinner (a x) (b y)) summable_on X\Y\ by (simp add: \finite X\ \finite Y\) have \
: "UNIV = ((-X) \ -Y) \ (X \ -Y) \ ((-X) \ Y) \ (X \ Y)" by auto show ?thesis using 1 2 3 4 by (force simp: \
intro!: summable_on_Un_disjoint) qed text \A variant of @{thm [source] Series.Cauchy_product_sums} with \<^term>\(*)\ replaced by \<^term>\cinner\. Differently from @{thm [source] Series.Cauchy_product_sums}, we do not require absolute summability of \<^term>\a\ and \<^term>\b\ individually but only unconditional summability of \<^term>\a\, \<^term>\b\, and their product. While on, e.g., reals, unconditional summability is equivalent to absolute summability, in general unconditional summability is a weaker requirement. Logically belong in \<^theory>\Complex_Bounded_Operators.Complex_Inner_Product\ but the proof uses facts from this theory.\ lemma fixes a b :: "nat \ 'a::complex_inner" assumes asum: \a summable_on UNIV\ assumes bsum: \b summable_on UNIV\ assumes absum: \(\(x, y). cinner (a x) (b y)) summable_on UNIV\ (* \ \See @{thm [source] Cauchy_cinner_product_summable} or @{thm [source] Cauchy_cinner_product_summable'} for a way to rewrite this premise.\ *) shows Cauchy_cinner_product_infsum: \(\\<^sub>\k. \i\k. cinner (a i) (b (k - i))) = cinner (\\<^sub>\k. a k) (\\<^sub>\k. b k)\ and Cauchy_cinner_product_infsum_exists: \(\k. \i\k. cinner (a i) (b (k - i))) summable_on UNIV\ proof - have img: \(\(k::nat, i). (i, k - i)) ` {(k, i). i \ k} = UNIV\ apply (auto simp: image_def) by (metis add.commute add_diff_cancel_right' diff_le_self) have inj: \inj_on (\(k::nat, i). (i, k - i)) {(k, i). i \ k}\ by (smt (verit, del_insts) Pair_inject case_prodE case_prod_conv eq_diff_iff inj_onI mem_Collect_eq) have sigma: \(SIGMA k:UNIV. {i. i \ k}) = {(k, i). i \ k}\ by auto from absum have \
: \(\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ by (rule Cauchy_cinner_product_summable'[THEN iffD1]) then have \(\k. \\<^sub>\i|i\k. cinner (a i) (b (k-i))) summable_on UNIV\ by (metis (mono_tags, lifting) sigma summable_on_Sigma_banach summable_on_cong) then show \(\k. \i\k. cinner (a i) (b (k - i))) summable_on UNIV\ by (metis (mono_tags, lifting) atMost_def finite_Collect_le_nat infsum_finite summable_on_cong) have \cinner (\\<^sub>\k. a k) (\\<^sub>\k. b k) = (\\<^sub>\k. \\<^sub>\l. cinner (a k) (b l))\ by (smt (verit, best) asum bsum infsum_cinner_left infsum_cinner_right infsum_cong) also have \\ = (\\<^sub>\(k,l). cinner (a k) (b l))\ by (smt (verit) UNIV_Times_UNIV infsum_Sigma'_banach infsum_cong local.absum) also have \\ = (\\<^sub>\(k, l)\(\(k, i). (i, k - i)) ` {(k, i). i \ k}. cinner (a k) (b l))\ by (simp only: img) also have \\ = infsum ((\(k, l). a k \\<^sub>C b l) \ (\(k, i). (i, k - i))) {(k, i). i \ k}\ using inj by (rule infsum_reindex) also have \\ = (\\<^sub>\(k,i)|i\k. a i \\<^sub>C b (k-i))\ by (simp add: o_def case_prod_unfold) also have \\ = (\\<^sub>\k. \\<^sub>\i|i\k. a i \\<^sub>C b (k-i))\ by (metis (no_types) \
infsum_Sigma'_banach sigma) also have \\ = (\\<^sub>\k. \i\k. a i \\<^sub>C b (k-i))\ by (simp add: atMost_def) finally show \(\\<^sub>\k. \i\k. a i \\<^sub>C b (k - i)) = (\\<^sub>\k. a k) \\<^sub>C (\\<^sub>\k. b k)\ by simp qed lemma CBlinfun_plus: assumes [simp]: \bounded_clinear f\ \bounded_clinear g\ shows \CBlinfun (f + g) = CBlinfun f + CBlinfun g\ by (auto intro!: cblinfun_eqI simp: plus_fun_def bounded_clinear_add CBlinfun_inverse cblinfun.add_left) lemma CBlinfun_scaleC: assumes \bounded_clinear f\ shows \CBlinfun (\y. c *\<^sub>C f y) = c *\<^sub>C CBlinfun f\ by (simp add: assms eq_onp_same_args scaleC_cblinfun.abs_eq) 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 norm_cblinfun_Sup: \norm A = (SUP \. norm (A *\<^sub>V \) / norm \)\ by (simp add: norm_cblinfun.rep_eq onorm_def) 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_closure 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" by (metis assms cblinfun_eqI cblinfun_eq_on ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) declare cnj_bounded_antilinear[bounded_antilinear] lemma Cblinfun_comp_bounded_cbilinear: \bounded_clinear (CBlinfun o p)\ if \bounded_cbilinear p\ by (metis bounded_cbilinear.bounded_clinear_prod_right bounded_cbilinear.prod_right_def comp_id map_fun_def that) lemma Cblinfun_comp_bounded_sesquilinear: \bounded_antilinear (CBlinfun o p)\ if \bounded_sesquilinear p\ by (metis (mono_tags, opaque_lifting) bounded_clinear_CBlinfun_apply bounded_sesquilinear.bounded_clinear_right comp_apply that transfer_bounded_sesquilinear_bounded_antilinearI) 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 by transfer (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" by transfer (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" by transfer (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" by transfer 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\ by transfer 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) instance blinfun :: (real_normed_vector, cbanach) cbanach.. 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" by transfer (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 by transfer auto lemma blinfun_cblinfun_eq_right_total[transfer_rule]: \right_total blinfun_cblinfun_eq\ unfolding right_total_def by transfer (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)" by transfer 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 by transfer 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 by transfer 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 by transfer 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 by transfer (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 by transfer 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 by transfer 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 by transfer (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 by transfer auto lemma cblinfun_blinfun_transfer_apply[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=) ===> (=)) blinfun_apply cblinfun_apply" unfolding rel_fun_def by transfer 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\ by transfer simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ by transfer auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ by transfer auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ by transfer auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ by transfer 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)\ by transfer auto 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)\ by transfer auto subsection \Composition\ 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)\ by transfer (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" by transfer auto lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" by transfer auto 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) lemma cblinfun_compose_sum_left: \(\i\S. g i) o\<^sub>C\<^sub>L x = (\i\S. g i o\<^sub>C\<^sub>L x)\ by (induction S rule:infinite_finite_induct) (auto simp: cblinfun_compose_add_left) lemma cblinfun_compose_sum_right: \x o\<^sub>C\<^sub>L (\i\S. g i) = (\i\S. x o\<^sub>C\<^sub>L g i)\ by (induction S rule:infinite_finite_induct) (auto simp: cblinfun_compose_add_right) lemma cblinfun_compose_minus_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.diff_right bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_minus_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.diff_left bounded_cbilinear_cblinfun_compose) lemma simp_a_oCL_b: \a o\<^sub>C\<^sub>L b = c \ a o\<^sub>C\<^sub>L (b o\<^sub>C\<^sub>L d) = c o\<^sub>C\<^sub>L d\ \ \A convenience lemma to transform simplification rules of the form \<^term>\a o\<^sub>C\<^sub>L b = c\. E.g., \simp_a_oCL_b[OF isometryD, simp]\ declares a simp-rule for simplifying \<^term>\adj x o\<^sub>C\<^sub>L (x o\<^sub>C\<^sub>L y) = id_cblinfun o\<^sub>C\<^sub>L y\.\ by (auto simp: cblinfun_compose_assoc) lemma simp_a_oCL_b': \a o\<^sub>C\<^sub>L b = c \ a *\<^sub>V (b *\<^sub>V d) = c *\<^sub>V d\ \ \A convenience lemma to transform simplification rules of the form \<^term>\a o\<^sub>C\<^sub>L b = c\. E.g., \simp_a_oCL_b'[OF isometryD, simp]\ declares a simp-rule for simplifying \<^term>\adj x *\<^sub>V x *\<^sub>V y = id_cblinfun *\<^sub>V y\.\ by auto lemma cblinfun_compose_uminus_left: \(- a) o\<^sub>C\<^sub>L b = - (a o\<^sub>C\<^sub>L b)\ by (simp add: bounded_cbilinear.minus_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_uminus_right: \a o\<^sub>C\<^sub>L (- b) = - (a o\<^sub>C\<^sub>L b)\ by (simp add: bounded_cbilinear.minus_right bounded_cbilinear_cblinfun_compose) lemma bounded_clinear_cblinfun_compose_left: \bounded_clinear (\x. x o\<^sub>C\<^sub>L y)\ by (simp add: bounded_cbilinear.bounded_clinear_left bounded_cbilinear_cblinfun_compose) lemma bounded_clinear_cblinfun_compose_right: \bounded_clinear (\y. x o\<^sub>C\<^sub>L y)\ by (simp add: bounded_cbilinear.bounded_clinear_right bounded_cbilinear_cblinfun_compose) lemma clinear_cblinfun_compose_left: \clinear (\x. x o\<^sub>C\<^sub>L y)\ by (simp add: bounded_cbilinear.bounded_clinear_left bounded_cbilinear_cblinfun_compose bounded_clinear.clinear) lemma clinear_cblinfun_compose_right: \clinear (\y. x o\<^sub>C\<^sub>L y)\ by (simp add: bounded_clinear.clinear bounded_clinear_cblinfun_compose_right) lemma additive_cblinfun_compose_left[simp]: \Modules.additive (\x. x o\<^sub>C\<^sub>L a)\ by (simp add: Modules.additive_def cblinfun_compose_add_left) lemma additive_cblinfun_compose_right[simp]: \Modules.additive (\x. a o\<^sub>C\<^sub>L x)\ by (simp add: Modules.additive_def cblinfun_compose_add_right) lemma isCont_cblinfun_compose_left: \isCont (\x. x o\<^sub>C\<^sub>L a) y\ apply (rule clinear_continuous_at) by (rule bounded_clinear_cblinfun_compose_left) lemma isCont_cblinfun_compose_right: \isCont (\x. a o\<^sub>C\<^sub>L x) y\ apply (rule clinear_continuous_at) by (rule bounded_clinear_cblinfun_compose_right) lemma cspan_compose_closed: assumes \\a b. a \ A \ b \ A \ a o\<^sub>C\<^sub>L b \ A\ assumes \a \ cspan A\ and \b \ cspan A\ shows \a o\<^sub>C\<^sub>L b \ cspan A\ proof - from \a \ cspan A\ obtain F f where \finite F\ and \F \ A\ and a_def: \a = (\x\F. f x *\<^sub>C x)\ by (smt (verit, del_insts) complex_vector.span_explicit mem_Collect_eq) from \b \ cspan A\ obtain G g where \finite G\ and \G \ A\ and b_def: \b = (\x\G. g x *\<^sub>C x)\ by (smt (verit, del_insts) complex_vector.span_explicit mem_Collect_eq) have \a o\<^sub>C\<^sub>L b = (\(x,y)\F\G. (f x *\<^sub>C x) o\<^sub>C\<^sub>L (g y *\<^sub>C y))\ apply (simp add: a_def b_def cblinfun_compose_sum_left) by (auto intro!: sum.cong simp add: cblinfun_compose_sum_right scaleC_sum_right sum.cartesian_product case_prod_beta) also have \\ = (\(x,y)\F\G. (f x * g y) *\<^sub>C (x o\<^sub>C\<^sub>L y))\ by (metis (no_types, opaque_lifting) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right scaleC_scaleC) also have \\ \ cspan A\ using assms \G \ A\ \F \ A\ apply (auto intro!: complex_vector.span_sum complex_vector.span_scale simp: complex_vector.span_clauses) using complex_vector.span_clauses(1) by blast finally show ?thesis by - qed 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" by (metis adj.rep_eq apply_id_cblinfun cadjoint_id cblinfun_apply_inject) 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 \\<^sub>C v) = (u \\<^sub>C (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*)" by transfer (simp add: 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 \\<^sub>C v) = (u \\<^sub>C 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_adj_left: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \(G* *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (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 \\<^sub>C (G* *\<^sub>V y) = (G *\<^sub>V x) \\<^sub>C 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)\ by intro_classes simp from True have c2: \class.not_singleton TYPE('c)\ by intro_classes 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\ by (simp add: adj_plus antilinearI) 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::complex_inner\ and F:: \'a \\<^sub>C\<^sub>L 'b\ assumes \\x y. ((cblinfun_apply F) x \\<^sub>C y) = (x \\<^sub>C (cblinfun_apply G) y)\ shows \F = G*\ using assms apply transfer using cadjoint_eqI by auto lemma adj_uminus: \(-A)* = - (A*)\ by (metis scaleR_adj scaleR_minus1_left scaleR_minus1_left) lemma cinner_real_hermiteanI: \ \Prop. II.2.12 in \<^cite>\conway2013course\\ assumes \\\. \ \\<^sub>C (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" apply (rule_tac sym) 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\ by (meson \ complex_of_real_mono mult_nonneg_nonneg norm_ge_zero power_mono real_sqrt_ge_zero \\ \ 0\) also have \\ \ cinner (A* *\<^sub>V \) (A* *\<^sub>V \)\ by (auto simp flip: power2_norm_eq_cinner) also have \
: \\ = cinner \ ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ by (auto simp: cinner_adj_left) also have \\ \ norm (A o\<^sub>C\<^sub>L A*)\ using \norm \ = 1\ by (smt (verit) Re_complex_of_real \
cdot_square_norm cinner_ge_zero cmod_Re complex_inner_class.Cauchy_Schwarz_ineq2 complex_of_real_mono 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 lemma sum_adj: \(sum a F)* = sum (\i. (a i)*) F\ by (induction rule:infinite_finite_induct) (auto simp add: adj_plus) lemma has_sum_adj: assumes \(f has_sum x) I\ shows \((\x. adj (f x)) has_sum adj x) I\ apply (rule has_sum_comm_additive[where f=adj, unfolded o_def]) apply (simp add: antilinear.axioms(1)) apply (metis (no_types, lifting) LIM_eq adj_plus adj_uminus norm_adj uminus_add_conv_diff) by (simp add: assms) lemma adj_minus: \(A - B)* = (A*) - (B*)\ by (metis add_implies_diff adj_plus diff_add_cancel) lemma cinner_hermitian_real: \x \\<^sub>C (A *\<^sub>V x) \ \\ if \A* = A\ by (metis Reals_cnj_iff cinner_adj_right cinner_commute' that) lemma adj_inject: \adj a = adj b \ a = b\ by (metis (no_types, opaque_lifting) adj_minus eq_iff_diff_eq_0 norm_adj norm_eq_zero) lemma norm_AadjA[simp]: \norm (A* o\<^sub>C\<^sub>L A) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ by (metis double_adj norm_AAadj norm_adj) lemma cspan_adj_closed: assumes \\a. a \ A \ a* \ A\ assumes \a \ cspan A\ shows \a* \ cspan A\ proof - from \a \ cspan A\ obtain F f where \finite F\ and \F \ A\ and \a = (\x\F. f x *\<^sub>C x)\ by (smt (verit, del_insts) complex_vector.span_explicit mem_Collect_eq) then have \a* = (\x\F. cnj (f x) *\<^sub>C x*)\ by (auto simp: sum_adj) also have \\ \ cspan A\ using assms \F \ A\ by (auto intro!: complex_vector.span_sum complex_vector.span_scale simp: complex_vector.span_clauses) finally show ?thesis by - qed subsection \Powers of operators\ lift_definition cblinfun_power :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'a \ nat \ 'a \\<^sub>C\<^sub>L 'a\ is \\(a::'a\'a) n. a ^^ n\ apply (rename_tac f n, induct_tac n, auto simp: Nat.funpow_code_def) by (simp add: bounded_clinear_compose) lemma cblinfun_power_0[simp]: \cblinfun_power A 0 = id_cblinfun\ by transfer auto lemma cblinfun_power_Suc': \cblinfun_power A (Suc n) = A o\<^sub>C\<^sub>L cblinfun_power A n\ by transfer auto lemma cblinfun_power_Suc: \cblinfun_power A (Suc n) = cblinfun_power A n o\<^sub>C\<^sub>L A\ apply (induction n) by (auto simp: cblinfun_power_Suc' simp flip: cblinfun_compose_assoc) lemma cblinfun_power_compose[simp]: \cblinfun_power A n o\<^sub>C\<^sub>L cblinfun_power A m = cblinfun_power A (n+m)\ apply (induction n) by (auto simp: cblinfun_power_Suc' cblinfun_compose_assoc) lemma cblinfun_power_scaleC: \cblinfun_power (c *\<^sub>C a) n = c^n *\<^sub>C cblinfun_power a n\ apply (induction n) by (auto simp: cblinfun_power_Suc) lemma cblinfun_power_scaleR: \cblinfun_power (c *\<^sub>R a) n = c^n *\<^sub>R cblinfun_power a n\ apply (induction n) by (auto simp: cblinfun_power_Suc) lemma cblinfun_power_uminus: \cblinfun_power (-a) n = (-1)^n *\<^sub>R cblinfun_power a n\ apply (subst asm_rl[of \-a = (-1) *\<^sub>R a\]) by simp (rule cblinfun_power_scaleR) lemma cblinfun_power_adj: \(cblinfun_power S n)* = cblinfun_power (S*) n\ apply (induction n) apply simp apply (subst cblinfun_power_Suc) apply (subst cblinfun_power_Suc') by auto 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 unitaryI: \unitary a\ if \a* o\<^sub>C\<^sub>L a = id_cblinfun\ and \a o\<^sub>C\<^sub>L a* = id_cblinfun\ unfolding unitary_def using that by simp 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 by transfer 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_closure[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 lemma isometry_preserves_norm: \isometry U \ norm (U *\<^sub>V \) = norm \\ by (metis (no_types, lifting) cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply cinner_adj_right cnorm_eq isometryD) lemma norm_isometry_compose: assumes \isometry U\ shows \norm (U o\<^sub>C\<^sub>L A) = norm A\ proof - have *: \norm (U *\<^sub>V A *\<^sub>V \) = norm (A *\<^sub>V \)\ for \ by (smt (verit, ccfv_threshold) assms cblinfun_apply_cblinfun_compose cinner_adj_right cnorm_eq id_cblinfun_apply isometryD) have \norm (U o\<^sub>C\<^sub>L A) = (SUP \. norm (U *\<^sub>V A *\<^sub>V \) / norm \)\ unfolding norm_cblinfun_Sup by auto also have \\ = (SUP \. norm (A *\<^sub>V \) / norm \)\ using * by auto also have \\ = norm A\ unfolding norm_cblinfun_Sup by auto finally show ?thesis by simp qed lemma norm_isometry: fixes U :: \'a::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ assumes \isometry U\ shows \norm U = 1\ apply (subst asm_rl[of \U = U o\<^sub>C\<^sub>L id_cblinfun\], simp) apply (subst norm_isometry_compose, simp add: assms) by simp lemma norm_preserving_isometry: \isometry U\ if \\\. norm (U *\<^sub>V \) = norm \\ by (smt (verit, ccfv_SIG) cblinfun_cinner_eqI cblinfun_id_cblinfun_apply cinner_adj_right cnorm_eq isometry_def simp_a_oCL_b' that) lemma norm_isometry_compose': \norm (A o\<^sub>C\<^sub>L U) = norm A\ if \isometry (U*)\ by (smt (verit) cblinfun_compose_assoc cblinfun_compose_id_right double_adj isometryD mult_cancel_left2 norm_AadjA norm_cblinfun_compose norm_isometry_compose norm_zero power2_eq_square right_diff_distrib that zero_less_norm_iff) lemma unitary_nonzero[simp]: \\ unitary (0 :: 'a::{chilbert_space, not_singleton} \\<^sub>C\<^sub>L _)\ by (simp add: unitary_def) lemma isometry_inj: assumes \isometry U\ shows \inj_on U X\ apply (rule inj_on_inverseI[where g=\U*\]) using assms by (simp flip: cblinfun_apply_cblinfun_compose) lemma unitary_inj: assumes \unitary U\ shows \inj_on U X\ apply (rule isometry_inj) using assms by simp lemma unitary_adj_inv: \unitary U \ cblinfun_apply (U*) = inv (cblinfun_apply U)\ apply (rule inj_imp_inv_eq[symmetric]) apply (simp add: unitary_inj) unfolding unitary_def by (simp flip: cblinfun_apply_cblinfun_compose) lemma isometry_cinner_both_sides: assumes \isometry U\ shows \cinner (U x) (U y) = cinner x y\ using assms by (simp add: flip: cinner_adj_right cblinfun_apply_cblinfun_compose) lemma isometry_image_is_ortho_set: assumes \is_ortho_set A\ assumes \isometry U\ shows \is_ortho_set (U ` A)\ using assms apply (auto simp add: is_ortho_set_def isometry_cinner_both_sides) by (metis cinner_eq_zero_iff isometry_cinner_both_sides) subsection \Product spaces\ lift_definition cblinfun_left :: \'a::complex_normed_vector \\<^sub>C\<^sub>L ('a\'b::complex_normed_vector)\ is \(\x. (x,0))\ by (auto intro!: bounded_clinearI[where K=1]) lift_definition cblinfun_right :: \'b::complex_normed_vector \\<^sub>C\<^sub>L ('a::complex_normed_vector\'b)\ is \(\x. (0,x))\ by (auto intro!: bounded_clinearI[where K=1]) lemma isometry_cblinfun_left[simp]: \isometry cblinfun_left\ apply (rule orthogonal_on_basis_is_isometry[of some_chilbert_basis]) apply simp by transfer simp lemma isometry_cblinfun_right[simp]: \isometry cblinfun_right\ apply (rule orthogonal_on_basis_is_isometry[of some_chilbert_basis]) apply simp by transfer simp lemma cblinfun_left_right_ortho[simp]: \cblinfun_left* o\<^sub>C\<^sub>L cblinfun_right = 0\ proof - have \x \\<^sub>C ((cblinfun_left* o\<^sub>C\<^sub>L cblinfun_right) *\<^sub>V y) = 0\ for x :: 'b and y :: 'a apply (simp add: cinner_adj_right) by transfer auto then show ?thesis by (metis cblinfun.zero_left cblinfun_eqI cinner_eq_zero_iff) qed lemma cblinfun_right_left_ortho[simp]: \cblinfun_right* o\<^sub>C\<^sub>L cblinfun_left = 0\ proof - have \x \\<^sub>C ((cblinfun_right* o\<^sub>C\<^sub>L cblinfun_left) *\<^sub>V y) = 0\ for x :: 'b and y :: 'a apply (simp add: cinner_adj_right) by transfer auto then show ?thesis by (metis cblinfun.zero_left cblinfun_eqI cinner_eq_zero_iff) qed lemma cblinfun_left_apply[simp]: \cblinfun_left *\<^sub>V \ = (\,0)\ by transfer simp lemma cblinfun_left_adj_apply[simp]: \cblinfun_left* *\<^sub>V \ = fst \\ apply (cases \) by (auto intro!: cinner_extensionality[of \_ *\<^sub>V _\] simp: cinner_adj_right) lemma cblinfun_right_apply[simp]: \cblinfun_right *\<^sub>V \ = (0,\)\ by transfer simp lemma cblinfun_right_adj_apply[simp]: \cblinfun_right* *\<^sub>V \ = snd \\ apply (cases \) by (auto intro!: cinner_extensionality[of \_ *\<^sub>V _\] simp: cinner_adj_right) lift_definition ccsubspace_Times :: \'a::complex_normed_vector ccsubspace \ 'b::complex_normed_vector ccsubspace \ ('a\'b) ccsubspace\ is Product_Type.Times proof - fix S :: \'a set\ and T :: \'b set\ assume [simp]: \closed_csubspace S\ \closed_csubspace T\ have \csubspace (S \ T)\ by (simp add: complex_vector.subspace_Times) moreover have \closed (S \ T)\ by (simp add: closed_Times closed_csubspace.closed) ultimately show \closed_csubspace (S \ T)\ by (rule closed_csubspace.intro) qed lemma ccspan_Times: \ccspan (S \ T) = ccsubspace_Times (ccspan S) (ccspan T)\ if \0 \ S\ and \0 \ T\ proof (transfer fixing: S T) from that have \closure (cspan (S \ T)) = closure (cspan S \ cspan T)\ by (simp add: cspan_Times) also have \\ = closure (cspan S) \ closure (cspan T)\ using closure_Times by blast finally show \closure (cspan (S \ T)) = closure (cspan S) \ closure (cspan T)\ by - qed lemma ccspan_Times_sing1: \ccspan ({0::'a::complex_normed_vector} \ B) = ccsubspace_Times 0 (ccspan B)\ proof (transfer fixing: B) have \closure (cspan ({0::'a} \ B)) = closure ({0} \ cspan B)\ by (simp add: complex_vector.span_Times_sing1) also have \\ = closure {0} \ closure (cspan B)\ using closure_Times by blast also have \\ = {0} \ closure (cspan B)\ by simp finally show \closure (cspan ({0::'a} \ B)) = {0} \ closure (cspan B)\ by - qed lemma ccspan_Times_sing2: \ccspan (B \ {0::'a::complex_normed_vector}) = ccsubspace_Times (ccspan B) 0\ proof (transfer fixing: B) have \closure (cspan (B \ {0::'a})) = closure (cspan B \ {0})\ by (simp add: complex_vector.span_Times_sing2) also have \\ = closure (cspan B) \ closure {0}\ using closure_Times by blast also have \\ = closure (cspan B) \ {0}\ by simp finally show \closure (cspan (B \ {0::'a})) = closure (cspan B) \ {0}\ by - qed lemma ccsubspace_Times_sup: \sup (ccsubspace_Times A B) (ccsubspace_Times C D) = ccsubspace_Times (sup A C) (sup B D)\ proof transfer fix A C :: \'a set\ and B D :: \'b set\ have \A \ B +\<^sub>M C \ D = closure ((A \ B) + (C \ D))\ using closed_sum_def by blast also have \\ = closure ((A + C) \ (B + D))\ by (simp add: set_Times_plus_distrib) also have \\ = closure (A + C) \ closure (B + D)\ by (simp add: closure_Times) also have \\ = (A +\<^sub>M C) \ (B +\<^sub>M D)\ by (simp add: closed_sum_def) finally show \A \ B +\<^sub>M C \ D = (A +\<^sub>M C) \ (B +\<^sub>M D)\ by - qed lemma ccsubspace_Times_top_top[simp]: \ccsubspace_Times top top = top\ by transfer simp +lemma is_ortho_set_prod: + assumes \is_ortho_set B\ \is_ortho_set B'\ + shows \is_ortho_set ((B \ {0}) \ ({0} \ B'))\ + using assms unfolding is_ortho_set_def + apply (auto simp: is_onb_def is_ortho_set_def zero_prod_def) + by (meson is_onb_def is_ortho_set_def)+ + +lemma ccsubspace_Times_ccspan: + assumes \ccspan B = S\ and \ccspan B' = S'\ + shows \ccspan ((B \ {0}) \ ({0} \ B')) = ccsubspace_Times S S'\ + by (smt (z3) Diff_eq_empty_iff Sigma_cong assms(1) assms(2) ccspan.rep_eq ccspan_0 ccspan_Times_sing1 ccspan_Times_sing2 ccspan_of_empty ccspan_remove_0 ccspan_superset ccspan_union ccsubspace_Times_sup complex_vector.span_insert_0 space_as_set_bot sup_bot_left sup_bot_right) + lemma is_onb_prod: assumes \is_onb B\ \is_onb B'\ shows \is_onb ((B \ {0}) \ ({0} \ B'))\ -proof - - from assms - have 1: \is_ortho_set ((B \ {0}) \ ({0} \ B'))\ - unfolding is_ortho_set_def - apply (auto simp: is_onb_def is_ortho_set_def zero_prod_def) - by (meson is_onb_def is_ortho_set_def)+ - - have 2: \(l, r) \ B \ {0} \ norm (l, r) = 1\ for l :: 'a and r :: 'b - using \is_onb B\ is_onb_def by auto - - have 3: \(l, r) \ {0} \ B' \ norm (l, r) = 1\ for l :: 'a and r :: 'b - using \is_onb B'\ is_onb_def by auto - - have [simp]: \ccspan B = top\ \ccspan B' = top\ - using assms is_onb_def by auto - - have 4: \ccspan ((B \ {0}) \ ({0} \ B')) = top\ - by (auto simp: ccspan_Times_sing1 ccspan_Times_sing2 ccsubspace_Times_sup simp flip: ccspan_union) - - from 1 2 3 4 - show \is_onb ((B \ {0}) \ ({0} \ B'))\ - by (auto simp add: is_onb_def) -qed + using assms by (auto intro!: is_ortho_set_prod simp add: is_onb_def ccsubspace_Times_ccspan) subsection \Images\ text \The following definition defines the image of a closed subspace \<^term>\S\ under a bounded operator \<^term>\A\. We do not define that image as the image of \<^term>\A\ seen as a function (\<^term>\A ` S\) but as the topological closure of that image. This is because \<^term>\A ` S\ might in general not be closed. For example, if $e_i$ ($i\in\mathbb N$) form an orthonormal basis, and $A$ maps $e_i$ to $e_i/i$, then all $e_i$ are in \<^term>\A ` S\, so the closure of \<^term>\A ` S\ is the whole space. However, $\sum_i e_i/i$ is not in \<^term>\A ` S\ because its preimage would have to be $\sum_i e_i$ which does not converge. So \<^term>\A ` S\ does not contain the whole space, hence it is not closed.\ 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 by transfer (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_ccsubspace (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_ccsubspace (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_ccsubspace (closure ((*\<^sub>V) (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_ccsubspace (closure ((*\<^sub>V) A ` space_as_set S))" using \Abs_ccsubspace (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_ccsubspace (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 \ = \" by transfer (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\X. V i) \ (INF i\X. 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 closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis 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)" by transfer (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" and \X \ {}\ shows "U *\<^sub>S (INF i\X. V i) = (INF i\X. U *\<^sub>S V i)" proof (rule antisym) show "U *\<^sub>S (INF i\X. V i) \ (INF i\X. 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\X. U *\<^sub>S V i)" and INFV_def: "INFV = (INF i\X. 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 fixing: i) fix V :: "'a \ 'b set" and Uinv :: "'c \ 'b" and U :: "'b \ 'c" assume "pred_fun \ closed_csubspace V" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ (Uinv \ U) \ = \" then show "closure ((Uinv \ U) ` V i) = V i" proof auto fix x from \pred_fun \ closed_csubspace V\ show "x \ V i" if "x \ closure (V i)" using that apply simp by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) with \pred_fun \ closed_csubspace V\ show "x \ closure (V i)" if "x \ V i" 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 using \X \ {}\ by (metis INF_eq_const INF_lower2) 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" using that proof auto fix x show "x \ INFUV" if "x \ closure INFUV" using that \closed_csubspace INFUV\ by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure INFUV" if "x \ INFUV" using that \closed_csubspace INFUV\ 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\X. 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 by transfer (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\ \X \ {}\ shows "U *\<^sub>S (INF i\X. V i) = (INF i\X. 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 using \X \ {}\ 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 and X=UNIV] unfolding INF_UNIV_bool_expand using assms by auto lemma cblinfun_image_ccspan: shows "A *\<^sub>S ccspan G = ccspan ((*\<^sub>V) A ` G)" by transfer (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\ by transfer (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 by auto (smt (z3) cblinfun_plus_image_distr inf_sup_aci(6) le_iff_sup) qed next case False then show ?thesis by auto qed lemma space_as_set_image_commute: assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ and VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ (* I think only one of them is needed, can the lemma be strengthened? *) shows \(*\<^sub>V) U ` space_as_set T = space_as_set (U *\<^sub>S T)\ proof - have \space_as_set (U *\<^sub>S T) = U ` V ` space_as_set (U *\<^sub>S T)\ by (simp add: image_image UV flip: cblinfun_apply_cblinfun_compose) also have \\ \ U ` space_as_set (V *\<^sub>S U *\<^sub>S T)\ by (metis cblinfun_image.rep_eq closure_subset image_mono) also have \\ = U ` space_as_set T\ by (simp add: VU cblinfun_assoc_left(2)) finally have 1: \space_as_set (U *\<^sub>S T) \ U ` space_as_set T\ by - have 2: \U ` space_as_set T \ space_as_set (U *\<^sub>S T)\ by (simp add: cblinfun_image.rep_eq closure_subset) from 1 2 show ?thesis by simp qed lemma right_total_rel_ccsubspace: fixes R :: \'a::complex_normed_vector \ 'b::complex_normed_vector \ bool\ assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ assumes VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ assumes R_def: \\x y. R x y \ x = U *\<^sub>V y\ shows \right_total (rel_ccsubspace R)\ proof (rule right_totalI) fix T :: \'b ccsubspace\ show \\S. rel_ccsubspace R S T\ apply (rule exI[of _ \U *\<^sub>S T\]) using UV VU by (auto simp add: rel_ccsubspace_def R_def rel_set_def simp flip: space_as_set_image_commute) qed lemma left_total_rel_ccsubspace: fixes R :: \'a::complex_normed_vector \ 'b::complex_normed_vector \ bool\ assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ assumes VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ assumes R_def: \\x y. R x y \ y = U *\<^sub>V x\ shows \left_total (rel_ccsubspace R)\ proof - have \right_total (rel_ccsubspace (conversep R))\ apply (rule right_total_rel_ccsubspace) using assms by auto then show ?thesis by (auto intro!: right_total_conversep[THEN iffD1] simp: converse_rel_ccsubspace) qed lemma cblinfun_image_bot_zero[simp]: \A *\<^sub>S top = bot \ A = 0\ by (metis Complex_Bounded_Linear_Function.zero_cblinfun_image bot_ccsubspace.rep_eq cblinfun_apply_in_image cblinfun_eqI empty_iff insert_iff zero_ccsubspace_def) lemma surj_isometry_is_unitary: \ \This lemma is a bit stronger than its name suggests: We actually only require that the image of U is dense. The converse is @{thm [source] unitary_surj}\ fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes \isometry U\ assumes \U *\<^sub>S \ = \\ shows \unitary U\ by (metis UNIV_I assms(1) assms(2) cblinfun_assoc_left(1) cblinfun_compose_id_right cblinfun_eqI cblinfun_fixes_range id_cblinfun_apply isometry_def space_as_set_top unitary_def) lemma cblinfun_apply_in_image': "A *\<^sub>V \ \ space_as_set (A *\<^sub>S S)" if \\ \ space_as_set S\ by (metis cblinfun_image.rep_eq closure_subset image_subset_iff that) lemma cblinfun_image_ccspan_leqI: assumes \\v. v \ M \ A *\<^sub>V v \ space_as_set T\ shows \A *\<^sub>S ccspan M \ T\ by (simp add: assms cblinfun_image_ccspan ccspan_leqI image_subsetI) lemma cblinfun_same_on_image: \A \ = B \\ if eq: \A o\<^sub>C\<^sub>L C = B o\<^sub>C\<^sub>L C\ and mem: \\ \ space_as_set (C *\<^sub>S \)\ proof - have \A \ = B \\ if \\ \ range C\ for \ by (metis (no_types, lifting) eq cblinfun_apply_cblinfun_compose image_iff that) moreover have \\ \ closure (range C)\ by (metis cblinfun_image.rep_eq mem top_ccsubspace.rep_eq) ultimately show ?thesis apply (rule on_closure_eqI) by (auto simp: continuous_on_eq_continuous_at) qed lemma lift_cblinfun_comp: \ \Utility lemma to lift a lemma of the form \<^term>\a o\<^sub>C\<^sub>L b = c\ to become a more general rewrite rule.\ assumes \a o\<^sub>C\<^sub>L b = c\ shows \a o\<^sub>C\<^sub>L b = c\ and \a o\<^sub>C\<^sub>L (b o\<^sub>C\<^sub>L d) = c o\<^sub>C\<^sub>L d\ and \a *\<^sub>S (b *\<^sub>S S) = c *\<^sub>S S\ and \a *\<^sub>V (b *\<^sub>V x) = c *\<^sub>V x\ apply (fact assms) apply (simp add: assms cblinfun_assoc_left(1)) using assms cblinfun_assoc_left(2) apply force using assms by force lemma cblinfun_image_def2: \A *\<^sub>S S = ccspan ((*\<^sub>V) A ` space_as_set S)\ apply (simp add: flip: cblinfun_image_ccspan) by (metis ccspan_leqI ccspan_superset less_eq_ccsubspace.rep_eq order_class.order_eq_iff) lemma unitary_image_onb: \ \Logically belongs in an earlier section but the proof uses results from this section.\ assumes \is_onb A\ assumes \unitary U\ shows \is_onb (U ` A)\ using assms by (auto simp add: is_onb_def isometry_image_is_ortho_set isometry_preserves_norm simp flip: cblinfun_image_ccspan) 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 sandwich_arg_compose: assumes \isometry U\ shows \sandwich U x o\<^sub>C\<^sub>L sandwich U y = sandwich U (x o\<^sub>C\<^sub>L y)\ apply (simp add: sandwich_apply) by (metis (no_types, lifting) lift_cblinfun_comp(2) assms cblinfun_compose_id_right isometryD) 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) -lemma sandwich_compose: \sandwich A o\<^sub>C\<^sub>L sandwich B = sandwich (A o\<^sub>C\<^sub>L B)\ +lemma sandwich_compose: \sandwich (A o\<^sub>C\<^sub>L B) = sandwich A o\<^sub>C\<^sub>L sandwich B\ by (auto intro!: cblinfun_eqI simp: sandwich_apply) lemma inj_sandwich_isometry: \inj (sandwich U)\ if [simp]: \isometry U\ for U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ apply (rule inj_on_inverseI[where g=\(*\<^sub>V) (sandwich (U*))\]) - by (auto simp: sandwich_compose simp flip: cblinfun_apply_cblinfun_compose) + by (auto simp flip: cblinfun_apply_cblinfun_compose sandwich_compose) lemma sandwich_isometry_id: \isometry (U*) \ sandwich U id_cblinfun = id_cblinfun\ by (simp add: sandwich_apply isometry_def) 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\ by transfer (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)\ by transfer 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::complex_normed_vector \\<^sub>C\<^sub>L 'a \ bool\ is \\P. \M. is_projection_on P M\ . lemma is_Proj_id[simp]: \is_Proj id_cblinfun\ apply transfer by (auto intro!: exI[of _ UNIV] simp: is_projection_on_def is_arg_min_def) lemma Proj_top[simp]: \Proj \ = id_cblinfun\ by (metis Proj_idempotent Proj_range cblinfun_eqI cblinfun_fixes_range id_cblinfun_apply iso_tuple_UNIV_I space_as_set_top) 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: \is_orthogonal (x - P *\<^sub>V x) y\ 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) \\<^sub>C y = (x - P *\<^sub>V x) \\<^sub>C (P *\<^sub>V t)\ by (simp add: t_def) also have \\ = (P *\<^sub>V (x - P *\<^sub>V x)) \\<^sub>C t\ by (metis \P = P*\ cinner_adj_left) also have \\ = (P *\<^sub>V x - P *\<^sub>V (P *\<^sub>V x)) \\<^sub>C t\ by (simp add: cblinfun.diff_right) also have \\ = (P *\<^sub>V x - P *\<^sub>V x) \\<^sub>C t\ by (metis assms(1) comp_apply cblinfun_compose.rep_eq) also have \\ = (0 \\<^sub>C 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))" apply (rule is_projection_on_closed[where f=\cblinfun_apply P\]) using assms is_Proj.rep_eq is_projection_on_image by auto 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 Proj_range_closed[OF that] is_projection_on_cadjoint[where \=P and M=\range P\] by transfer (metis bounded_clinear.axioms(1) closed_csubspace_UNIV closed_csubspace_def complex_vector.linear_subspace_image is_projection_on_image) 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\ by transfer auto lemma ccsubspace_supI_via_Proj: fixes A B C::"'a::chilbert_space ccsubspace" assumes a1: \Proj (- C) *\<^sub>S A \ B\ shows "A \ 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 force 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_ccsubspace (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 apply transfer using is_projection_on_fixes_image is_projection_on_in_image by fastforce 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" apply transfer apply (rule exI[of _ 0]) by (simp add: is_projection_on_zero) lemma is_Proj_complement[simp]: fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ 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) lemma norm_Proj_leq1: \norm (Proj M) \ 1\ for M :: \'a :: chilbert_space ccsubspace\ by transfer (metis (no_types, opaque_lifting) mult.left_neutral onorm_bound projection_reduces_norm zero_less_one_class.zero_le_one) lemma Proj_orthog_ccspan_insert: assumes "\y. y \ Y \ is_orthogonal x y" shows \Proj (ccspan (insert x Y)) = proj x + Proj (ccspan Y)\ apply (subst asm_rl[of \insert x Y = {x} \ Y\], simp) apply (rule Proj_orthog_ccspan_union) using assms by auto lemma Proj_fixes_image: \Proj S *\<^sub>V \ = \\ if \\ \ space_as_set S\ by (metis Proj_idempotent Proj_range that cblinfun_fixes_range) lemma norm_is_Proj: \norm P \ 1\ if \is_Proj P\ for P :: \'a :: chilbert_space \\<^sub>C\<^sub>L 'a\ by (metis Proj_on_own_range norm_Proj_leq1 that) lemma Proj_sup: \orthogonal_spaces S T \ Proj (sup S T) = Proj S + Proj T\ unfolding orthogonal_spaces_def by transfer (simp add: projection_plus) lemma Proj_sum_spaces: assumes \finite X\ assumes \\x y. x\X \ y\X \ x\y \ orthogonal_spaces (J x) (J y)\ shows \Proj (\x\X. J x) = (\x\X. Proj (J x))\ using assms proof induction case empty show ?case by auto next case (insert x F) have \Proj (sum J (insert x F)) = Proj (J x \ sum J F)\ by (simp add: insert) also have \\ = Proj (J x) + Proj (sum J F)\ apply (rule Proj_sup) apply (rule orthogonal_sum) using insert by auto also have \\ = (\x\insert x F. Proj (J x))\ by (simp add: insert.IH insert.hyps(1) insert.hyps(2) insert.prems) finally show ?case by - qed lemma is_Proj_reduces_norm: fixes P :: \'a::complex_inner \\<^sub>C\<^sub>L 'a\ assumes \is_Proj P\ shows \norm (P *\<^sub>V \) \ norm \\ apply (rule is_projection_on_reduces_norm[where M=\range P\]) using assms is_Proj.rep_eq is_projection_on_image by blast (simp add: Proj_range_closed assms closed_csubspace.intro) lemma norm_Proj_apply: \norm (Proj T *\<^sub>V \) = norm \ \ \ \ space_as_set T\ proof (rule iffI) show \norm (Proj T *\<^sub>V \) = norm \\ if \\ \ space_as_set T\ by (simp add: Proj_fixes_image that) assume assm: \norm (Proj T *\<^sub>V \) = norm \\ have \_decomp: \\ = Proj T \ + Proj (-T) \\ by (simp add: Proj_ortho_compl cblinfun.real.diff_left) have \(norm (Proj (-T) \))\<^sup>2 = (norm (Proj T \))\<^sup>2 + (norm (Proj (-T) \))\<^sup>2 - (norm (Proj T \))\<^sup>2\ by auto also have \\ = (norm (Proj T \ + Proj (-T) \))\<^sup>2 - (norm (Proj T \))\<^sup>2\ apply (subst pythagorean_theorem) apply (metis (no_types, lifting) Proj_idempotent \_decomp add_cancel_right_left adj_Proj cblinfun.real.add_right cblinfun_apply_cblinfun_compose cinner_adj_left cinner_zero_left) by simp also with \_decomp have \\ = (norm \)\<^sup>2 - (norm (Proj T \))\<^sup>2\ by metis also with assm have \\ = 0\ by simp finally have \norm (Proj (-T) \) = 0\ by auto with \_decomp have \\ = Proj T \\ by auto then show \\ \ space_as_set T\ by (metis Proj_range cblinfun_apply_in_image) qed lemma norm_Proj_apply_1: \norm \ = 1 \ norm (Proj T *\<^sub>V \) = 1 \ \ \ space_as_set T\ using norm_Proj_apply by metis lemma norm_is_Proj_nonzero: \norm P = 1\ if \is_Proj P\ and \P \ 0\ for P :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ proof (rule antisym) show \norm P \ 1\ by (simp add: norm_is_Proj that(1)) from \P \ 0\ have \range P \ 0\ by (metis cblinfun_eq_0_on_UNIV_span complex_vector.span_UNIV rangeI set_zero singletonD) then obtain \ where \\ \ range P\ and \\ \ 0\ by force then have \P \ = \\ using is_Proj.rep_eq is_projection_on_fixes_image is_projection_on_image that(1) by blast then show \norm P \ 1\ apply (rule_tac cblinfun_norm_geqI[of _ _ \]) using \\ \ 0\ by simp qed lemma Proj_compose_cancelI: assumes \A *\<^sub>S \ \ S\ shows \Proj S o\<^sub>C\<^sub>L A = A\ apply (rule cblinfun_eqI) proof - fix x have \(Proj S o\<^sub>C\<^sub>L A) *\<^sub>V x = Proj S *\<^sub>V (A *\<^sub>V x)\ by simp also have \\ = A *\<^sub>V x\ apply (rule Proj_fixes_image) using assms cblinfun_apply_in_image less_eq_ccsubspace.rep_eq by blast finally show \(Proj S o\<^sub>C\<^sub>L A) *\<^sub>V x = A *\<^sub>V x\ by - qed lemma space_as_setI_via_Proj: assumes \Proj M *\<^sub>V x = x\ shows \x \ space_as_set M\ using assms norm_Proj_apply by fastforce lemma unitary_image_ortho_compl: \ \Logically, this lemma belongs in an earlier section but its proof uses projectors.\ fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes [simp]: \unitary U\ shows \U *\<^sub>S (- A) = - (U *\<^sub>S A)\ proof - have \Proj (U *\<^sub>S (- A)) = sandwich U (Proj (- A))\ by (simp add: Proj_sandwich) also have \\ = sandwich U (id_cblinfun - Proj A)\ by (simp add: Proj_ortho_compl) also have \\ = id_cblinfun - sandwich U (Proj A)\ by (metis assms cblinfun.diff_right sandwich_isometry_id unitary_twosided_isometry) also have \\ = id_cblinfun - Proj (U *\<^sub>S A)\ by (simp add: Proj_sandwich) also have \\ = Proj (- (U *\<^sub>S A))\ by (simp add: Proj_ortho_compl) finally show ?thesis by (simp add: Proj_inj) qed lemma Proj_on_image [simp]: \Proj S *\<^sub>S S = S\ by (metis Proj_idempotent Proj_range cblinfun_compose_image) subsection \Kernel / eigenspaces\ -lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector +lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a ccsubspace" - is "\ f. f -` {0}" + 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" by transfer auto lemma kernel_id[simp]: "kernel id_cblinfun = 0" by transfer 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 by transfer auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms by transfer 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 by transfer auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms by transfer auto lemma kernel_Proj[simp]: \kernel (Proj S) = - S\ apply transfer apply auto apply (metis diff_0_right is_projection_on_iff_orthog projection_is_projection_on') by (simp add: complex_vector.subspace_0 projection_eqI) lemma orthogonal_projectors_orthogonal_spaces: \ \Logically belongs in section "Projectors".\ fixes S T :: \'a::chilbert_space ccsubspace\ shows \orthogonal_spaces S T \ Proj S o\<^sub>C\<^sub>L Proj T = 0\ proof (intro ballI iffI) assume \Proj S o\<^sub>C\<^sub>L Proj T = 0\ then have \is_orthogonal x y\ if \x \ space_as_set S\ \y \ space_as_set T\ for x y by (metis (no_types, opaque_lifting) Proj_fixes_image adj_Proj cblinfun.zero_left cblinfun_apply_cblinfun_compose cinner_adj_right cinner_zero_right that(1) that(2)) then show \orthogonal_spaces S T\ by (simp add: orthogonal_spaces_def) next assume \orthogonal_spaces S T\ then have \S \ - T\ by (simp add: orthogonal_spaces_leq_compl) then show \Proj S o\<^sub>C\<^sub>L Proj T = 0\ by (metis (no_types, opaque_lifting) Proj_range adj_Proj adj_cblinfun_compose basic_trans_rules(31) cblinfun.zero_left cblinfun_apply_cblinfun_compose cblinfun_apply_in_image cblinfun_eqI kernel_Proj kernel_memberD less_eq_ccsubspace.rep_eq) qed lemma cblinfun_compose_Proj_kernel[simp]: \a o\<^sub>C\<^sub>L Proj (kernel a) = 0\ apply (rule cblinfun_eqI) by simp (metis Proj_range cblinfun_apply_in_image kernel_memberD) lemma kernel_compl_adj_range: shows \kernel a = - (a* *\<^sub>S top)\ proof (rule ccsubspace_eqI) fix x have \x \ space_as_set (kernel a) \ a x = 0\ by transfer simp also have \a x = 0 \ (\y. is_orthogonal y (a x))\ by (metis cinner_gt_zero_iff cinner_zero_right) also have \\ \ (\y. is_orthogonal (a* *\<^sub>V y) x)\ by (simp add: cinner_adj_left) also have \\ \ x \ space_as_set (- (a* *\<^sub>S top))\ by transfer (metis (mono_tags, opaque_lifting) UNIV_I image_iff is_orthogonal_sym orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') finally show \x \ space_as_set (kernel a) \ x \ space_as_set (- (a* *\<^sub>S top))\ by - qed lemma kernel_apply_self: \A *\<^sub>S kernel A = 0\ proof transfer fix A :: \'b \ 'a\ assume \bounded_clinear A\ then have \A 0 = 0\ by (simp add: bounded_clinear_def complex_vector.linear_0) then have \A ` A -` {0} = {0}\ by fastforce then show \closure (A ` A -` {0}) = {0}\ by auto qed lemma leq_kernel_iff: shows \A \ kernel B \ B *\<^sub>S A = 0\ proof (rule iffI) assume \A \ kernel B\ then have \B *\<^sub>S A \ B *\<^sub>S kernel B\ by (simp add: cblinfun_image_mono) also have \\ = 0\ by (simp add: kernel_apply_self) finally show \B *\<^sub>S A = 0\ by (simp add: bot.extremum_unique) next assume \B *\<^sub>S A = 0\ then show \A \ kernel B\ apply transfer by (metis closure_subset image_subset_iff_subset_vimage) qed lemma cblinfun_image_kernel: assumes \C *\<^sub>S A *\<^sub>S kernel B \ kernel B\ assumes \A o\<^sub>C\<^sub>L C = id_cblinfun\ shows \A *\<^sub>S kernel B = kernel (B o\<^sub>C\<^sub>L C)\ proof (rule antisym) show \A *\<^sub>S kernel B \ kernel (B o\<^sub>C\<^sub>L C)\ using assms(1) by (simp add: leq_kernel_iff cblinfun_compose_image) show \kernel (B o\<^sub>C\<^sub>L C) \ A *\<^sub>S kernel B\ proof (insert assms(2), transfer, intro subsetI) fix A :: \'a \ 'b\ and B :: \'a \ 'c\ and C :: \'b \ 'a\ and x assume \x \ (B \ C) -` {0}\ then have BCx: \B (C x) = 0\ by simp assume \A \ C = (\x. x)\ then have \x = A (C x)\ apply (simp add: o_def) by metis then show \x \ closure (A ` B -` {0})\ using \B (C x) = 0\ closure_subset by fastforce qed qed lemma cblinfun_image_kernel_unitary: assumes \unitary U\ shows \U *\<^sub>S kernel B = kernel (B o\<^sub>C\<^sub>L U*)\ apply (rule cblinfun_image_kernel) using assms by (auto simp flip: cblinfun_compose_image) lemma kernel_cblinfun_compose: assumes \kernel B = 0\ shows \kernel A = kernel (B o\<^sub>C\<^sub>L A)\ using assms apply transfer by auto lemma eigenspace_0[simp]: \eigenspace 0 A = kernel A\ by (simp add: eigenspace_def) lemma kernel_isometry: \kernel U = 0\ if \isometry U\ by (simp add: kernel_compl_adj_range range_adjoint_isometry that) lemma cblinfun_image_eigenspace_isometry: assumes [simp]: \isometry A\ and \c \ 0\ shows \A *\<^sub>S eigenspace c B = eigenspace c (sandwich A B)\ proof (rule antisym) show \A *\<^sub>S eigenspace c B \ eigenspace c (sandwich A B)\ proof (unfold cblinfun_image_def2, rule ccspan_leqI, rule subsetI) fix x assume \x \ (*\<^sub>V) A ` space_as_set (eigenspace c B)\ then obtain y where x_def: \x = A y\ and \y \ space_as_set (eigenspace c B)\ by auto then have \B y = c *\<^sub>C y\ by (simp add: eigenspace_memberD) then have \sandwich A B x = c *\<^sub>C x\ apply (simp add: sandwich_apply x_def cblinfun_compose_assoc flip: cblinfun_apply_cblinfun_compose) by (simp add: cblinfun.scaleC_right) then show \x \ space_as_set (eigenspace c (sandwich A B))\ by (simp add: eigenspace_memberI) qed show \eigenspace c (sandwich A *\<^sub>V B) \ A *\<^sub>S eigenspace c B\ proof (rule ccsubspace_leI_unit) fix x assume \x \ space_as_set (eigenspace c (sandwich A B))\ then have *: \sandwich A B x = c *\<^sub>C x\ by (simp add: eigenspace_memberD) then have \c *\<^sub>C x \ range A\ apply (simp add: sandwich_apply) by (metis rangeI) then have \(inverse c * c) *\<^sub>C x \ range A\ apply (simp flip: scaleC_scaleC) by (metis (no_types, lifting) cblinfun.scaleC_right rangeE rangeI) with \c \ 0\ have \x \ range A\ by simp then obtain y where x_def: \x = A y\ by auto have \B *\<^sub>V y = A* *\<^sub>V sandwich A B x\ apply (simp add: sandwich_apply x_def) by (metis assms cblinfun_apply_cblinfun_compose id_cblinfun.rep_eq isometryD) also have \\ = c *\<^sub>C y\ apply (simp add: * cblinfun.scaleC_right) apply (simp add: x_def) by (metis assms(1) cblinfun_apply_cblinfun_compose id_cblinfun_apply isometry_def) finally have \y \ space_as_set (eigenspace c B)\ by (simp add: eigenspace_memberI) then show \x \ space_as_set (A *\<^sub>S eigenspace c B) \ by (simp add: x_def cblinfun_apply_in_image') qed qed lemma cblinfun_image_eigenspace_unitary: assumes [simp]: \unitary A\ shows \A *\<^sub>S eigenspace c B = eigenspace c (sandwich A B)\ apply (cases \c = 0\) apply (simp add: sandwich_apply cblinfun_image_kernel_unitary kernel_isometry cblinfun_compose_assoc flip: kernel_cblinfun_compose) by (simp add: cblinfun_image_eigenspace_isometry) lemma kernel_member_iff: \x \ space_as_set (kernel A) \ A *\<^sub>V x = 0\ using kernel_memberD kernel_memberI by blast lemma kernel_square[simp]: \kernel (A* o\<^sub>C\<^sub>L A) = kernel A\ proof (intro ccsubspace_eqI iffI) fix x assume \x \ space_as_set (kernel A)\ then show \x \ space_as_set (kernel (A* o\<^sub>C\<^sub>L A))\ by (simp add: kernel.rep_eq) next fix x assume \x \ space_as_set (kernel (A* o\<^sub>C\<^sub>L A))\ then have \A* *\<^sub>V A *\<^sub>V x = 0\ by (simp add: kernel.rep_eq) then have \(A *\<^sub>V x) \\<^sub>C (A *\<^sub>V x) = 0\ by (metis cinner_adj_right cinner_zero_right) then have \A *\<^sub>V x = 0\ by auto then show \x \ space_as_set (kernel A)\ by (simp add: kernel.rep_eq) qed subsection \Partial isometries\ definition partial_isometry where \partial_isometry A \ (\h \ space_as_set (- kernel A). norm (A h) = norm h)\ lemma partial_isometryI: assumes \\h. h \ space_as_set (- kernel A) \ norm (A h) = norm h\ shows \partial_isometry A\ using assms partial_isometry_def by blast lemma fixes A :: \'a :: chilbert_space \\<^sub>C\<^sub>L 'b :: complex_normed_vector\ assumes iso: \\\. \ \ space_as_set V \ norm (A *\<^sub>V \) = norm \\ assumes zero: \\\. \ \ space_as_set (- V) \ A *\<^sub>V \ = 0\ shows partial_isometryI': \partial_isometry A\ and partial_isometry_initial: \kernel A = - V\ proof - from zero have \- V \ kernel A\ by (simp add: kernel_memberI less_eq_ccsubspace.rep_eq subsetI) moreover have \kernel A \ -V\ by (smt (verit, ccfv_threshold) Proj_ortho_compl Proj_range assms(1) cblinfun.diff_left cblinfun.diff_right cblinfun_apply_in_image cblinfun_id_cblinfun_apply ccsubspace_leI kernel_Proj kernel_memberD kernel_memberI norm_eq_zero ortho_involution subsetI zero) ultimately show kerA: \kernel A = -V\ by simp show \partial_isometry A\ apply (rule partial_isometryI) by (simp add: kerA iso) qed lemma Proj_partial_isometry[simp]: \partial_isometry (Proj S)\ apply (rule partial_isometryI) by (simp add: Proj_fixes_image) lemma is_Proj_partial_isometry: \is_Proj P \ partial_isometry P\ for P :: \_ :: chilbert_space \\<^sub>C\<^sub>L _\ by (metis Proj_on_own_range Proj_partial_isometry) lemma isometry_partial_isometry: \isometry P \ partial_isometry P\ by (simp add: isometry_preserves_norm partial_isometry_def) lemma unitary_partial_isometry: \unitary P \ partial_isometry P\ using isometry_partial_isometry unitary_isometry by blast lemma norm_partial_isometry: fixes A :: \'a :: chilbert_space \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \partial_isometry A\ and \A \ 0\ shows \norm A = 1\ proof - from \A \ 0\ have \- (kernel A) \ 0\ by (metis cblinfun_eqI diff_zero id_cblinfun_apply kernel_id kernel_memberD ortho_involution orthog_proj_exists orthogonal_complement_closed_subspace uminus_ccsubspace.rep_eq zero_cblinfun.rep_eq) then obtain h where \h \ space_as_set (- kernel A)\ and \h \ 0\ by (metis cblinfun_id_cblinfun_apply ccsubspace_eqI closed_csubspace.subspace complex_vector.subspace_0 kernel_id kernel_memberD kernel_memberI orthogonal_complement_closed_subspace uminus_ccsubspace.rep_eq) with \partial_isometry A\ have \norm (A h) = norm h\ using partial_isometry_def by blast then have \norm A \ 1\ by (smt (verit) \h \ 0\ mult_cancel_right1 mult_left_le_one_le norm_cblinfun norm_eq_zero norm_ge_zero) have \norm A \ 1\ proof (rule norm_cblinfun_bound) show \0 \ (1::real)\ by simp fix \ :: 'a define g h where \g = Proj (kernel A) \\ and \h = Proj (- kernel A) \\ have \A g = 0\ by (metis Proj_range cblinfun_apply_in_image g_def kernel_memberD) moreover from \partial_isometry A\ have \norm (A h) = norm h\ by (metis Proj_range cblinfun_apply_in_image h_def partial_isometry_def) ultimately have \norm (A \) = norm h\ by (simp add: Proj_ortho_compl cblinfun.diff_left cblinfun.diff_right g_def h_def) also have \norm h \ norm \\ by (smt (verit, del_insts) h_def mult_left_le_one_le norm_Proj_leq1 norm_cblinfun norm_ge_zero) finally show \norm (A *\<^sub>V \) \ 1 * norm \\ by simp qed from \norm A \ 1\ and \norm A \ 1\ show \norm A = 1\ by auto qed lemma partial_isometry_adj_a_o_a: assumes \partial_isometry a\ shows \a* o\<^sub>C\<^sub>L a = Proj (- kernel a)\ proof (rule cblinfun_cinner_eqI) define P where \P = Proj (- kernel a)\ have aP: \a o\<^sub>C\<^sub>L P = a\ by (auto intro!: simp: cblinfun_compose_minus_right P_def Proj_ortho_compl) have is_Proj_P[simp]: \is_Proj P\ by (simp add: P_def) fix \ :: 'a have \\ \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V \) = a \ \\<^sub>C a \\ by (simp add: cinner_adj_right) also have \\ = a (P \) \\<^sub>C a (P \)\ by (metis aP cblinfun_apply_cblinfun_compose) also have \\ = P \ \\<^sub>C P \\ by (metis P_def Proj_range assms cblinfun_apply_in_image cdot_square_norm partial_isometry_def) also have \\ = \ \\<^sub>C P \\ by (simp flip: cinner_adj_right add: is_proj_selfadj is_Proj_idempotent[THEN simp_a_oCL_b']) finally show \\ \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V \) = \ \\<^sub>C P \\ by - qed lemma partial_isometry_square_proj: \is_Proj (a* o\<^sub>C\<^sub>L a)\ if \partial_isometry a\ by (simp add: partial_isometry_adj_a_o_a that) lemma partial_isometry_adj[simp]: \partial_isometry (a*)\ if \partial_isometry a\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ proof - have ran_ker: \a *\<^sub>S top = - kernel (a*)\ by (simp add: kernel_compl_adj_range) have \norm (a* *\<^sub>V h) = norm h\ if \h \ range a\ for h proof - from that obtain x where h: \h = a x\ by auto have \norm (a* *\<^sub>V h) = norm (a* *\<^sub>V a *\<^sub>V x)\ by (simp add: h) also have \\ = norm (Proj (- kernel a) *\<^sub>V x)\ by (simp add: \partial_isometry a\ partial_isometry_adj_a_o_a simp_a_oCL_b') also have \\ = norm (a *\<^sub>V Proj (- kernel a) *\<^sub>V x)\ by (metis Proj_range \partial_isometry a\ cblinfun_apply_in_image partial_isometry_def) also have \\ = norm (a *\<^sub>V x)\ by (smt (verit, best) Proj_idempotent \partial_isometry a\ adj_Proj cblinfun_apply_cblinfun_compose cinner_adj_right cnorm_eq partial_isometry_adj_a_o_a) also have \\ = norm h\ using h by auto finally show ?thesis by - qed then have norm_pres: \norm (a* *\<^sub>V h) = norm h\ if \h \ closure (range a)\ for h using that apply (rule on_closure_eqI) by assumption (intro continuous_intros)+ show ?thesis apply (rule partial_isometryI) by (auto simp: cblinfun_image.rep_eq norm_pres simp flip: ran_ker) qed 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 \invertible_cblinfun A \ (\B. 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)\ + \cblinfun_inv A = (if invertible_cblinfun A then SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun else 0)\ lemma cblinfun_inv_left: assumes \invertible_cblinfun A\ shows \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ - unfolding cblinfun_inv_def apply (rule someI_ex) + apply (simp add: assms cblinfun_inv_def) + apply (rule someI_ex) using assms by (simp add: invertible_cblinfun_def) lemma inv_cblinfun_invertible: \iso_cblinfun A \ invertible_cblinfun A\ by (auto simp: iso_cblinfun_def invertible_cblinfun_def) lemma cblinfun_inv_right: assumes \iso_cblinfun A\ shows \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) + by (simp add: assms cblinfun_inv_left inv_cblinfun_invertible) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB show \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 inv_cblinfun_invertible cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) lemma iso_cblinfun_unitary: \unitary A \ iso_cblinfun A\ using iso_cblinfun_def unitary_def by blast lemma invertible_cblinfun_isometry: \isometry A \ invertible_cblinfun A\ using invertible_cblinfun_def isometryD by blast lemma summable_cblinfun_apply_invertible: assumes \invertible_cblinfun A\ shows \(\x. A *\<^sub>V g x) summable_on S \ g summable_on S\ proof (rule iffI) assume \g summable_on S\ then show \(\x. A *\<^sub>V g x) summable_on S\ by (rule summable_on_cblinfun_apply) next assume \(\x. A *\<^sub>V g x) summable_on S\ then have \(\x. cblinfun_inv A *\<^sub>V A *\<^sub>V g x) summable_on S\ by (rule summable_on_cblinfun_apply) then show \g summable_on S\ by (simp add: cblinfun_inv_left assms flip: cblinfun_apply_cblinfun_compose) qed lemma infsum_cblinfun_apply_invertible: assumes \invertible_cblinfun A\ - shows \infsum (\x. A *\<^sub>V g x) S = A *\<^sub>V (infsum g S)\ + shows \(\\<^sub>\x\S. A *\<^sub>V g x) = A *\<^sub>V (\\<^sub>\x\S. g x)\ proof (cases \g summable_on S\) case True then show ?thesis by (rule infsum_cblinfun_apply) next case False then have \\ (\x. A *\<^sub>V g x) summable_on S\ using assms by (simp add: summable_cblinfun_apply_invertible) with False show ?thesis by (simp add: infsum_not_exists) qed 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 \\<^sub>C B) = cnj (B \\<^sub>C A)" unfolding cinner_cblinfun_def by auto show "(A + B) \\<^sub>C C = (A \\<^sub>C C) + (B \\<^sub>C C)" by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) show "(c *\<^sub>C A \\<^sub>C B) = cnj c * (A \\<^sub>C B)" by (simp add: cblinfun.scaleC_left cinner_cblinfun_def) show "0 \ (A \\<^sub>C 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 \\<^sub>C A) = 0) = (A = 0)" by (auto simp: cinner_cblinfun_def) show "norm A = sqrt (cmod (A \\<^sub>C A))" unfolding cinner_cblinfun_def by transfer (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]" definition \canonical_basis_length_cblinfun (_ :: ('a \\<^sub>C\<^sub>L 'b) itself) = (1::nat)\ 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" by transfer 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)" by transfer 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)" by transfer 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" by transfer (simp add: o_def one_dim_inverse) show \canonical_basis_length TYPE('a \\<^sub>C\<^sub>L 'b) = length (canonical_basis :: ('a \\<^sub>C\<^sub>L 'b) list)\ by (simp add: canonical_basis_length_cblinfun_def canonical_basis_cblinfun_def) qed end lemma id_cblinfun_eq_1[simp]: \id_cblinfun = 1\ by transfer 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" by transfer 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" by transfer simp 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 = one_dim_iso (x *\<^sub>C y)\ by (metis of_complex_def cblinfun.scaleC_right one_cblinfun.rep_eq scaleC_cblinfun.rep_eq) lemma cblinfun_compose_1_left[simp]: \1 o\<^sub>C\<^sub>L x = x\ by transfer auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ by transfer 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) lemma is_onb_one_dim[simp]: \norm x = 1 \ is_onb {x}\ for x :: \_ :: one_dim\ by (auto simp: is_onb_def intro!: ccspan_one_dim) lemma one_dim_iso_cblinfun_comp: \one_dim_iso (a o\<^sub>C\<^sub>L b) = of_complex (cinner (a* *\<^sub>V 1) (b *\<^sub>V 1))\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::one_dim\ and b :: \'c::one_dim \\<^sub>C\<^sub>L 'a\ by (simp add: cinner_adj_left cinner_cblinfun_def one_dim_iso_def) lemma one_dim_iso_cblinfun_apply[simp]: \one_dim_iso \ *\<^sub>V \ = one_dim_iso (one_dim_iso \ *\<^sub>C \)\ by (smt (verit) cblinfun.scaleC_left one_cblinfun.rep_eq one_dim_iso_of_one one_dim_iso_scaleC one_dim_scaleC_1) 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" by transfer 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_def_heterogenous: \A \ B = +definition less_eq_cblinfun_def_heterogenous: \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 \(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 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\ by atomize_elim (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 cblinfun_leI: assumes \\x. norm x = 1 \ x \\<^sub>C (A *\<^sub>V x) \ x \\<^sub>C (B *\<^sub>V x)\ shows \A \ B\ proof (unfold less_eq_cblinfun_def, intro allI, case_tac \\ = 0\) fix \ :: 'a assume \\ = 0\ then show \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ by simp next fix \ :: 'a assume \\ \ 0\ define \ where \\ = \ /\<^sub>R norm \\ have \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ apply (rule assms) unfolding \_def by (simp add: \\ \ 0\) with \\ \ 0\ show \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ unfolding \_def by (smt (verit) cinner_adj_left cinner_scaleR_left cinner_simps(6) complex_of_real_nn_iff mult_cancel_right1 mult_left_mono norm_eq_zero norm_ge_zero of_real_1 right_inverse scaleR_scaleC scaleR_scaleR) qed lemma positive_cblinfunI: \A \ 0\ if \\x. norm x = 1 \ cinner x (A *\<^sub>V x) \ 0\ apply (rule cblinfun_leI) using that by simp lemma less_eq_scaled_id_norm: assumes \norm A \ c\ and \A = A*\ shows \A \ complex_of_real c *\<^sub>C id_cblinfun\ proof - have \x \\<^sub>C (A *\<^sub>V x) \ complex_of_real c\ if \norm x = 1\ for x proof - have \norm (x \\<^sub>C (A *\<^sub>V x)) \ norm (A *\<^sub>V x)\ by (metis complex_inner_class.Cauchy_Schwarz_ineq2 mult_cancel_right1 that) also have \\ \ norm A\ by (metis more_arith_simps(6) norm_cblinfun that) also have \\ \ c\ by (rule assms) finally have \norm (x \\<^sub>C (A *\<^sub>V x)) \ c\ by - moreover have \x \\<^sub>C (A *\<^sub>V x) \ \\ by (metis assms(2) cinner_hermitian_real) ultimately show ?thesis by (metis cnorm_le_square complex_of_real_cmod complex_of_real_mono complex_of_real_nn_iff dual_order.trans reals_zero_comparable) qed then show ?thesis by (smt (verit) cblinfun.scaleC_left cblinfun_id_cblinfun_apply cblinfun_leI cinner_scaleC_right cnorm_eq_1 mult_cancel_left2) qed (* 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 - 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 lemma op_square_nondegenerate: \a = 0\ if \a* o\<^sub>C\<^sub>L a = 0\ proof (rule cblinfun_eq_0_on_UNIV_span[where basis=UNIV]; simp) fix s from that have \s \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V s) = 0\ by simp then have \(a *\<^sub>V s) \\<^sub>C (a *\<^sub>V s) = 0\ by (simp add: cinner_adj_right) then show \a *\<^sub>V s = 0\ by simp qed lemma comparable_hermitean: assumes \a \ b\ assumes \a* = a\ shows \b* = b\ by (smt (verit, best) assms(1) assms(2) cinner_hermitian_real cinner_real_hermiteanI comparable complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma comparable_hermitean': assumes \a \ b\ assumes \b* = b\ shows \a* = a\ by (smt (verit, best) assms(1) assms(2) cinner_hermitian_real cinner_real_hermiteanI comparable complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma Proj_mono: \Proj S \ Proj T \ S \ T\ proof (rule iffI) assume \S \ T\ define D where \D = Proj T - Proj S\ from \S \ T\ have TS_S[simp]: \Proj T o\<^sub>C\<^sub>L Proj S = Proj S\ by (smt (verit, ccfv_threshold) Proj_idempotent Proj_range cblinfun_apply_cblinfun_compose cblinfun_apply_in_image cblinfun_eqI cblinfun_fixes_range less_eq_ccsubspace.rep_eq subset_iff) then have ST_S[simp]: \Proj S o\<^sub>C\<^sub>L Proj T = Proj S\ by (metis adj_Proj adj_cblinfun_compose) have \D* o\<^sub>C\<^sub>L D = D\ by (simp add: D_def cblinfun_compose_minus_left cblinfun_compose_minus_right adj_minus adj_Proj) then have \D \ 0\ by (metis positive_cblinfun_squareI) then show \Proj S \ Proj T\ by (simp add: D_def) next assume PS_PT: \Proj S \ Proj T\ show \S \ T\ proof (rule ccsubspace_leI_unit) fix \ assume \\ \ space_as_set S\ and [simp]: \norm \ = 1\ then have \1 = norm (Proj S *\<^sub>V \)\ by (simp add: Proj_fixes_image) also from PS_PT have \\ \ norm (Proj T *\<^sub>V \)\ by (metis (no_types, lifting) Proj_idempotent adj_Proj cblinfun_apply_cblinfun_compose cinner_adj_left cnorm_le less_eq_cblinfun_def) also have \\ \ 1\ by (metis Proj_is_Proj \norm \ = 1\ is_Proj_reduces_norm) ultimately have \norm (Proj T *\<^sub>V \) = 1\ by auto then show \\ \ space_as_set T\ by (simp add: norm_Proj_apply_1) qed qed 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_compose[simp]: "A o\<^sub>C\<^sub>L (vector_to_cblinfun \) = vector_to_cblinfun (A *\<^sub>V \)" 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\ by transfer (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 by (intro clinear.scaleC bounded_clinear.clinear bounded_clinear_vector_to_cblinfun) lemma vector_to_cblinfun_apply_one_dim[simp]: shows "vector_to_cblinfun \ *\<^sub>V \ = one_dim_iso \ *\<^sub>C \" by transfer (rule refl) lemma vector_to_cblinfun_one_dim_iso[simp]: \vector_to_cblinfun = one_dim_iso\ by (auto intro!: ext cblinfun_eqI) 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_compose) lemma image_vector_to_cblinfun[simp]: "vector_to_cblinfun x *\<^sub>S \ = ccspan {x}" \ \Not that the general case \<^term>\vector_to_cblinfun x *\<^sub>S S\ can be handled by using that \S = \\ or \S = \\ by @{thm [source] one_dim_ccsubspace_all_or_nothing}\ 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 \)) = (\ \\<^sub>C \) *\<^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 \) *\<^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 lemma image_vector_to_cblinfun_adj: assumes \\ \ space_as_set (- S)\ shows \(vector_to_cblinfun \)* *\<^sub>S S = \\ proof - from assms obtain \ where \\ \ space_as_set S\ and \\ is_orthogonal \ \\ by (metis orthogonal_complementI uminus_ccsubspace.rep_eq) have \((vector_to_cblinfun \)* *\<^sub>S S :: 'b ccsubspace) \ (vector_to_cblinfun \)* *\<^sub>S ccspan {\}\ (is \_ \ \\) by (simp add: \\ \ space_as_set S\ cblinfun_image_mono ccspan_leqI) also have \\ = ccspan {(vector_to_cblinfun \)* *\<^sub>V \}\ by (auto simp: cblinfun_image_ccspan) also have \\ = ccspan {of_complex (\ \\<^sub>C \)}\ by auto also have \\ > \\ by (simp add: \\ \\<^sub>C \ \ 0\ flip: bot.not_eq_extremum ) finally(dual_order.strict_trans1) show ?thesis using one_dim_ccsubspace_all_or_nothing bot.not_eq_extremum by auto qed lemma image_vector_to_cblinfun_adj': assumes \\ \ 0\ shows \(vector_to_cblinfun \)* *\<^sub>S \ = \\ apply (rule image_vector_to_cblinfun_adj) using assms by simp subsection \Rank-1 operators / butterflies\ definition rank1 where \rank1 A \ (\\\0. A *\<^sub>S \ = ccspan {\})\ definition "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 flip: vector_to_cblinfun_cblinfun_compose) lemma cblinfun_comp_butterfly: "a o\<^sub>C\<^sub>L butterfly \ \ = butterfly (a *\<^sub>V \) \" unfolding butterfly_def by (simp add: cblinfun_compose_assoc flip: vector_to_cblinfun_cblinfun_compose) lemma butterfly_apply[simp]: "butterfly \ \' *\<^sub>V \ = (\' \\<^sub>C \) *\<^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 \\<^sub>C \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 butterfly_is_rank1: assumes \\ \ 0\ shows \butterfly \ \ *\<^sub>S \ = ccspan {\}\ using assms by (simp add: butterfly_def cblinfun_compose_image image_vector_to_cblinfun_adj') lemma rank1_is_butterfly: assumes \A *\<^sub>S \ = ccspan {\::_::chilbert_space}\ shows \\\. A = butterfly \ \\ proof (rule exI[of _ \A* *\<^sub>V (\ /\<^sub>R (norm \)\<^sup>2)\], rule cblinfun_eqI) fix \ :: 'b from assms have \A *\<^sub>V \ \ space_as_set (ccspan {\})\ by (simp flip: assms) then obtain c where c: \A *\<^sub>V \ = c *\<^sub>C \\ apply atomize_elim apply (auto simp: ccspan.rep_eq) by (metis complex_vector.span_breakdown_eq complex_vector.span_empty eq_iff_diff_eq_0 singletonD) have \A *\<^sub>V \ = butterfly \ (\ /\<^sub>R (norm \)\<^sup>2) *\<^sub>V (A *\<^sub>V \)\ apply (auto simp: c simp flip: scaleC_scaleC) by (metis cinner_eq_zero_iff divideC_field_simps(1) power2_norm_eq_cinner scaleC_left_commute scaleC_zero_right) also have \\ = (butterfly \ (\ /\<^sub>R (norm \)\<^sup>2) o\<^sub>C\<^sub>L A) *\<^sub>V \\ by simp also have \\ = butterfly \ (A* *\<^sub>V (\ /\<^sub>R (norm \)\<^sup>2)) *\<^sub>V \\ by (simp add: cinner_adj_left) finally show \A *\<^sub>V \ = \\ by - qed lemma zero_not_rank1[simp]: \\ rank1 0\ unfolding rank1_def by auto (metis ccspan_superset insert_not_empty singleton_insert_inj_eq space_as_set_bot subset_singletonD) lemma rank1_iff_butterfly: \rank1 A \ (\\ \. A = butterfly \ \) \ A \ 0\ for A :: \_::complex_inner \\<^sub>C\<^sub>L _::chilbert_space\ proof (rule iffI) assume \rank1 A\ then obtain \ where \A *\<^sub>S \ = ccspan {\}\ using rank1_def by auto then have \\\. A = butterfly \ \\ by (rule rank1_is_butterfly) moreover from \rank1 A\ have \A \ 0\ by auto ultimately show \(\\ \. A = butterfly \ \) \ A \ 0\ by auto next assume asm: \(\\ \. A = butterfly \ \) \ A \ 0\ then obtain \ \ where A: \A = butterfly \ \\ by auto from asm have \A \ 0\ by simp with A have \\ \ 0\ and \\ \ 0\ by auto then have \butterfly \ \ *\<^sub>S \ = ccspan {\}\ by (rule_tac butterfly_is_rank1) with A \\ \ 0\ show \rank1 A\ by (auto intro!: exI[of _ \] simp: rank1_def) qed lemma butterfly_if_rank1: \(\\ \. A = butterfly \ \) \ rank1 A \ A = 0\ for A :: \_::complex_inner \\<^sub>C\<^sub>L _::chilbert_space\ by (metis butterfly_0_left rank1_iff_butterfly) 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 (\ \\<^sub>C 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 \\<^sub>C x) / (x \\<^sub>C x)" have "(x \\<^sub>C 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 \\<^sub>C 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 \\<^sub>C x) *\<^sub>C x = selfbutter x *\<^sub>V x\ \selfbutter y *\<^sub>V x = (y \\<^sub>C 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_sgn_eq_proj: shows "selfbutter (sgn x) = proj x" proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \selfbutter (sgn x) = proj (sgn x)\ by (simp add: butterfly_eq_proj norm_sgn) also have \ccspan {sgn x} = ccspan {x}\ by (metis ccspan_singleton_scaleC scaleC_eq_0_iff scaleR_scaleC sgn_div_norm sgn_zero_iff) finally show ?thesis by - 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 lemma sum_butterfly_is_Proj: assumes \is_ortho_set E\ assumes \\e. e\E \ norm e = 1\ shows \is_Proj (\e\E. butterfly e e)\ proof (cases \finite E\) case True show ?thesis proof (rule is_Proj_I) show \(\e\E. butterfly e e)* = (\e\E. butterfly e e)\ by (simp add: sum_adj) have ortho: \f \ e \ e \ E \ f \ E \ is_orthogonal f e\ for f e by (meson assms(1) is_ortho_set_def) have unit: \e \\<^sub>C e = 1\ if \e \ E\ for e using assms(2) cnorm_eq_1 that by blast have *: \(\f\E. (f \\<^sub>C e) *\<^sub>C butterfly f e) = butterfly e e\ if \e \ E\ for e apply (subst sum_single[where i=e]) by (auto intro!: simp: that ortho unit True) show \(\e\E. butterfly e e) o\<^sub>C\<^sub>L (\e\E. butterfly e e) = (\e\E. butterfly e e)\ by (auto simp: * cblinfun_compose_sum_right cblinfun_compose_sum_left) qed next case False then show ?thesis by simp qed lemma rank1_compose_left: \rank1 (a o\<^sub>C\<^sub>L b)\ if \rank1 b\ and \a o\<^sub>C\<^sub>L b \ 0\ proof - from \rank1 b\ obtain \ where \b *\<^sub>S \ = ccspan {\}\ using rank1_def by blast then have *: \(a o\<^sub>C\<^sub>L b) *\<^sub>S \ = ccspan {a \}\ by (metis cblinfun_assoc_left(2) cblinfun_image_ccspan image_empty image_insert) with \a o\<^sub>C\<^sub>L b \ 0\ have \a \ \ 0\ by auto with * show \rank1 (a o\<^sub>C\<^sub>L b)\ using rank1_def by blast qed lemma rank1_compose_right: \rank1 (a o\<^sub>C\<^sub>L b)\ if \rank1 a\ and \a o\<^sub>C\<^sub>L b \ 0\ proof - from \rank1 a\ obtain \ where \a *\<^sub>S \ = ccspan {\}\ and \\ \ 0\ using rank1_def by blast then have *: \(a o\<^sub>C\<^sub>L b) *\<^sub>S \ \ ccspan {\}\ by (metis cblinfun_assoc_left(2) cblinfun_image_mono top_greatest) from \a o\<^sub>C\<^sub>L b \ 0\ obtain \ where \_ab: \\ \ space_as_set ((a o\<^sub>C\<^sub>L b) *\<^sub>S \)\ and \\ \ 0\ by (metis cblinfun.real.zero_left cblinfun_apply_in_image cblinfun_eqI) with * have \\ \ space_as_set (ccspan {\})\ using less_eq_ccsubspace.rep_eq by blast with \\ \ 0\ obtain c where \\ = c *\<^sub>C \\ and \c \ 0\ apply (simp add: ccspan.rep_eq) by (auto simp add: complex_vector.span_singleton) with \_ab have \(a o\<^sub>C\<^sub>L b) *\<^sub>S \ \ ccspan {\}\ by (metis ccspan_leqI ccspan_singleton_scaleC empty_subsetI insert_subset) with * have \(a o\<^sub>C\<^sub>L b) *\<^sub>S \ = ccspan {\}\ by fastforce with \\ \ 0\ show \rank1 (a o\<^sub>C\<^sub>L b)\ using rank1_def by blast qed lemma rank1_scaleC: \rank1 (c *\<^sub>C a)\ if \rank1 a\ and \c \ 0\ using rank1_compose_left[OF \rank1 a\, where a=\c *\<^sub>C id_cblinfun\] using that by force lemma rank1_uminus: \rank1 (-a)\ if \rank1 a\ using that rank1_scaleC[where c=\-1\ and a=a] by simp lemma rank1_uminus_iff[simp]: \rank1 (-a) \ rank1 a\ using rank1_uminus by force lemma rank1_adj: \rank1 (a*)\ if \rank1 a\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ by (metis adj_0 butterfly_adjoint rank1_iff_butterfly that) lemma rank1_adj_iff[simp]: \rank1 (a*) \ rank1 a\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ by (metis double_adj rank1_adj) lemma butterflies_sum_id_finite: \id_cblinfun = (\x\B. selfbutter x)\ if \is_onb B\ for B :: \'a :: {cfinite_dim, chilbert_space} set\ proof (rule cblinfun_eq_gen_eqI) from \is_onb B\ show \ccspan B = \\ by (simp add: is_onb_def) from \is_onb B\ have \cindependent B\ by (simp add: is_onb_def is_ortho_set_cindependent) then have [simp]: \finite B\ using cindependent_cfinite_dim_finite by blast from \is_onb B\ have 1: \j \ b \ j \ B \ is_orthogonal j b\ if \b \ B\ for j b using that by (auto simp add: is_onb_def is_ortho_set_def) from \is_onb B\ have 2: \b \\<^sub>C b = 1\ if \b \ B\ for b using that by (simp add: is_onb_def cnorm_eq_1) fix b assume \b \ B\ then show \id_cblinfun *\<^sub>V b = (\x\B. selfbutter x) *\<^sub>V b\ using 1 2 by (simp add: cblinfun.sum_left sum_single[where i=b]) qed lemma butterfly_sum_left: \butterfly (\i\M. \ i) \ = (\i\M. butterfly (\ i) \)\ apply (induction M rule:infinite_finite_induct) by (auto simp add: butterfly_add_left) lemma butterfly_sum_right: \butterfly \ (\i\M. \ i) = (\i\M. butterfly \ (\ i))\ apply (induction M rule:infinite_finite_induct) by (auto simp add: butterfly_add_right) 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 fail? (* Clears "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 riesz_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 \\<^sub>C x)\ - by transfer (rule riesz_frechet_representation_existence) - -lemma riesz_frechet_representation_cblinfun_unique: + by transfer (rule riesz_representation_existence) + +lemma riesz_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 \\<^sub>C x)\ assumes \\x. f *\<^sub>V x = (u \\<^sub>C x)\ shows \t = u\ - using assms by (rule riesz_frechet_representation_unique) - -theorem riesz_frechet_representation_cblinfun_norm: + using assms by (rule riesz_representation_unique) + +theorem riesz_representation_cblinfun_norm: includes notation_norm fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = (t \\<^sub>C x)\ shows \\f\ = \t\\ using assms proof transfer fix f::\'a \ complex\ and t assume \bounded_clinear f\ and \\x. f x = (t \\<^sub>C x)\ from \\x. f x = (t \\<^sub>C 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 \\<^sub>C x))\ using \\x. f x = (t \\<^sub>C x)\ by simp also have \norm (t \\<^sub>C 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 \\<^sub>C t)\ using \\x. f x = (t \\<^sub>C 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 definition the_riesz_rep :: \('a::chilbert_space \\<^sub>C\<^sub>L complex) \ 'a\ where \the_riesz_rep f = (SOME t. \x. f *\<^sub>V x = t \\<^sub>C x)\ lemma the_riesz_rep[simp]: \the_riesz_rep f \\<^sub>C x = f *\<^sub>V x\ unfolding the_riesz_rep_def apply (rule someI2_ex) - by (simp_all add: riesz_frechet_representation_cblinfun_existence) + by (simp_all add: riesz_representation_cblinfun_existence) lemma the_riesz_rep_unique: assumes \\x. f *\<^sub>V x = t \\<^sub>C x\ shows \t = the_riesz_rep f\ - using assms riesz_frechet_representation_cblinfun_unique the_riesz_rep by metis + using assms riesz_representation_cblinfun_unique the_riesz_rep by metis lemma the_riesz_rep_scaleC: \the_riesz_rep (c *\<^sub>C f) = cnj c *\<^sub>C the_riesz_rep f\ apply (rule the_riesz_rep_unique[symmetric]) by auto lemma the_riesz_rep_add: \the_riesz_rep (f + g) = the_riesz_rep f + the_riesz_rep g\ apply (rule the_riesz_rep_unique[symmetric]) by (auto simp: cinner_add_left cblinfun.add_left) lemma the_riesz_rep_norm[simp]: \norm (the_riesz_rep f) = norm f\ - apply (rule riesz_frechet_representation_cblinfun_norm[symmetric]) + apply (rule riesz_representation_cblinfun_norm[symmetric]) by simp lemma bounded_antilinear_the_riesz_rep[bounded_antilinear]: \bounded_antilinear the_riesz_rep\ by (metis (no_types, opaque_lifting) bounded_antilinear_intro dual_order.refl mult.commute mult_1 the_riesz_rep_add the_riesz_rep_norm the_riesz_rep_scaleC) lift_definition the_riesz_rep_sesqui :: \('a::complex_normed_vector \ 'b::chilbert_space \ complex) \ ('a \\<^sub>C\<^sub>L 'b)\ is \\p. if bounded_sesquilinear p then the_riesz_rep o CBlinfun o p else 0\ by (metis (mono_tags, lifting) Cblinfun_comp_bounded_sesquilinear bounded_antilinear_o_bounded_antilinear' bounded_antilinear_the_riesz_rep bounded_clinear_0 fun.map_comp) lemma the_riesz_rep_sesqui_apply: assumes \bounded_sesquilinear p\ shows \(the_riesz_rep_sesqui p *\<^sub>V x) \\<^sub>C y = p x y\ apply (transfer fixing: p x y) by (auto simp add: CBlinfun_inverse bounded_sesquilinear_apply_bounded_clinear assms) subsection \Bidual\ lift_definition bidual_embedding :: \'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 bidual_embedding_apply[simp]: \(bidual_embedding *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ by (transfer fixing: x f, simp) lemma bidual_embedding_isometric[simp]: \norm (bidual_embedding *\<^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_bidual_embedding[simp]: \norm (bidual_embedding :: '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 lemma isometry_bidual_embedding[simp]: \isometry bidual_embedding\ by (simp add: norm_preserving_isometry) lemma bidual_embedding_surj[simp]: \surj (bidual_embedding :: 'a::chilbert_space \\<^sub>C\<^sub>L _)\ proof - have \\y''. \f. (bidual_embedding *\<^sub>V y'') *\<^sub>V f = y *\<^sub>V f\ for y :: \('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex\ proof - define y' where \y' = CBlinfun (\z. cnj (y (cblinfun_cinner_right z)))\ have y'_apply: \y' z = cnj (y (cblinfun_cinner_right z))\ for z unfolding y'_def apply (subst CBlinfun_inverse) by (auto intro!: bounded_linear_intros) obtain y'' where \y' z = y'' \\<^sub>C z\ for z - using riesz_frechet_representation_cblinfun_existence by blast + using riesz_representation_cblinfun_existence by blast then have y'': \z \\<^sub>C y'' = cnj (y' z)\ for z by auto have \(bidual_embedding *\<^sub>V y'') *\<^sub>V f = y *\<^sub>V f\ for f :: \'a \\<^sub>C\<^sub>L complex\ proof - obtain f' where f': \f z = f' \\<^sub>C z\ for z - using riesz_frechet_representation_cblinfun_existence by blast + using riesz_representation_cblinfun_existence by blast then have f'2: \f = cblinfun_cinner_right f'\ using cblinfun_apply_inject by force show ?thesis by (auto simp add: f' f'2 y'' y'_apply) qed then show ?thesis by blast qed then show ?thesis by (metis cblinfun_eqI surj_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_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) lemma 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_bounded_dense: \cblinfun_extension_exists S f\ and cblinfun_extension_norm_bounded_dense: \B \ 0 \ norm (cblinfun_extension S f) \ B\ proof - define B' where \B' = (if B\0 then 1 else B)\ then have bounded': \\x. x \ S \ norm (f x) \ B' * norm x\ using bounded by (metis mult_1 mult_le_0_iff norm_ge_zero order_trans) have \B' > 0\ by (simp add: B'_def) 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 by atomize_elim (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 then have \cblinfun_extension S f *\<^sub>V \ = CBlinfun g *\<^sub>V \\ if \\ \ S\ for \ by (simp add: cblinfun_extension_apply that) then have ext_is_g: \(*\<^sub>V) (cblinfun_extension S f) = g\ apply (rule_tac ext) apply (rule on_closure_eqI[where S=S]) using \closure S = UNIV\ \bounded_clinear g\ by (auto simp add: continuous_at_imp_continuous_on clinear_continuous_within) show \norm (cblinfun_extension S f) \ B\ if \B \ 0\ proof (cases \B > 0\) case True then have \B = B'\ unfolding B'_def by auto moreover have *: \norm (cblinfun_extension S f) \ B'\ by (metis ext_is_g \0 < B'\ g_bounded norm_cblinfun_bound order_le_less) ultimately show ?thesis by simp next case False with bounded have \f x = 0\ if \x \ S\ for x by (smt (verit) mult_nonpos_nonneg norm_ge_zero norm_le_zero_iff that) then have \g x = (\_. 0) x\ if \x \ S\ for x using that by simp then have \g x = 0\ for x apply (rule on_closure_eqI[where S=S]) using \closure S = UNIV\ \bounded_clinear g\ by (auto simp add: continuous_at_imp_continuous_on clinear_continuous_within) with ext_is_g have \cblinfun_extension S f = 0\ by (simp add: cblinfun_eqI) then show ?thesis using that by simp qed qed lemma cblinfun_extension_cong: assumes \cspan A = cspan B\ assumes \B \ A\ assumes fg: \\x. x\B \ f x = g x\ assumes \cblinfun_extension_exists A f\ shows \cblinfun_extension A f = cblinfun_extension B g\ proof - from \cblinfun_extension_exists A f\ fg \B \ A\ have \cblinfun_extension_exists B g\ by (metis assms(2) cblinfun_extension_exists_def subset_eq) have \(\x\A. C *\<^sub>V x = f x) \ (\x\B. C *\<^sub>V x = f x)\ for C by (smt (verit, ccfv_SIG) assms(1) assms(2) assms(4) cblinfun_eq_on_span cblinfun_extension_exists_def complex_vector.span_eq subset_iff) also from fg have \\ C \ (\x\B. C *\<^sub>V x = g x)\ for C by auto finally show \cblinfun_extension A f = cblinfun_extension B g\ unfolding cblinfun_extension_def by auto qed lemma fixes f :: \'a::complex_inner \ 'b::chilbert_space\ and S assumes \is_ortho_set S\ and \closure (cspan S) = UNIV\ assumes ortho_f: \\x y. x\S \ y\S \ x\y \ is_orthogonal (f x) (f y)\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows cblinfun_extension_exists_ortho: \cblinfun_extension_exists S f\ and cblinfun_extension_exists_ortho_norm: \B \ 0 \ norm (cblinfun_extension S f) \ B\ proof - define g where \g = cconstruct S f\ have \cindependent S\ using assms(1) is_ortho_set_cindependent by blast have g_f: \g x = f x\ if \x\S\ for x unfolding g_def using \cindependent S\ that by (rule complex_vector.construct_basis) have [simp]: \clinear g\ unfolding g_def using \cindependent S\ by (rule complex_vector.linear_construct) then have g_add: \g (x + y) = g x + g y\ if \x \ cspan S\ and \y \ cspan S\ for x y using clinear_iff by blast from \clinear g\ have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ if \x \ cspan S\ for x c by (simp add: complex_vector.linear_scale) moreover have g_bounded: \norm (g x) \ abs B * norm x\ if \x \ cspan S\ for x proof - from that obtain t r where x_sum: \x = (\a\t. r a *\<^sub>C a)\ and \finite t\ and \t \ S\ unfolding complex_vector.span_explicit by auto have \(norm (g x))\<^sup>2 = (norm (\a\t. r a *\<^sub>C g a))\<^sup>2\ by (simp add: x_sum complex_vector.linear_sum clinear.scaleC) also have \\ = (norm (\a\t. r a *\<^sub>C f a))\<^sup>2\ by (smt (verit) \t \ S\ g_f in_mono sum.cong) also have \\ = (\a\t. (norm (r a *\<^sub>C f a))\<^sup>2)\ using _ \finite t\ apply (rule pythagorean_theorem_sum) using \t \ S\ ortho_f in_mono by fastforce also have \\ = (\a\t. (cmod (r a) * norm (f a))\<^sup>2)\ by simp also have \\ \ (\a\t. (cmod (r a) * B * norm a)\<^sup>2)\ apply (rule sum_mono) by (metis \t \ S\ assms(4) in_mono mult_left_mono mult_nonneg_nonneg norm_ge_zero power_mono vector_space_over_itself.scale_scale) also have \\ = B\<^sup>2 * (\a\t. (norm (r a *\<^sub>C a))\<^sup>2)\ by (simp add: sum_distrib_left mult.commute vector_space_over_itself.scale_left_commute flip: power_mult_distrib) also have \\ = B\<^sup>2 * (norm (\a\t. (r a *\<^sub>C a)))\<^sup>2\ apply (subst pythagorean_theorem_sum) using \finite t\ by auto (meson \t \ S\ assms(1) is_ortho_set_def subsetD) also have \\ = (abs B * norm x)\<^sup>2\ by (simp add: power_mult_distrib x_sum) finally show \norm (g x) \ abs B * norm x\ by auto qed from g_add g_scale g_bounded have extg_exists: \cblinfun_extension_exists (cspan S) g\ apply (rule_tac cblinfun_extension_exists_bounded_dense[where B=\abs B\]) using \closure (cspan S) = UNIV\ by auto then show \cblinfun_extension_exists S f\ by (metis (mono_tags, opaque_lifting) g_f cblinfun_extension_apply cblinfun_extension_existsI complex_vector.span_base) have norm_extg: \norm (cblinfun_extension (cspan S) g) \ B\ if \B \ 0\ apply (rule cblinfun_extension_norm_bounded_dense) using g_add g_scale g_bounded \closure (cspan S) = UNIV\ that by auto have extg_extf: \cblinfun_extension (cspan S) g = cblinfun_extension S f\ apply (rule cblinfun_extension_cong) by (auto simp add: complex_vector.span_base g_f extg_exists) from norm_extg extg_extf show \norm (cblinfun_extension S f) \ B\ if \B \ 0\ using that by simp qed lemma cblinfun_extension_exists_proj: fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes ex_P: \\P :: 'a \\<^sub>C\<^sub>L 'a. is_Proj P \ range P = closure S\ 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\ \ \We cannot give a statement about the norm. While there is an extension with norm \<^term>\B\, there is no guarantee that \<^term>\cblinfun_extension S f\ returns that specific extension since the extension is only determined on \<^term>\ccspan S\.\ proof (cases \B \ 0\) case True note True[simp] obtain P :: \'a \\<^sub>C\<^sub>L 'a\ where P_proj: \is_Proj P\ and P_im: \range P = closure S\ using ex_P by blast define f' S' where \f' \ = f (P \)\ and \S' = S + space_as_set (kernel P)\ for \ have PS': \P *\<^sub>V x \ S\ if \x \ S'\ for x proof - obtain x1 x2 where x12: \x = x1 + x2\ and x1: \x1 \ S\ and x2: \x2 \ space_as_set (kernel P)\ using that S'_def \x \ S'\ set_plus_elim by blast have \P *\<^sub>V x = P *\<^sub>V x1\ using x2 by (simp add: x12 cblinfun.add_right kernel_memberD) also have \\ = x1\ by (metis (no_types, opaque_lifting) P_im P_proj cblinfun_apply_cblinfun_compose closure_insert image_iff insertI1 insert_absorb is_Proj_idempotent x1) also have \\ \ S\ by (fact x1) finally show ?thesis by - qed have SS': \S \ S'\ by (metis S'_def ordered_field_class.sign_simps(2) set_zero_plus2 zero_space_as_set) have \csubspace S'\ by (simp add: S'_def assms(1) csubspace_set_plus) moreover have \closure S' = UNIV\ proof auto fix \ have \\ = P \ + (id - P) \\ by simp also have \\ \ closure S + space_as_set (kernel P)\ by (smt (verit) P_im P_proj calculation cblinfun.real.add_right eq_add_iff is_Proj_idempotent kernel_memberI rangeI set_plus_intro simp_a_oCL_b') also have \\ \ closure (closure S + space_as_set (kernel P))\ using closure_subset by blast also have \\ = closure (S + space_as_set (kernel P))\ using closed_sum_closure_left closed_sum_def by blast also have \\ = closure S'\ using S'_def by fastforce finally show \\ \ closure S'\ by - qed moreover have \f' (x + y) = f' x + f' y\ if \x \ S'\ and \y \ S'\ for x y using that by (auto simp: f'_def cblinfun.add_right f_add PS') moreover have \f' (c *\<^sub>C x) = c *\<^sub>C f' x\ if \x \ S'\ for c x using that by (auto simp: f'_def cblinfun.scaleC_right f_scale PS') moreover have \norm (f' x) \ (B * norm P) * norm x\ if \x \ S'\ for x proof - have \norm (f' x) \ B* norm (P x)\ by (auto simp: f'_def PS' bounded that) also have \\ \ B * norm P * norm x\ by (simp add: mult.assoc mult_left_mono norm_cblinfun) finally show ?thesis by auto qed ultimately have F_ex: \cblinfun_extension_exists S' f'\ by (rule cblinfun_extension_exists_bounded_dense) define F where \F = cblinfun_extension S' f'\ have \F \ = f \\ if \\ \ S\ for \ proof - from F_ex that SS' have \F \ = f' \\ by (auto simp add: F_def cblinfun_extension_apply) also have \\ = f (P *\<^sub>V \)\ by (simp add: f'_def) also have \\ = f \\ using P_proj P_im apply (transfer fixing: \ S f) by (metis closure_subset in_mono is_projection_on_fixes_image is_projection_on_image that) finally show ?thesis by - qed then show \cblinfun_extension_exists S f\ using cblinfun_extension_exists_def by blast next case False then have \S \ {0}\ using bounded by auto (meson norm_ge_zero norm_le_zero_iff order_trans zero_le_mult_iff) then show \cblinfun_extension_exists S f\ apply (rule_tac cblinfun_extension_existsI[where B=0]) apply auto using bounded by fastforce qed lemma cblinfun_extension_exists_hilbert: fixes f :: \'a::chilbert_space \ 'b::cbanach\ assumes \csubspace S\ 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\ \ \We cannot give a statement about the norm. While there is an extension with norm \<^term>\B\, there is no guarantee that \<^term>\cblinfun_extension S f\ returns that specific extension since the extension is only determined on \<^term>\ccspan S\.\ proof - have \\P. is_Proj P \ range ((*\<^sub>V) P) = closure S\ apply (rule exI[of _ \Proj (ccspan S)\]) apply (rule conjI) by simp (metis Proj_is_Proj Proj_range Proj_range_closed assms(1) cblinfun_image.rep_eq ccspan.rep_eq closure_closed complex_vector.span_eq_iff space_as_set_top) from assms(1) this assms(2-4) show ?thesis by (rule cblinfun_extension_exists_proj) qed lemma cblinfun_extension_exists_restrict: assumes \B \ A\ assumes \\x. x\B \ f x = g x\ assumes \cblinfun_extension_exists A f\ shows \cblinfun_extension_exists B g\ by (metis assms cblinfun_extension_exists_def subset_eq) subsection \Bijections between different ONBs\ text \Some of the theorems here logically belong into \<^theory>\Complex_Bounded_Operators.Complex_Inner_Product\ but the proof uses some concepts from the present theory.\ lemma all_ortho_bases_same_card: \ \Follows \<^cite>\conway2013course\, Proposition 4.14\ fixes E F :: \'a::chilbert_space set\ assumes \is_ortho_set E\ \is_ortho_set F\ \ccspan E = \\ \ccspan F = \\ shows \\f. bij_betw f E F\ proof - have \|F| \o |E|\ if \infinite E\ and \is_ortho_set E\ \is_ortho_set F\ \ccspan E = top\ \ccspan F = top\ for E F :: \'a set\ proof - define F' where \F' e = {f\F. \ is_orthogonal f e}\ for e have \\e\E. cinner f e \ 0\ if \f \ F\ for f proof (rule ccontr, simp) assume \\e\E. is_orthogonal f e\ then have \f \ orthogonal_complement E\ by (simp add: orthogonal_complementI) also have \\ = orthogonal_complement (closure (cspan E))\ using orthogonal_complement_of_closure orthogonal_complement_of_cspan by blast also have \\ = space_as_set (- ccspan E)\ by transfer simp also have \\ = space_as_set (- top)\ by (simp add: \ccspan E = top\) also have \\ = {0}\ by (auto simp add: top_ccsubspace.rep_eq uminus_ccsubspace.rep_eq) finally have \f = 0\ by simp with \f \ F\ \is_ortho_set F\ show False by (simp add: is_onb_def is_ortho_set_def) qed then have F'_union: \F = (\e\E. F' e)\ unfolding F'_def by auto have \countable (F' e)\ for e proof - have \(\f\M. (cmod (cinner (sgn f) e))\<^sup>2) \ (norm e)\<^sup>2\ if \finite M\ and \M \ F\ for M proof - have [simp]: \is_ortho_set M\ using \is_ortho_set F\ is_ortho_set_antimono that(2) by blast have [simp]: \norm (sgn f) = 1\ if \f \ M\ for f by (metis \is_ortho_set M\ is_ortho_set_def norm_sgn that) have \(\f\M. (cmod (cinner (sgn f) e))\<^sup>2) = (\f\M. (norm ((cinner (sgn f) e) *\<^sub>C sgn f))\<^sup>2)\ apply (rule sum.cong[OF refl]) by simp also have \\ = (norm (\f\M. ((cinner (sgn f) e) *\<^sub>C sgn f)))\<^sup>2\ apply (rule pythagorean_theorem_sum[symmetric]) using that by auto (meson \is_ortho_set M\ is_ortho_set_def) also have \\ = (norm (\f\M. proj f e))\<^sup>2\ by (metis butterfly_apply butterfly_sgn_eq_proj) also have \\ = (norm (Proj (ccspan M) e))\<^sup>2\ apply (rule arg_cong[where f=\\x. (norm x)\<^sup>2\]) using \finite M\ \is_ortho_set M\ apply induction by simp (smt (verit, ccfv_threshold) Proj_orthog_ccspan_insert insertCI is_ortho_set_def plus_cblinfun.rep_eq sum.insert) also have \\ \ (norm (Proj (ccspan M)) * norm e)\<^sup>2\ by (auto simp: norm_cblinfun intro!: power_mono) also have \\ \ (norm e)\<^sup>2\ apply (rule power_mono) apply (metis norm_Proj_leq1 mult_left_le_one_le norm_ge_zero) by simp finally show ?thesis by - qed then have \(\f. (cmod (cinner (sgn f) e))\<^sup>2) abs_summable_on F\ apply (intro nonneg_bdd_above_summable_on bdd_aboveI) by auto then have \countable {f \ F. (cmod (sgn f \\<^sub>C e))\<^sup>2 \ 0}\ by (rule abs_summable_countable) then have \countable {f \ F. \ is_orthogonal (sgn f) e}\ by force then have \countable {f \ F. \ is_orthogonal f e}\ by force then show ?thesis unfolding F'_def by simp qed then have F'_leq: \|F' e| \o natLeq\ for e using countable_leq_natLeq by auto from F'_union have \|F| \o |Sigma E F'|\ (is \_ \o \\) using card_of_UNION_Sigma by blast also have \\ \o |E \ (UNIV::nat set)|\ (is \_ \o \\) apply (rule card_of_Sigma_mono1) using F'_leq using card_of_nat ordIso_symmetric ordLeq_ordIso_trans by blast also have \\ =o |E| *c natLeq\ (is \ordIso2 _ \\) by (metis Field_card_of Field_natLeq card_of_ordIso_subst cprod_def) also have \\ =o |E|\ apply (rule card_prod_omega) using that by (simp add: cinfinite_def) finally show \|F| \o |E|\ by - qed then have infinite: \|E| =o |F|\ if \infinite E\ and \infinite F\ by (simp add: assms ordIso_iff_ordLeq that(1) that(2)) have \|E| =o |F|\ if \finite E\ and \is_ortho_set E\ \is_ortho_set F\ \ccspan E = top\ \ccspan F = top\ for E F :: \'a set\ proof - have \card E = card F\ using that by (metis bij_betw_same_card ccspan.rep_eq closure_finite_cspan complex_vector.bij_if_span_eq_span_bases complex_vector.independent_span_bound is_ortho_set_cindependent top_ccsubspace.rep_eq top_greatest) with \finite E\ have \finite F\ by (metis ccspan.rep_eq closure_finite_cspan complex_vector.independent_span_bound is_ortho_set_cindependent that(3) that(4) top_ccsubspace.rep_eq top_greatest) with \card E = card F\ that show ?thesis apply (rule_tac finite_card_of_iff_card[THEN iffD2]) by auto qed with infinite assms have \|E| =o |F|\ by (meson ordIso_symmetric) then show ?thesis by (simp add: card_of_ordIso) qed lemma all_onbs_same_card: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows \\f. bij_betw f E F\ apply (rule all_ortho_bases_same_card) using assms by (auto simp: is_onb_def) definition bij_between_bases where \bij_between_bases E F = (SOME f. bij_betw f E F)\ for E F :: \'a::chilbert_space set\ lemma bij_between_bases_bij: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows \bij_betw (bij_between_bases E F) E F\ using all_onbs_same_card by (metis assms(1) assms(2) bij_between_bases_def someI) definition unitary_between where \unitary_between E F = cblinfun_extension E (bij_between_bases E F)\ lemma unitary_between_apply: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ \e \ E\ shows \unitary_between E F *\<^sub>V e = bij_between_bases E F e\ proof - from \is_onb E\ \is_onb F\ have EF: \bij_between_bases E F e \ F\ if \e \ E\ for e by (meson bij_betwE bij_between_bases_bij that) have ortho: \is_orthogonal (bij_between_bases E F x) (bij_between_bases E F y)\ if \x \ y\ and \x \ E\ and \y \ E\ for x y by (smt (verit, del_insts) assms(1) assms(2) bij_betw_iff_bijections bij_between_bases_bij is_onb_def is_ortho_set_def that(1) that(2) that(3)) have spanE: \closure (cspan E) = UNIV\ by (metis assms(1) ccspan.rep_eq is_onb_def top_ccsubspace.rep_eq) show ?thesis unfolding unitary_between_def apply (rule cblinfun_extension_apply) apply (rule cblinfun_extension_exists_ortho[where B=1]) using assms EF ortho spanE by (auto simp: is_onb_def) qed lemma unitary_between_unitary: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows \unitary (unitary_between E F)\ proof - have \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V c) = b \\<^sub>C c\ if \b \ E\ and \c \ E\ for b c proof (cases \b = c\) case True from \is_onb E\ that have 1: \b \\<^sub>C b = 1\ using cnorm_eq_1 is_onb_def by blast from that have \unitary_between E F *\<^sub>V b \ F\ by (metis assms(1) assms(2) bij_betw_apply bij_between_bases_bij unitary_between_apply) with \is_onb F\ have 2: \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V b) = 1\ by (simp add: cnorm_eq_1 is_onb_def) from 1 2 True show ?thesis by simp next case False from \is_onb E\ that have 1: \b \\<^sub>C c = 0\ by (simp add: False is_onb_def is_ortho_set_def) from that have inF: \unitary_between E F *\<^sub>V b \ F\ \unitary_between E F *\<^sub>V c \ F\ by (metis assms(1) assms(2) bij_betw_apply bij_between_bases_bij unitary_between_apply)+ have neq: \unitary_between E F *\<^sub>V b \ unitary_between E F *\<^sub>V c\ by (metis (no_types, lifting) False assms(1) assms(2) bij_betw_iff_bijections bij_between_bases_bij that(1) that(2) unitary_between_apply) from inF neq \is_onb F\ have 2: \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V c) = 0\ by (simp add: is_onb_def is_ortho_set_def) from 1 2 show ?thesis by simp qed then have iso: \isometry (unitary_between E F)\ apply (rule_tac orthogonal_on_basis_is_isometry[where B=E]) using assms(1) is_onb_def by auto have \unitary_between E F *\<^sub>S top = unitary_between E F *\<^sub>S ccspan E\ by (metis assms(1) is_onb_def) also have \\ \ ccspan (unitary_between E F ` E)\ (is \_ \ \\) by (simp add: cblinfun_image_ccspan) also have \\ = ccspan (bij_between_bases E F ` E)\ by (metis assms(1) assms(2) image_cong unitary_between_apply) also have \\ = ccspan F\ by (metis assms(1) assms(2) bij_betw_imp_surj_on bij_between_bases_bij) also have \\ = top\ using assms(2) is_onb_def by blast finally have surj: \unitary_between E F *\<^sub>S top = top\ by (simp add: top.extremum_unique) from iso surj show ?thesis by (rule surj_isometry_is_unitary) qed subsection \Notation\ bundle cblinfun_notation begin notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 67) notation cblinfun_apply (infixr "*\<^sub>V" 70) notation cblinfun_image (infixr "*\<^sub>S" 70) notation adj ("_*" [99] 100) type_notation cblinfun ("(_ \\<^sub>C\<^sub>L /_)" [22, 21] 21) end bundle no_cblinfun_notation begin no_notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 67) no_notation cblinfun_apply (infixr "*\<^sub>V" 70) no_notation cblinfun_image (infixr "*\<^sub>S" 70) no_notation adj ("_*" [99] 100) no_type_notation cblinfun ("(_ \\<^sub>C\<^sub>L /_)" [22, 21] 21) end unbundle no_cblinfun_notation unbundle no_lattice_syntax end diff --git a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy --- a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy +++ b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy @@ -1,2588 +1,2588 @@ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) section \\Complex_Inner_Product\ -- Complex Inner Product Spaces\ theory Complex_Inner_Product imports Complex_Inner_Product0 begin subsection \Complex inner product spaces\ unbundle cinner_syntax lemma cinner_real: "cinner x x \ \" by (simp add: cdot_square_norm) lemmas cinner_commute' [simp] = cinner_commute[symmetric] lemma (in complex_inner) cinner_eq_flip: \(cinner x y = cinner z w) \ (cinner y x = cinner w z)\ by (metis cinner_commute) lemma Im_cinner_x_x[simp]: "Im (x \\<^sub>C x) = 0" using comp_Im_same[OF cinner_ge_zero] by simp lemma of_complex_inner_1' [simp]: "cinner (1 :: 'a :: {complex_inner, complex_normed_algebra_1}) (of_complex x) = x" by (metis cinner_commute complex_cnj_cnj of_complex_inner_1) class chilbert_space = complex_inner + complete_space begin subclass cbanach by standard end instantiation complex :: "chilbert_space" begin instance .. end subsection \Misc facts\ lemma cinner_scaleR_left [simp]: "cinner (scaleR r x) y = of_real r * (cinner x y)" by (simp add: scaleR_scaleC) lemma cinner_scaleR_right [simp]: "cinner x (scaleR r y) = of_real r * (cinner x y)" by (simp add: scaleR_scaleC) text \This is a useful rule for establishing the equality of vectors\ lemma cinner_extensionality: assumes \\\. \ \\<^sub>C \ = \ \\<^sub>C \\ shows \\ = \\ by (metis assms cinner_eq_zero_iff cinner_simps(3) right_minus_eq) lemma polar_identity: includes notation_norm shows \\x + y\^2 = \x\^2 + \y\^2 + 2 * Re (x \\<^sub>C y)\ \ \Shown in the proof of Corollary 1.5 in \<^cite>\conway2013course\\ proof - have \(x \\<^sub>C y) + (y \\<^sub>C x) = (x \\<^sub>C y) + cnj (x \\<^sub>C y)\ by simp hence \(x \\<^sub>C y) + (y \\<^sub>C x) = 2 * Re (x \\<^sub>C y) \ using complex_add_cnj by presburger have \\x + y\^2 = (x+y) \\<^sub>C (x+y)\ by (simp add: cdot_square_norm) hence \\x + y\^2 = (x \\<^sub>C x) + (x \\<^sub>C y) + (y \\<^sub>C x) + (y \\<^sub>C y)\ by (simp add: cinner_add_left cinner_add_right) thus ?thesis using \(x \\<^sub>C y) + (y \\<^sub>C x) = 2 * Re (x \\<^sub>C y)\ by (smt (verit, ccfv_SIG) Re_complex_of_real plus_complex.simps(1) power2_norm_eq_cinner') qed lemma polar_identity_minus: includes notation_norm shows \\x - y\^2 = \x\^2 + \y\^2 - 2 * Re (x \\<^sub>C y)\ proof- have \\x + (-y)\^2 = \x\^2 + \-y\^2 + 2 * Re (x \\<^sub>C -y)\ using polar_identity by blast hence \\x - y\^2 = \x\^2 + \y\^2 - 2*Re (x \\<^sub>C y)\ by simp thus ?thesis by blast qed proposition parallelogram_law: includes notation_norm fixes x y :: "'a::complex_inner" shows \\x+y\^2 + \x-y\^2 = 2*( \x\^2 + \y\^2 )\ \ \Shown in the proof of Theorem 2.3 in \<^cite>\conway2013course\\ by (simp add: polar_identity_minus polar_identity) theorem pythagorean_theorem: includes notation_norm shows \(x \\<^sub>C y) = 0 \ \ x + y \^2 = \ x \^2 + \ y \^2\ \ \Shown in the proof of Theorem 2.2 in \<^cite>\conway2013course\\ by (simp add: polar_identity) lemma pythagorean_theorem_sum: assumes q1: "\a a'. a \ t \ a' \ t \ a \ a' \ f a \\<^sub>C f a' = 0" and q2: "finite t" shows "(norm (\a\t. f a))^2 = (\a\t.(norm (f a))^2)" proof (insert q1, use q2 in induction) case empty show ?case by auto next case (insert x F) have r1: "f x \\<^sub>C f a = 0" if "a \ F" for a using that insert.hyps(2) insert.prems by auto have "sum f F = (\a\F. f a)" by simp hence s4: "f x \\<^sub>C sum f F = f x \\<^sub>C (\a\F. f a)" by simp also have s3: "\ = (\a\F. f x \\<^sub>C f a)" using cinner_sum_right by auto also have s2: "\ = (\a\F. 0)" using r1 by simp also have s1: "\ = 0" by simp finally have xF_ortho: "f x \\<^sub>C sum f F = 0" using s2 s3 by auto have "(norm (sum f (insert x F)))\<^sup>2 = (norm (f x + sum f F))\<^sup>2" by (simp add: insert.hyps(1) insert.hyps(2)) also have "\ = (norm (f x))\<^sup>2 + (norm (sum f F))\<^sup>2" using xF_ortho by (rule pythagorean_theorem) also have "\ = (norm (f x))\<^sup>2 + (\a\F.(norm (f a))^2)" apply (subst insert.IH) using insert.prems by auto also have "\ = (\a\insert x F.(norm (f a))^2)" by (simp add: insert.hyps(1) insert.hyps(2)) finally show ?case by simp qed lemma Cauchy_cinner_Cauchy: fixes x y :: \nat \ 'a::complex_inner\ assumes a1: \Cauchy x\ and a2: \Cauchy y\ shows \Cauchy (\ n. x n \\<^sub>C y n)\ proof- have \bounded (range x)\ using a1 by (simp add: Elementary_Metric_Spaces.cauchy_imp_bounded) hence b1: \\M. \n. norm (x n) < M\ by (meson bounded_pos_less rangeI) have \bounded (range y)\ using a2 by (simp add: Elementary_Metric_Spaces.cauchy_imp_bounded) hence b2: \\ M. \ n. norm (y n) < M\ by (meson bounded_pos_less rangeI) have \\M. \n. norm (x n) < M \ norm (y n) < M\ using b1 b2 by (metis dual_order.strict_trans linorder_neqE_linordered_idom) then obtain M where M1: \\n. norm (x n) < M\ and M2: \\n. norm (y n) < M\ by blast have M3: \M > 0\ by (smt M2 norm_not_less_zero) have \\N. \n \ N. \m \ N. norm ( (\ i. x i \\<^sub>C y i) n - (\ i. x i \\<^sub>C y i) m ) < e\ if "e > 0" for e proof- have \e / (2*M) > 0\ using M3 by (simp add: that) hence \\N. \n\N. \m\N. norm (x n - x m) < e / (2*M)\ using a1 by (simp add: Cauchy_iff) then obtain N1 where N1_def: \\n m. n\N1 \ m\N1 \ norm (x n - x m) < e / (2*M)\ by blast have x1: \\N. \ n\N. \ m\N. norm (y n - y m) < e / (2*M)\ using a2 \e / (2*M) > 0\ by (simp add: Cauchy_iff) obtain N2 where N2_def: \\n m. n\N2 \ m\N2 \ norm (y n - y m) < e / (2*M)\ using x1 by blast define N where N_def: \N = N1 + N2\ hence \N \ N1\ by auto have \N \ N2\ using N_def by auto have \norm (x n \\<^sub>C y n - x m \\<^sub>C y m) < e\ if \n \ N\ and \m \ N\ for n m proof - have \x n \\<^sub>C y n - x m \\<^sub>C y m = (x n \\<^sub>C y n - x m \\<^sub>C y n) + (x m \\<^sub>C y n - x m \\<^sub>C y m)\ by simp hence y1: \norm (x n \\<^sub>C y n - x m \\<^sub>C y m) \ norm (x n \\<^sub>C y n - x m \\<^sub>C y n) + norm (x m \\<^sub>C y n - x m \\<^sub>C y m)\ by (metis norm_triangle_ineq) have \x n \\<^sub>C y n - x m \\<^sub>C y n = (x n - x m) \\<^sub>C y n\ by (simp add: cinner_diff_left) hence \norm (x n \\<^sub>C y n - x m \\<^sub>C y n) = norm ((x n - x m) \\<^sub>C y n)\ by simp moreover have \norm ((x n - x m) \\<^sub>C y n) \ norm (x n - x m) * norm (y n)\ using complex_inner_class.Cauchy_Schwarz_ineq2 by blast moreover have \norm (y n) < M\ by (simp add: M2) moreover have \norm (x n - x m) < e/(2*M)\ using \N \ m\ \N \ n\ \N1 \ N\ N1_def by auto ultimately have \norm ((x n \\<^sub>C y n) - (x m \\<^sub>C y n)) < (e/(2*M)) * M\ by (smt linordered_semiring_strict_class.mult_strict_mono norm_ge_zero) moreover have \ (e/(2*M)) * M = e/2\ using \M > 0\ by simp ultimately have \norm ((x n \\<^sub>C y n) - (x m \\<^sub>C y n)) < e/2\ by simp hence y2: \norm (x n \\<^sub>C y n - x m \\<^sub>C y n) < e/2\ by blast have \x m \\<^sub>C y n - x m \\<^sub>C y m = x m \\<^sub>C (y n - y m)\ by (simp add: cinner_diff_right) hence \norm ((x m \\<^sub>C y n) - (x m \\<^sub>C y m)) = norm (x m \\<^sub>C (y n - y m))\ by simp moreover have \norm (x m \\<^sub>C (y n - y m)) \ norm (x m) * norm (y n - y m)\ by (meson complex_inner_class.Cauchy_Schwarz_ineq2) moreover have \norm (x m) < M\ by (simp add: M1) moreover have \norm (y n - y m) < e/(2*M)\ using \N \ m\ \N \ n\ \N2 \ N\ N2_def by auto ultimately have \norm ((x m \\<^sub>C y n) - (x m \\<^sub>C y m)) < M * (e/(2*M))\ by (smt linordered_semiring_strict_class.mult_strict_mono norm_ge_zero) moreover have \M * (e/(2*M)) = e/2\ using \M > 0\ by simp ultimately have \norm ((x m \\<^sub>C y n) - (x m \\<^sub>C y m)) < e/2\ by simp hence y3: \norm ((x m \\<^sub>C y n) - (x m \\<^sub>C y m)) < e/2\ by blast show \norm ( (x n \\<^sub>C y n) - (x m \\<^sub>C y m) ) < e\ using y1 y2 y3 by simp qed thus ?thesis by blast qed thus ?thesis by (simp add: CauchyI) qed lemma cinner_sup_norm: \norm \ = (SUP \. cmod (cinner \ \) / norm \)\ proof (rule sym, rule cSup_eq_maximum) have \norm \ = cmod (cinner \ \) / norm \\ by (metis norm_eq_sqrt_cinner norm_ge_zero real_div_sqrt) then show \norm \ \ range (\\. cmod (cinner \ \) / norm \)\ by blast next fix n assume \n \ range (\\. cmod (cinner \ \) / norm \)\ then obtain \ where n\: \n = cmod (cinner \ \) / norm \\ by auto show \n \ norm \\ unfolding n\ by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2 divide_le_eq ordered_field_class.sign_simps(33)) qed lemma cinner_sup_onorm: fixes A :: \'a::{real_normed_vector,not_singleton} \ 'b::complex_inner\ assumes \bounded_linear A\ shows \onorm A = (SUP (\,\). cmod (cinner \ (A \)) / (norm \ * norm \))\ proof (unfold onorm_def, rule cSup_eq_cSup) show \bdd_above (range (\x. norm (A x) / norm x))\ by (meson assms bdd_aboveI2 le_onorm) next fix a assume \a \ range (\\. norm (A \) / norm \)\ then obtain \ where \a = norm (A \) / norm \\ by auto then have \a \ cmod (cinner (A \) (A \)) / (norm (A \) * norm \)\ apply auto by (smt (verit) divide_divide_eq_left norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) then show \\b\range (\(\, \). cmod (cinner \ (A \)) / (norm \ * norm \)). a \ b\ by force next fix b assume \b \ range (\(\, \). cmod (cinner \ (A \)) / (norm \ * norm \))\ then obtain \ \ where b: \b = cmod (cinner \ (A \)) / (norm \ * norm \)\ by auto then have \b \ norm (A \) / norm \\ apply auto by (smt (verit, ccfv_threshold) complex_inner_class.Cauchy_Schwarz_ineq2 division_ring_divide_zero linordered_field_class.divide_right_mono mult_cancel_left1 nonzero_mult_divide_mult_cancel_left2 norm_imp_pos_and_ge ordered_field_class.sign_simps(33) zero_le_divide_iff) then show \\a\range (\x. norm (A x) / norm x). b \ a\ by auto qed lemma sum_cinner: fixes f :: "'a \ 'b::complex_inner" shows "cinner (sum f A) (sum g B) = (\i\A. \j\B. cinner (f i) (g j))" by (simp add: cinner_sum_right cinner_sum_left) (rule sum.swap) lemma Cauchy_cinner_product_summable': fixes a b :: "nat \ 'a::complex_inner" shows \(\(x, y). cinner (a x) (b y)) summable_on UNIV \ (\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ proof - have img: \(\(k::nat, i). (i, k - i)) ` {(k, i). i \ k} = UNIV\ apply (auto simp: image_def) by (metis add.commute add_diff_cancel_right' diff_le_self) have inj: \inj_on (\(k::nat, i). (i, k - i)) {(k, i). i \ k}\ by (smt (verit, del_insts) Pair_inject case_prodE case_prod_conv eq_diff_iff inj_onI mem_Collect_eq) have \(\(x, y). cinner (a x) (b y)) summable_on UNIV \ (\(k, l). cinner (a k) (b l)) summable_on (\(k, i). (i, k - i)) ` {(k, i). i \ k}\ by (simp only: img) also have \\ \ ((\(k, l). cinner (a k) (b l)) \ (\(k, i). (i, k - i))) summable_on {(k, i). i \ k}\ using inj by (rule summable_on_reindex) also have \\ \ (\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ by (simp add: o_def case_prod_unfold) finally show ?thesis by - qed instantiation prod :: (complex_inner, complex_inner) complex_inner begin definition cinner_prod_def: "cinner x y = cinner (fst x) (fst y) + cinner (snd x) (snd y)" instance proof fix r :: complex fix x y z :: "'a::complex_inner \ 'b::complex_inner" show "cinner x y = cnj (cinner y x)" unfolding cinner_prod_def by simp show "cinner (x + y) z = cinner x z + cinner y z" unfolding cinner_prod_def by (simp add: cinner_add_left) show "cinner (scaleC r x) y = cnj r * cinner x y" unfolding cinner_prod_def by (simp add: distrib_left) show "0 \ cinner x x" unfolding cinner_prod_def by (intro add_nonneg_nonneg cinner_ge_zero) show "cinner x x = 0 \ x = 0" unfolding cinner_prod_def prod_eq_iff by (metis antisym cinner_eq_zero_iff cinner_ge_zero fst_zero le_add_same_cancel2 snd_zero verit_sum_simplify) show "norm x = sqrt (cmod (cinner x x))" unfolding norm_prod_def cinner_prod_def by (metis (no_types, lifting) Re_complex_of_real add_nonneg_nonneg cinner_ge_zero complex_of_real_cmod plus_complex.simps(1) power2_norm_eq_cinner') qed end lemma sgn_cinner[simp]: \sgn \ \\<^sub>C \ = norm \\ apply (cases \\ = 0\) apply (auto simp: sgn_div_norm) by (smt (verit, ccfv_SIG) cinner_scaleR_left cinner_scaleR_right cnorm_eq cnorm_eq_1 complex_of_real_cmod complex_of_real_nn_iff left_inverse mult.right_neutral mult_scaleR_right norm_eq_zero norm_not_less_zero norm_one of_real_def of_real_eq_iff) instance prod :: (chilbert_space, chilbert_space) chilbert_space.. subsection \Orthogonality\ definition "orthogonal_complement S = {x| x. \y\S. cinner x y = 0}" lemma orthogonal_complement_orthoI: \x \ orthogonal_complement M \ y \ M \ x \\<^sub>C y = 0\ unfolding orthogonal_complement_def by auto lemma orthogonal_complement_orthoI': \x \ M \ y \ orthogonal_complement M \ x \\<^sub>C y = 0\ by (metis cinner_commute' complex_cnj_zero orthogonal_complement_orthoI) lemma orthogonal_complementI: \(\x. x \ M \ y \\<^sub>C x = 0) \ y \ orthogonal_complement M\ unfolding orthogonal_complement_def by simp abbreviation is_orthogonal::\'a::complex_inner \ 'a \ bool\ where \is_orthogonal x y \ x \\<^sub>C y = 0\ bundle orthogonal_notation begin notation is_orthogonal (infixl "\" 69) end bundle no_orthogonal_notation begin no_notation is_orthogonal (infixl "\" 69) end lemma is_orthogonal_sym: "is_orthogonal \ \ = is_orthogonal \ \" by (metis cinner_commute' complex_cnj_zero) lemma is_orthogonal_sgn_right[simp]: \is_orthogonal e (sgn f) \ is_orthogonal e f\ proof (cases \f = 0\) case True then show ?thesis by simp next case False have \cinner e (sgn f) = cinner e f / norm f\ by (simp add: sgn_div_norm divide_inverse scaleR_scaleC) moreover have \norm f \ 0\ by (simp add: False) ultimately show ?thesis by force qed lemma is_orthogonal_sgn_left[simp]: \is_orthogonal (sgn e) f \ is_orthogonal e f\ by (simp add: is_orthogonal_sym) lemma orthogonal_complement_closed_subspace[simp]: "closed_csubspace (orthogonal_complement A)" for A :: \('a::complex_inner) set\ proof (intro closed_csubspace.intro complex_vector.subspaceI) fix x y and c show \0 \ orthogonal_complement A\ by (rule orthogonal_complementI, simp) show \x + y \ orthogonal_complement A\ if \x \ orthogonal_complement A\ and \y \ orthogonal_complement A\ using that by (auto intro!: orthogonal_complementI dest!: orthogonal_complement_orthoI simp add: cinner_add_left) show \c *\<^sub>C x \ orthogonal_complement A\ if \x \ orthogonal_complement A\ using that by (auto intro!: orthogonal_complementI dest!: orthogonal_complement_orthoI) show "closed (orthogonal_complement A)" proof (auto simp add: closed_sequential_limits, rename_tac an a) fix an a assume ortho: \\n::nat. an n \ orthogonal_complement A\ assume lim: \an \ a\ have \\ y \ A. \ n. is_orthogonal y (an n)\ using orthogonal_complement_orthoI' by (simp add: orthogonal_complement_orthoI' ortho) moreover have \isCont (\ x. y \\<^sub>C x) a\ for y using bounded_clinear_cinner_right clinear_continuous_at by (simp add: clinear_continuous_at bounded_clinear_cinner_right) ultimately have \(\ n. (\ v. y \\<^sub>C v) (an n)) \ (\ v. y \\<^sub>C v) a\ for y using isCont_tendsto_compose by (simp add: isCont_tendsto_compose lim) hence \\ y\A. (\ n. y \\<^sub>C an n) \ y \\<^sub>C a\ by simp hence \\ y\A. (\ n. 0) \ y \\<^sub>C a\ using \\ y \ A. \ n. is_orthogonal y (an n)\ by fastforce hence \\ y \ A. is_orthogonal y a\ using limI by fastforce then show \a \ orthogonal_complement A\ by (simp add: orthogonal_complementI is_orthogonal_sym) qed qed lemma orthogonal_complement_zero_intersection: assumes "0\M" shows \M \ (orthogonal_complement M) = {0}\ proof - have "x=0" if "x\M" and "x\orthogonal_complement M" for x proof - from that have "is_orthogonal x x" unfolding orthogonal_complement_def by auto thus "x=0" by auto qed with assms show ?thesis unfolding orthogonal_complement_def by auto qed lemma is_orthogonal_closure_cspan: assumes "\x y. x \ X \ y \ Y \ is_orthogonal x y" assumes \x \ closure (cspan X)\ \y \ closure (cspan Y)\ shows "is_orthogonal x y" proof - have *: \cinner x y = 0\ if \y \ Y\ for y using bounded_antilinear_cinner_left apply (rule bounded_antilinear_eq_on[where G=X]) using assms that by auto show \cinner x y = 0\ using bounded_clinear_cinner_right apply (rule bounded_clinear_eq_on_closure[where G=Y]) using * assms by auto qed instantiation ccsubspace :: (complex_inner) "uminus" begin lift_definition uminus_ccsubspace::\'a ccsubspace \ 'a ccsubspace\ is \orthogonal_complement\ by simp instance .. end lemma orthocomplement_top[simp]: \- top = (bot :: 'a::complex_inner ccsubspace)\ \ \For \<^typ>\'a\ of sort \<^class>\chilbert_space\, this is covered by @{thm [source] orthocomplemented_lattice_class.compl_top_eq} already. But here we give it a wider sort.\ apply transfer by (metis Int_UNIV_left UNIV_I orthogonal_complement_zero_intersection) instantiation ccsubspace :: (complex_inner) minus begin lift_definition minus_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" is "\A B. A \ (orthogonal_complement B)" by simp instance.. end definition is_ortho_set :: "'a::complex_inner set \ bool" where \ \Orthogonal set\ \is_ortho_set S \ (\x\S. \y\S. x \ y \ (x \\<^sub>C y) = 0) \ 0 \ S\ definition is_onb where \is_onb E \ is_ortho_set E \ (\b\E. norm b = 1) \ ccspan E = top\ lemma is_ortho_set_empty[simp]: "is_ortho_set {}" unfolding is_ortho_set_def by auto lemma is_ortho_set_antimono: \A \ B \ is_ortho_set B \ is_ortho_set A\ unfolding is_ortho_set_def by auto lemma orthogonal_complement_of_closure: fixes A ::"('a::complex_inner) set" shows "orthogonal_complement A = orthogonal_complement (closure A)" proof- have s1: \is_orthogonal y x\ if a1: "x \ (orthogonal_complement A)" and a2: \y \ closure A\ for x y proof- have \\ y \ A. is_orthogonal y x\ by (simp add: a1 orthogonal_complement_orthoI') then obtain yy where \\ n. yy n \ A\ and \yy \ y\ using a2 closure_sequential by blast have \isCont (\ t. t \\<^sub>C x) y\ by simp hence \(\ n. yy n \\<^sub>C x) \ y \\<^sub>C x\ using \yy \ y\ isCont_tendsto_compose by fastforce hence \(\ n. 0) \ y \\<^sub>C x\ using \\ y \ A. is_orthogonal y x\ \\ n. yy n \ A\ by simp thus ?thesis using limI by force qed hence "x \ orthogonal_complement (closure A)" if a1: "x \ (orthogonal_complement A)" for x using that by (meson orthogonal_complementI is_orthogonal_sym) moreover have \x \ (orthogonal_complement A)\ if "x \ (orthogonal_complement (closure A))" for x using that by (meson closure_subset orthogonal_complement_orthoI orthogonal_complementI subset_eq) ultimately show ?thesis by blast qed lemma is_orthogonal_closure: assumes \\s. s \ S \ is_orthogonal a s\ assumes \x \ closure S\ shows \is_orthogonal a x\ by (metis assms(1) assms(2) orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI) lemma is_orthogonal_cspan: assumes a1: "\s. s \ S \ is_orthogonal a s" and a3: "x \ cspan S" shows "is_orthogonal a x" proof- have "\t r. finite t \ t \ S \ (\a\t. r a *\<^sub>C a) = x" using complex_vector.span_explicit by (smt a3 mem_Collect_eq) then obtain t r where b1: "finite t" and b2: "t \ S" and b3: "(\a\t. r a *\<^sub>C a) = x" by blast have x1: "is_orthogonal a i" if "i\t" for i using b2 a1 that by blast have "a \\<^sub>C x = a \\<^sub>C (\i\t. r i *\<^sub>C i)" by (simp add: b3) also have "\ = (\i\t. r i *\<^sub>C (a \\<^sub>C i))" by (simp add: cinner_sum_right) also have "\ = 0" using x1 by simp finally show ?thesis. qed lemma ccspan_leq_ortho_ccspan: assumes "\s t. s\S \ t\T \ is_orthogonal s t" shows "ccspan S \ - (ccspan T)" using assms apply transfer by (smt (verit, ccfv_threshold) is_orthogonal_closure is_orthogonal_cspan is_orthogonal_sym orthogonal_complementI subsetI) lemma double_orthogonal_complement_increasing[simp]: shows "M \ orthogonal_complement (orthogonal_complement M)" proof (rule subsetI) fix x assume s1: "x \ M" have \\ y \ (orthogonal_complement M). is_orthogonal x y\ using s1 orthogonal_complement_orthoI' by auto hence \x \ orthogonal_complement (orthogonal_complement M)\ by (simp add: orthogonal_complement_def) then show "x \ orthogonal_complement (orthogonal_complement M)" by blast qed lemma orthonormal_basis_of_cspan: fixes S::"'a::complex_inner set" assumes "finite S" shows "\A. is_ortho_set A \ (\x\A. norm x = 1) \ cspan A = cspan S \ finite A" proof (use assms in induction) case empty show ?case apply (rule exI[of _ "{}"]) by auto next case (insert s S) from insert.IH obtain A where orthoA: "is_ortho_set A" and normA: "\x. x\A \ norm x = 1" and spanA: "cspan A = cspan S" and finiteA: "finite A" by auto show ?case proof (cases \s \ cspan S\) case True then have \cspan (insert s S) = cspan S\ by (simp add: complex_vector.span_redundant) with orthoA normA spanA finiteA show ?thesis by auto next case False obtain a where a_ortho: \\x. x\A \ is_orthogonal x a\ and sa_span: \s - a \ cspan A\ proof (atomize_elim, use \finite A\ \is_ortho_set A\ in induction) case empty then show ?case by auto next case (insert x A) then obtain a where orthoA: \\x. x \ A \ is_orthogonal x a\ and sa: \s - a \ cspan A\ by (meson is_ortho_set_antimono subset_insertI) define a' where \a' = a - cinner x a *\<^sub>C inverse (cinner x x) *\<^sub>C x\ have \is_orthogonal x a'\ unfolding a'_def cinner_diff_right cinner_scaleC_right apply (cases \cinner x x = 0\) by auto have orthoA: \is_orthogonal y a'\ if \y \ A\ for y unfolding a'_def cinner_diff_right cinner_scaleC_right apply auto by (metis insert.prems insertCI is_ortho_set_def mult_not_zero orthoA that) have \s - a' \ cspan (insert x A)\ unfolding a'_def apply auto by (metis (no_types, lifting) complex_vector.span_breakdown_eq diff_add_cancel diff_diff_add sa) with \is_orthogonal x a'\ orthoA show ?case apply (rule_tac exI[of _ a']) by auto qed from False sa_span have \a \ 0\ unfolding spanA by auto define a' where \a' = inverse (norm a) *\<^sub>C a\ with \a \ 0\ have \norm a' = 1\ by (simp add: norm_inverse) have a: \a = norm a *\<^sub>C a'\ by (simp add: \a \ 0\ a'_def) from sa_span spanA have a'_span: \a' \ cspan (insert s S)\ unfolding a'_def by (metis complex_vector.eq_span_insert_eq complex_vector.span_scale complex_vector.span_superset in_mono insertI1) from sa_span have s_span: \s \ cspan (insert a' A)\ apply (subst (asm) a) using complex_vector.span_breakdown_eq by blast from \a \ 0\ a_ortho orthoA have ortho: "is_ortho_set (insert a' A)" unfolding is_ortho_set_def a'_def apply auto by (meson is_orthogonal_sym) have span: \cspan (insert a' A) = cspan (insert s S)\ using a'_span s_span spanA apply auto apply (metis (full_types) complex_vector.span_breakdown_eq complex_vector.span_redundant insert_commute s_span) by (metis (full_types) complex_vector.span_breakdown_eq complex_vector.span_redundant insert_commute s_span) show ?thesis apply (rule exI[of _ \insert a' A\]) by (simp add: ortho \norm a' = 1\ normA finiteA span) qed qed lemma is_ortho_set_cindependent: assumes "is_ortho_set A" shows "cindependent A" proof - have "u v = 0" if b1: "finite t" and b2: "t \ A" and b3: "(\v\t. u v *\<^sub>C v) = 0" and b4: "v \ t" for t u v proof - have "is_orthogonal v v'" if c1: "v'\t-{v}" for v' by (metis DiffE assms b2 b4 insertI1 is_ortho_set_antimono is_ortho_set_def that) hence sum0: "(\v'\t-{v}. u v' * (v \\<^sub>C v')) = 0" by simp have "v \\<^sub>C (\v'\t. u v' *\<^sub>C v') = (\v'\t. u v' * (v \\<^sub>C v'))" using b1 by (metis (mono_tags, lifting) cinner_scaleC_right cinner_sum_right sum.cong) also have "\ = u v * (v \\<^sub>C v) + (\v'\t-{v}. u v' * (v \\<^sub>C v'))" by (meson b1 b4 sum.remove) also have "\ = u v * (v \\<^sub>C v)" using sum0 by simp finally have "v \\<^sub>C (\v'\t. u v' *\<^sub>C v') = u v * (v \\<^sub>C v)" by blast hence "u v * (v \\<^sub>C v) = 0" using b3 by simp moreover have "(v \\<^sub>C v) \ 0" using assms is_ortho_set_def b2 b4 by auto ultimately show "u v = 0" by simp qed thus ?thesis using complex_vector.independent_explicit_module by (smt cdependent_raw_def) qed lemma onb_expansion_finite: includes notation_norm fixes T::\'a::{complex_inner,cfinite_dim} set\ assumes a1: \cspan T = UNIV\ and a3: \is_ortho_set T\ and a4: \\t. t\T \ \t\ = 1\ shows \x = (\t\T. (t \\<^sub>C x) *\<^sub>C t)\ proof - have \finite T\ apply (rule cindependent_cfinite_dim_finite) by (simp add: a3 is_ortho_set_cindependent) have \closure (complex_vector.span T) = complex_vector.span T\ by (simp add: a1) have \{\a\t. r a *\<^sub>C a |t r. finite t \ t \ T} = {\a\T. r a *\<^sub>C a |r. True}\ apply auto apply (rule_tac x=\\a. if a \ t then r a else 0\ in exI) apply (simp add: \finite T\ sum.mono_neutral_cong_right) using \finite T\ by blast have f1: "\A. {a. \Aa f. (a::'a) = (\a\Aa. f a *\<^sub>C a) \ finite Aa \ Aa \ A} = cspan A" by (simp add: complex_vector.span_explicit) have f2: "\a. (\f. a = (\a\T. f a *\<^sub>C a)) \ (\A. (\f. a \ (\a\A. f a *\<^sub>C a)) \ infinite A \ \ A \ T)" using \{\a\t. r a *\<^sub>C a |t r. finite t \ t \ T} = {\a\T. r a *\<^sub>C a |r. True}\ by auto have f3: "\A a. (\Aa f. (a::'a) = (\a\Aa. f a *\<^sub>C a) \ finite Aa \ Aa \ A) \ a \ cspan A" using f1 by blast have "cspan T = UNIV" by (metis (full_types, lifting) \complex_vector.span T = UNIV\) hence \\ r. x = (\ a\T. r a *\<^sub>C a)\ using f3 f2 by blast then obtain r where \x = (\ a\T. r a *\<^sub>C a)\ by blast have \r a = a \\<^sub>C x\ if \a \ T\ for a proof- have \norm a = 1\ using a4 by (simp add: \a \ T\) moreover have \norm a = sqrt (norm (a \\<^sub>C a))\ using norm_eq_sqrt_cinner by auto ultimately have \sqrt (norm (a \\<^sub>C a)) = 1\ by simp hence \norm (a \\<^sub>C a) = 1\ using real_sqrt_eq_1_iff by blast moreover have \(a \\<^sub>C a) \ \\ by (simp add: cinner_real) moreover have \(a \\<^sub>C a) \ 0\ using cinner_ge_zero by blast ultimately have w1: \(a \\<^sub>C a) = 1\ by (metis \0 \ (a \\<^sub>C a)\ \cmod (a \\<^sub>C a) = 1\ complex_of_real_cmod of_real_1) have \r t * (a \\<^sub>C t) = 0\ if \t \ T-{a}\ for t by (metis DiffD1 DiffD2 \a \ T\ a3 is_ortho_set_def mult_eq_0_iff singletonI that) hence s1: \(\ t\T-{a}. r t * (a \\<^sub>C t)) = 0\ by (simp add: \\t. t \ T - {a} \ r t * (a \\<^sub>C t) = 0\) have \(a \\<^sub>C x) = a \\<^sub>C (\ t\T. r t *\<^sub>C t)\ using \x = (\ a\T. r a *\<^sub>C a)\ by simp also have \\ = (\ t\T. a \\<^sub>C (r t *\<^sub>C t))\ using cinner_sum_right by blast also have \\ = (\ t\T. r t * (a \\<^sub>C t))\ by simp also have \\ = r a * (a \\<^sub>C a) + (\ t\T-{a}. r t * (a \\<^sub>C t))\ using \a \ T\ by (meson \finite T\ sum.remove) also have \\ = r a * (a \\<^sub>C a)\ using s1 by simp also have \\ = r a\ by (simp add: w1) finally show ?thesis by auto qed thus ?thesis using \x = (\ a\T. r a *\<^sub>C a)\ by fastforce qed lemma is_ortho_set_singleton[simp]: \is_ortho_set {x} \ x \ 0\ by (simp add: is_ortho_set_def) lemma orthogonal_complement_antimono[simp]: fixes A B :: \('a::complex_inner) set\ assumes "A \ B" shows \orthogonal_complement A \ orthogonal_complement B\ by (meson assms orthogonal_complementI orthogonal_complement_orthoI' subsetD subsetI) lemma orthogonal_complement_UNIV[simp]: "orthogonal_complement UNIV = {0}" by (metis Int_UNIV_left complex_vector.subspace_UNIV complex_vector.subspace_def orthogonal_complement_zero_intersection) lemma orthogonal_complement_zero[simp]: "orthogonal_complement {0} = UNIV" unfolding orthogonal_complement_def by auto lemma mem_ortho_ccspanI: assumes \\y. y \ S \ is_orthogonal x y\ shows \x \ space_as_set (- ccspan S)\ proof - have \x \ space_as_set (ccspan {x})\ using ccspan_superset by blast also have \\ \ space_as_set (- ccspan S)\ apply (simp add: flip: less_eq_ccsubspace.rep_eq) apply (rule ccspan_leq_ortho_ccspan) using assms by auto finally show ?thesis by - qed subsection \Projections\ lemma smallest_norm_exists: \ \Theorem 2.5 in \<^cite>\conway2013course\ (inside the proof)\ includes notation_norm fixes M :: \'a::chilbert_space set\ assumes q1: \convex M\ and q2: \closed M\ and q3: \M \ {}\ shows \\k. is_arg_min (\ x. \x\) (\ t. t \ M) k\ proof - define d where \d = Inf { \x\^2 | x. x \ M }\ have w4: \{ \x\^2 | x. x \ M } \ {}\ by (simp add: assms(3)) have \\ x. \x\^2 \ 0\ by simp hence bdd_below1: \bdd_below { \x\^2 | x. x \ M }\ by fastforce have \d \ \x\^2\ if a1: "x \ M" for x proof- have "\v. (\w. Re (v \\<^sub>C v) = \w\\<^sup>2 \ w \ M) \ v \ M" by (metis (no_types) power2_norm_eq_cinner') hence "Re (x \\<^sub>C x) \ {\v\\<^sup>2 |v. v \ M}" using a1 by blast thus ?thesis unfolding d_def by (metis (lifting) bdd_below1 cInf_lower power2_norm_eq_cinner') qed have \\ \ > 0. \ t \ { \x\^2 | x. x \ M }. t < d + \\ unfolding d_def using w4 bdd_below1 by (meson cInf_lessD less_add_same_cancel1) hence \\ \ > 0. \ x \ M. \x\^2 < d + \\ by auto hence \\ \ > 0. \ x \ M. \x\^2 < d + \\ by (simp add: \\x. x \ M \ d \ \x\\<^sup>2\) hence w1: \\ n::nat. \ x \ M. \x\^2 < d + 1/(n+1)\ by auto then obtain r::\nat \ 'a\ where w2: \\ n. r n \ M \ \ r n \^2 < d + 1/(n+1)\ by metis have w3: \\ n. r n \ M\ by (simp add: w2) have \\ n. \ r n \^2 < d + 1/(n+1)\ by (simp add: w2) have w5: \\ (r n) - (r m) \^2 < 2*(1/(n+1) + 1/(m+1))\ for m n proof- have w6: \\ r n \^2 < d + 1/(n+1)\ by (metis w2 of_nat_1 of_nat_add) have \ \ r m \^2 < d + 1/(m+1)\ by (metis w2 of_nat_1 of_nat_add) have \(r n) \ M\ by (simp add: \\n. r n \ M\) moreover have \(r m) \ M\ by (simp add: \\n. r n \ M\) ultimately have \(1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \ M\ using \convex M\ by (simp add: convexD) hence \\ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2 \ d\ by (simp add: \\x. x \ M \ d \ \x\\<^sup>2\) have \\ (1/2) *\<^sub>R (r n) - (1/2) *\<^sub>R (r m) \^2 = (1/2)*( \ r n \^2 + \ r m \^2 ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ by (smt (z3) div_by_1 field_sum_of_halves nonzero_mult_div_cancel_left parallelogram_law polar_identity power2_norm_eq_cinner' scaleR_collapse times_divide_eq_left) also have \... < (1/2)*( d + 1/(n+1) + \ r m \^2 ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ using \\r n\\<^sup>2 < d + 1 / real (n + 1)\ by auto also have \... < (1/2)*( d + 1/(n+1) + d + 1/(m+1) ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ using \\r m\\<^sup>2 < d + 1 / real (m + 1)\ by auto also have \... \ (1/2)*( d + 1/(n+1) + d + 1/(m+1) ) - d\ by (simp add: \d \ \(1 / 2) *\<^sub>R r n + (1 / 2) *\<^sub>R r m\\<^sup>2\) also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) + 2*d ) - d\ by simp also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + (1/2)*(2*d) - d\ by (simp add: distrib_left) also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + d - d\ by simp also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) )\ by simp finally have \ \(1 / 2) *\<^sub>R r n - (1 / 2) *\<^sub>R r m\\<^sup>2 < 1 / 2 * (1 / real (n + 1) + 1 / real (m + 1)) \ by blast hence \ \(1 / 2) *\<^sub>R (r n - r m) \\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (simp add: real_vector.scale_right_diff_distrib) hence \ ((1 / 2)*\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by simp hence \ (1 / 2)^2*(\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (metis power_mult_distrib) hence \ (1 / 4) *(\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (simp add: power_divide) hence \ \ (r n - r m) \\<^sup>2 < 2 * (1 / real (n + 1) + 1 / real (m + 1)) \ by simp thus ?thesis by (metis of_nat_1 of_nat_add) qed hence "\ N. \ n m. n \ N \ m \ N \ \ (r n) - (r m) \^2 < \^2" if "\ > 0" for \ proof- obtain N::nat where \1/(N + 1) < \^2/4\ using LIMSEQ_ignore_initial_segment[OF lim_inverse_n', where k=1] by (metis Suc_eq_plus1 \0 < \\ nat_approx_posE zero_less_divide_iff zero_less_numeral zero_less_power ) hence \4/(N + 1) < \^2\ by simp have "2*(1/(n+1) + 1/(m+1)) < \^2" if f1: "n \ N" and f2: "m \ N" for m n::nat proof- have \1/(n+1) \ 1/(N+1)\ by (simp add: f1 linordered_field_class.frac_le) moreover have \1/(m+1) \ 1/(N+1)\ by (simp add: f2 linordered_field_class.frac_le) ultimately have \2*(1/(n+1) + 1/(m+1)) \ 4/(N+1)\ by simp thus ?thesis using \4/(N + 1) < \^2\ by linarith qed hence "\ (r n) - (r m) \^2 < \^2" if y1: "n \ N" and y2: "m \ N" for m n::nat using that by (smt \\n m. \r n - r m\\<^sup>2 < 2 * (1 / (real n + 1) + 1 / (real m + 1))\ of_nat_1 of_nat_add) thus ?thesis by blast qed hence \\ \ > 0. \ N::nat. \ n m::nat. n \ N \ m \ N \ \ (r n) - (r m) \^2 < \^2\ by blast hence \\ \ > 0. \ N::nat. \ n m::nat. n \ N \ m \ N \ \ (r n) - (r m) \ < \\ by (meson less_eq_real_def power_less_imp_less_base) hence \Cauchy r\ using CauchyI by fastforce then obtain k where \r \ k\ using convergent_eq_Cauchy by auto have \k \ M\ using \closed M\ using \\n. r n \ M\ \r \ k\ closed_sequentially by auto have \(\ n. \ r n \^2) \ \ k \^2\ by (simp add: \r \ k\ tendsto_norm tendsto_power) moreover have \(\ n. \ r n \^2) \ d\ proof- have \\\ r n \^2 - d\ < 1/(n+1)\ for n :: nat using \\x. x \ M \ d \ \x\\<^sup>2\ \\n. r n \ M \ \r n\\<^sup>2 < d + 1 / (real n + 1)\ of_nat_1 of_nat_add by smt moreover have \(\n. 1 / real (n + 1)) \ 0\ using LIMSEQ_ignore_initial_segment[OF lim_inverse_n', where k=1] by blast ultimately have \(\ n. \\ r n \^2 - d\ ) \ 0\ by (simp add: LIMSEQ_norm_0) hence \(\ n. \ r n \^2 - d ) \ 0\ by (simp add: tendsto_rabs_zero_iff) moreover have \(\ n. d ) \ d\ by simp ultimately have \(\ n. (\ r n \^2 - d)+d ) \ 0+d\ using tendsto_add by fastforce thus ?thesis by simp qed ultimately have \d = \ k \^2\ using LIMSEQ_unique by auto hence \t \ M \ \ k \^2 \ \ t \^2\ for t using \\x. x \ M \ d \ \x\\<^sup>2\ by auto hence q1: \\ k. is_arg_min (\ x. \x\^2) (\ t. t \ M) k\ using \k \ M\ is_arg_min_def \d = \k\\<^sup>2\ by smt thus \\ k. is_arg_min (\ x. \x\) (\ t. t \ M) k\ by (smt is_arg_min_def norm_ge_zero power2_eq_square power2_le_imp_le) qed lemma smallest_norm_unique: \ \Theorem 2.5 in \<^cite>\conway2013course\ (inside the proof)\ includes notation_norm fixes M :: \'a::complex_inner set\ assumes q1: \convex M\ assumes r: \is_arg_min (\ x. \x\) (\ t. t \ M) r\ assumes s: \is_arg_min (\ x. \x\) (\ t. t \ M) s\ shows \r = s\ proof - have \r \ M\ using \is_arg_min (\x. \x\) (\ t. t \ M) r\ by (simp add: is_arg_min_def) moreover have \s \ M\ using \is_arg_min (\x. \x\) (\ t. t \ M) s\ by (simp add: is_arg_min_def) ultimately have \((1/2) *\<^sub>R r + (1/2) *\<^sub>R s) \ M\ using \convex M\ by (simp add: convexD) hence \\r\ \ \ (1/2) *\<^sub>R r + (1/2) *\<^sub>R s \\ by (metis is_arg_min_linorder r) hence u2: \\r\^2 \ \ (1/2) *\<^sub>R r + (1/2) *\<^sub>R s \^2\ using norm_ge_zero power_mono by blast have \\r\ \ \s\\ using r s is_arg_min_def by (metis is_arg_min_linorder) moreover have \\s\ \ \r\\ using r s is_arg_min_def by (metis is_arg_min_linorder) ultimately have u3: \\r\ = \s\\ by simp have \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \^2 \ 0\ using u2 u3 parallelogram_law by (smt (verit, ccfv_SIG) polar_identity_minus power2_norm_eq_cinner' scaleR_add_right scaleR_half_double) hence \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \^2 = 0\ by simp hence \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \ = 0\ by auto hence \(1/2) *\<^sub>R r - (1/2) *\<^sub>R s = 0\ using norm_eq_zero by blast thus ?thesis by simp qed theorem smallest_dist_exists: \ \Theorem 2.5 in \<^cite>\conway2013course\\ fixes M::\'a::chilbert_space set\ and h assumes a1: \convex M\ and a2: \closed M\ and a3: \M \ {}\ shows \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ proof - have *: "is_arg_min (\x. dist x h) (\x. x\M) (k+h) \ is_arg_min (\x. norm x) (\x. x\(\x. x-h) ` M) k" for k unfolding dist_norm is_arg_min_def apply auto using add_implies_diff by blast have \\k. is_arg_min (\x. dist x h) (\x. x\M) (k+h)\ apply (subst *) apply (rule smallest_norm_exists) using assms by (auto simp: closed_translation_subtract) then show \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ by metis qed theorem smallest_dist_unique: \ \Theorem 2.5 in \<^cite>\conway2013course\\ fixes M::\'a::complex_inner set\ and h assumes a1: \convex M\ assumes \is_arg_min (\ x. dist x h) (\ x. x \ M) r\ assumes \is_arg_min (\ x. dist x h) (\ x. x \ M) s\ shows \r = s\ proof- have *: "is_arg_min (\x. dist x h) (\x. x\M) k \ is_arg_min (\x. norm x) (\x. x\(\x. x-h) ` M) (k-h)" for k unfolding dist_norm is_arg_min_def by auto have \r - h = s - h\ using _ assms(2,3)[unfolded *] apply (rule smallest_norm_unique) by (simp add: a1) thus \r = s\ by auto qed \ \Theorem 2.6 in \<^cite>\conway2013course\\ theorem smallest_dist_is_ortho: fixes M::\'a::complex_inner set\ and h k::'a assumes b1: \closed_csubspace M\ shows \(is_arg_min (\ x. dist x h) (\ x. x \ M) k) \ h - k \ orthogonal_complement M \ k \ M\ proof - include notation_norm have \csubspace M\ using \closed_csubspace M\ unfolding closed_csubspace_def by blast have r1: \2 * Re ((h - k) \\<^sub>C f) \ \ f \^2\ if "f \ M" and \k \ M\ and \is_arg_min (\x. dist x h) (\ x. x \ M) k\ for f proof- have \k + f \ M\ using \csubspace M\ by (simp add:complex_vector.subspace_add that) have "\f A a b. \ is_arg_min f (\ x. x \ A) (a::'a) \ (f a::real) \ f b \ b \ A" by (metis (no_types) is_arg_min_linorder) hence "dist k h \ dist (f + k) h" by (metis \is_arg_min (\x. dist x h) (\ x. x \ M) k\ \k + f \ M\ add.commute) hence \dist h k \ dist h (k + f)\ by (simp add: add.commute dist_commute) hence \\ h - k \ \ \ h - (k + f) \\ by (simp add: dist_norm) hence \\ h - k \^2 \ \ h - (k + f) \^2\ by (simp add: power_mono) also have \... \ \ (h - k) - f \^2\ by (simp add: diff_diff_add) also have \... \ \ (h - k) \^2 + \ f \^2 - 2 * Re ((h - k) \\<^sub>C f)\ by (simp add: polar_identity_minus) finally have \\ (h - k) \^2 \ \ (h - k) \^2 + \ f \^2 - 2 * Re ((h - k) \\<^sub>C f)\ by simp thus ?thesis by simp qed have q4: \\ c > 0. 2 * Re ((h - k) \\<^sub>C f) \ c\ if \\c>0. 2 * Re ((h - k) \\<^sub>C f) \ c * \f\\<^sup>2\ for f proof (cases \\ f \^2 > 0\) case True hence \\ c > 0. 2 * Re (((h - k) \\<^sub>C f)) \ (c/\ f \^2)*\ f \^2\ using that linordered_field_class.divide_pos_pos by blast thus ?thesis using True by auto next case False hence \\ f \^2 = 0\ by simp thus ?thesis by auto qed have q3: \\ c::real. c > 0 \ 2 * Re (((h - k) \\<^sub>C f)) \ 0\ if a3: \\f. f \ M \ (\c>0. 2 * Re ((h - k) \\<^sub>C f) \ c * \f\\<^sup>2)\ and a2: "f \ M" and a1: "is_arg_min (\ x. dist x h) (\ x. x \ M) k" for f proof- have \\ c > 0. 2 * Re (((h - k) \\<^sub>C f)) \ c*\ f \^2\ by (simp add: that ) thus ?thesis using q4 by smt qed have w2: "h - k \ orthogonal_complement M \ k \ M" if a1: "is_arg_min (\ x. dist x h) (\ x. x \ M) k" proof- have \k \ M\ using is_arg_min_def that by fastforce hence \\ f. f \ M \ 2 * Re (((h - k) \\<^sub>C f)) \ \ f \^2\ using r1 by (simp add: that) have \\ f. f \ M \ (\ c::real. 2 * Re ((h - k) \\<^sub>C (c *\<^sub>R f)) \ \ c *\<^sub>R f \^2)\ using assms scaleR_scaleC complex_vector.subspace_def \csubspace M\ by (metis \\f. f \ M \ 2 * Re ((h - k) \\<^sub>C f) \ \f\\<^sup>2\) hence \\ f. f \ M \ (\ c::real. c * (2 * Re (((h - k) \\<^sub>C f))) \ \ c *\<^sub>R f \^2)\ by (metis Re_complex_of_real cinner_scaleC_right complex_add_cnj complex_cnj_complex_of_real complex_cnj_mult of_real_mult scaleR_scaleC semiring_normalization_rules(34)) hence \\ f. f \ M \ (\ c::real. c * (2 * Re (((h - k) \\<^sub>C f))) \ \c\^2*\ f \^2)\ by (simp add: power_mult_distrib) hence \\ f. f \ M \ (\ c::real. c * (2 * Re (((h - k) \\<^sub>C f))) \ c^2*\ f \^2)\ by auto hence \\ f. f \ M \ (\ c::real. c > 0 \ c * (2 * Re (((h - k) \\<^sub>C f))) \ c^2*\ f \^2)\ by simp hence \\ f. f \ M \ (\ c::real. c > 0 \ c*(2 * Re (((h - k) \\<^sub>C f))) \ c*(c*\ f \^2))\ by (simp add: power2_eq_square) hence q4: \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (((h - k) \\<^sub>C f)) \ c*\ f \^2)\ by simp have \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (((h - k) \\<^sub>C f)) \ 0)\ using q3 by (simp add: q4 that) hence \\ f. f \ M \ (\ c::real. c > 0 \ (2 * Re ((h - k) \\<^sub>C (-1 *\<^sub>R f))) \ 0)\ using assms scaleR_scaleC complex_vector.subspace_def by (metis \csubspace M\) hence \\ f. f \ M \ (\ c::real. c > 0 \ -(2 * Re (((h - k) \\<^sub>C f))) \ 0)\ by simp hence \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (((h - k) \\<^sub>C f)) \ 0)\ by simp hence \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (((h - k) \\<^sub>C f)) = 0)\ using \\ f. f \ M \ (\ c::real. c > 0 \ (2 * Re (((h - k) \\<^sub>C f))) \ 0)\ by fastforce have \\ f. f \ M \ ((1::real) > 0 \ 2 * Re (((h - k) \\<^sub>C f)) = 0)\ using \\f. f \ M \ (\c>0. 2 * Re (((h - k) \\<^sub>C f) ) = 0)\ by blast hence \\ f. f \ M \ 2 * Re (((h - k) \\<^sub>C f)) = 0\ by simp hence \\ f. f \ M \ Re (((h - k) \\<^sub>C f)) = 0\ by simp have \\ f. f \ M \ Re ((h - k) \\<^sub>C ((Complex 0 (-1)) *\<^sub>C f)) = 0\ using assms complex_vector.subspace_def \csubspace M\ by (metis \\f. f \ M \ Re ((h - k) \\<^sub>C f) = 0\) hence \\ f. f \ M \ Re ( (Complex 0 (-1))*(((h - k) \\<^sub>C f)) ) = 0\ by simp hence \\ f. f \ M \ Im (((h - k) \\<^sub>C f)) = 0\ using Complex_eq_neg_1 Re_i_times cinner_scaleC_right complex_of_real_def by auto have \\ f. f \ M \ (((h - k) \\<^sub>C f)) = 0\ using complex_eq_iff by (simp add: \\f. f \ M \ Im ((h - k) \\<^sub>C f) = 0\ \\f. f \ M \ Re ((h - k) \\<^sub>C f) = 0\) hence \h - k \ orthogonal_complement M \ k \ M\ by (simp add: \k \ M\ orthogonal_complementI) have \\ c. c *\<^sub>R f \ M\ if "f \ M" for f using that scaleR_scaleC \csubspace M\ complex_vector.subspace_def by (simp add: complex_vector.subspace_def scaleR_scaleC) have \((h - k) \\<^sub>C f) = 0\ if "f \ M" for f using \h - k \ orthogonal_complement M \ k \ M\ orthogonal_complement_orthoI that by auto hence \h - k \ orthogonal_complement M\ by (simp add: orthogonal_complement_def) thus ?thesis using \k \ M\ by auto qed have q1: \dist h k \ dist h f \ if "f \ M" and \h - k \ orthogonal_complement M \ k \ M\ for f proof- have \(h - k) \\<^sub>C (k - f) = 0\ by (metis (no_types, lifting) that cinner_diff_right diff_0_right orthogonal_complement_orthoI that) have \\ h - f \^2 = \ (h - k) + (k - f) \^2\ by simp also have \... = \ h - k \^2 + \ k - f \^2\ using \((h - k) \\<^sub>C (k - f)) = 0\ pythagorean_theorem by blast also have \... \ \ h - k \^2\ by simp finally have \\h - k\\<^sup>2 \ \h - f\\<^sup>2 \ by blast hence \\h - k\ \ \h - f\\ using norm_ge_zero power2_le_imp_le by blast thus ?thesis by (simp add: dist_norm) qed have w1: "is_arg_min (\ x. dist x h) (\ x. x \ M) k" if "h - k \ orthogonal_complement M \ k \ M" proof- have \h - k \ orthogonal_complement M\ using that by blast have \k \ M\ using \h - k \ orthogonal_complement M \ k \ M\ by blast thus ?thesis by (metis (no_types, lifting) dist_commute is_arg_min_linorder q1 that) qed show ?thesis using w1 w2 by blast qed corollary orthog_proj_exists: fixes M :: \'a::chilbert_space set\ assumes \closed_csubspace M\ shows \\k. h - k \ orthogonal_complement M \ k \ M\ proof - from \closed_csubspace M\ have \M \ {}\ using closed_csubspace.subspace complex_vector.subspace_0 by blast have \closed M\ using \closed_csubspace M\ by (simp add: closed_csubspace.closed) have \convex M\ using \closed_csubspace M\ by (simp) have \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ by (simp add: smallest_dist_exists \closed M\ \convex M\ \M \ {}\) thus ?thesis by (simp add: assms smallest_dist_is_ortho) qed corollary orthog_proj_unique: fixes M :: \'a::complex_inner set\ assumes \closed_csubspace M\ assumes \h - r \ orthogonal_complement M \ r \ M\ assumes \h - s \ orthogonal_complement M \ s \ M\ shows \r = s\ using _ assms(2,3) unfolding smallest_dist_is_ortho[OF assms(1), symmetric] apply (rule smallest_dist_unique) using assms(1) by (simp) definition is_projection_on::\('a \ 'a) \ ('a::metric_space) set \ bool\ where \is_projection_on \ M \ (\h. is_arg_min (\ x. dist x h) (\ x. x \ M) (\ h))\ lemma is_projection_on_iff_orthog: \closed_csubspace M \ is_projection_on \ M \ (\h. h - \ h \ orthogonal_complement M \ \ h \ M)\ by (simp add: is_projection_on_def smallest_dist_is_ortho) lemma is_projection_on_exists: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows "\\. is_projection_on \ M" unfolding is_projection_on_def apply (rule choice) using smallest_dist_exists[OF assms] by auto lemma is_projection_on_unique: fixes M :: \'a::complex_inner set\ assumes \convex M\ assumes "is_projection_on \\<^sub>1 M" assumes "is_projection_on \\<^sub>2 M" shows "\\<^sub>1 = \\<^sub>2" using smallest_dist_unique[OF assms(1)] using assms(2,3) unfolding is_projection_on_def by blast definition projection :: \'a::metric_space set \ ('a \ 'a)\ where \projection M = (SOME \. is_projection_on \ M)\ lemma projection_is_projection_on: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows "is_projection_on (projection M) M" by (metis assms(1) assms(2) assms(3) is_projection_on_exists projection_def someI) lemma projection_is_projection_on'[simp]: \ \Common special case of @{thm projection_is_projection_on}\ fixes M :: \'a::chilbert_space set\ assumes \closed_csubspace M\ shows "is_projection_on (projection M) M" apply (rule projection_is_projection_on) apply (auto simp add: assms closed_csubspace.closed) using assms closed_csubspace.subspace complex_vector.subspace_0 by blast lemma projection_orthogonal: fixes M :: \'a::chilbert_space set\ assumes "closed_csubspace M" and \m \ M\ shows \is_orthogonal (h - projection M h) m\ by (metis assms(1) assms(2) closed_csubspace.closed closed_csubspace.subspace csubspace_is_convex empty_iff is_projection_on_iff_orthog orthogonal_complement_orthoI projection_is_projection_on) lemma is_projection_on_in_image: assumes "is_projection_on \ M" shows "\ h \ M" using assms by (simp add: is_arg_min_def is_projection_on_def) lemma is_projection_on_image: assumes "is_projection_on \ M" shows "range \ = M" using assms apply (auto simp: is_projection_on_in_image) by (smt (verit, ccfv_threshold) dist_pos_lt dist_self is_arg_min_def is_projection_on_def rangeI) lemma projection_in_image[simp]: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows \projection M h \ M\ by (simp add: assms(1) assms(2) assms(3) is_projection_on_in_image projection_is_projection_on) lemma projection_image[simp]: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows \range (projection M) = M\ by (simp add: assms(1) assms(2) assms(3) is_projection_on_image projection_is_projection_on) lemma projection_eqI': fixes M :: \'a::complex_inner set\ assumes \convex M\ assumes \is_projection_on f M\ shows \projection M = f\ by (metis assms(1) assms(2) is_projection_on_unique projection_def someI_ex) lemma is_projection_on_eqI: fixes M :: \'a::complex_inner set\ assumes a1: \closed_csubspace M\ and a2: \h - x \ orthogonal_complement M\ and a3: \x \ M\ and a4: \is_projection_on \ M\ shows \\ h = x\ by (meson a1 a2 a3 a4 closed_csubspace.subspace csubspace_is_convex is_projection_on_def smallest_dist_is_ortho smallest_dist_unique) lemma projection_eqI: fixes M :: \('a::chilbert_space) set\ assumes \closed_csubspace M\ and \h - x \ orthogonal_complement M\ and \x \ M\ shows \projection M h = x\ by (metis assms(1) assms(2) assms(3) is_projection_on_iff_orthog orthog_proj_exists projection_def is_projection_on_eqI tfl_some) lemma is_projection_on_fixes_image: fixes M :: \'a::metric_space set\ assumes a1: "is_projection_on \ M" and a3: "x \ M" shows "\ x = x" by (metis a1 a3 dist_pos_lt dist_self is_arg_min_def is_projection_on_def) lemma projection_fixes_image: fixes M :: \('a::chilbert_space) set\ assumes "closed_csubspace M" and "x \ M" shows "projection M x = x" using is_projection_on_fixes_image \ \Theorem 2.7 in \<^cite>\conway2013course\\ by (simp add: assms complex_vector.subspace_0 projection_eqI) lemma is_projection_on_closed: assumes cont_f: \\x. x \ closure M \ isCont f x\ assumes \is_projection_on f M\ shows \closed M\ proof - have \x \ M\ if \s \ x\ and \range s \ M\ for s x proof - from \is_projection_on f M\ \range s \ M\ have \s = (f o s)\ by (simp add: comp_def is_projection_on_fixes_image range_subsetD) also from cont_f \s \ x\ have \(f o s) \ f x\ apply (rule continuous_imp_tendsto) using \s \ x\ \range s \ M\ by (meson closure_sequential range_subsetD) finally have \x = f x\ using \s \ x\ by (simp add: LIMSEQ_unique) then have \x \ range f\ by simp with \is_projection_on f M\ show \x \ M\ by (simp add: is_projection_on_image) qed then show ?thesis by (metis closed_sequential_limits image_subset_iff) qed proposition is_projection_on_reduces_norm: includes notation_norm fixes M :: \('a::complex_inner) set\ assumes \is_projection_on \ M\ and \closed_csubspace M\ shows \\ \ h \ \ \ h \\ proof- have \h - \ h \ orthogonal_complement M\ using assms is_projection_on_iff_orthog by blast hence \\ k \ M. is_orthogonal (h - \ h) k\ using orthogonal_complement_orthoI by blast also have \\ h \ M\ using \is_projection_on \ M\ by (simp add: is_projection_on_in_image) ultimately have \is_orthogonal (h - \ h) (\ h)\ by auto hence \\ \ h \^2 + \ h - \ h \^2 = \ h \^2\ using pythagorean_theorem by fastforce hence \\\ h \^2 \ \ h \^2\ by (smt zero_le_power2) thus ?thesis using norm_ge_zero power2_le_imp_le by blast qed proposition projection_reduces_norm: includes notation_norm fixes M :: \'a::chilbert_space set\ assumes a1: "closed_csubspace M" shows \\ projection M h \ \ \ h \\ using assms is_projection_on_iff_orthog orthog_proj_exists is_projection_on_reduces_norm projection_eqI by blast \ \Theorem 2.7 (version) in \<^cite>\conway2013course\\ theorem is_projection_on_bounded_clinear: fixes M :: \'a::complex_inner set\ assumes a1: "is_projection_on \ M" and a2: "closed_csubspace M" shows "bounded_clinear \" proof have b1: \csubspace (orthogonal_complement M)\ by (simp add: a2) have f1: "\a. a - \ a \ orthogonal_complement M \ \ a \ M" using a1 a2 is_projection_on_iff_orthog by blast hence "c *\<^sub>C x - c *\<^sub>C \ x \ orthogonal_complement M" for c x by (metis (no_types) b1 add_diff_cancel_right' complex_vector.subspace_def diff_add_cancel scaleC_add_right) thus r1: \\ (c *\<^sub>C x) = c *\<^sub>C (\ x)\ for x c using f1 by (meson a2 a1 closed_csubspace.subspace complex_vector.subspace_def is_projection_on_eqI) show r2: \\ (x + y) = (\ x) + (\ y)\ for x y proof- have "\A. \ closed_csubspace (A::'a set) \ csubspace A" by (metis closed_csubspace.subspace) hence "csubspace M" using a2 by auto hence \\ (x + y) - ( (\ x) + (\ y) ) \ M\ by (simp add: complex_vector.subspace_add complex_vector.subspace_diff f1) have \closed_csubspace (orthogonal_complement M)\ using a2 by simp have f1: "\a b. (b::'a) + (a - b) = a" by (metis add.commute diff_add_cancel) have f2: "\a b. (b::'a) - b = a - a" by auto hence f3: "\a. a - a \ orthogonal_complement M" by (simp add: complex_vector.subspace_0) have "\a b. (a \ orthogonal_complement M \ a + b \ orthogonal_complement M) \ b \ orthogonal_complement M" using add_diff_cancel_right' b1 complex_vector.subspace_diff by metis hence "\a b c. (a \ orthogonal_complement M \ c - (b + a) \ orthogonal_complement M) \ c - b \ orthogonal_complement M" using f1 by (metis diff_diff_add) hence f4: "\a b f. (f a - b \ orthogonal_complement M \ a - b \ orthogonal_complement M) \ \ is_projection_on f M" using f1 by (metis a2 is_projection_on_iff_orthog) have f5: "\a b c d. (d::'a) - (c + (b - a)) = d + (a - (b + c))" by auto have "x - \ x \ orthogonal_complement M" using a1 a2 is_projection_on_iff_orthog by blast hence q1: \\ (x + y) - ( (\ x) + (\ y) ) \ orthogonal_complement M\ using f5 f4 f3 by (metis \csubspace (orthogonal_complement M)\ \is_projection_on \ M\ add_diff_eq complex_vector.subspace_diff diff_diff_add diff_diff_eq2) hence \\ (x + y) - ( (\ x) + (\ y) ) \ M \ (orthogonal_complement M)\ by (simp add: \\ (x + y) - (\ x + \ y) \ M\) moreover have \M \ (orthogonal_complement M) = {0}\ by (simp add: \closed_csubspace M\ complex_vector.subspace_0 orthogonal_complement_zero_intersection) ultimately have \\ (x + y) - ( (\ x) + (\ y) ) = 0\ by auto thus ?thesis by simp qed from is_projection_on_reduces_norm show t1: \\ K. \ x. norm (\ x) \ norm x * K\ by (metis a1 a2 mult.left_neutral ordered_field_class.sign_simps(5)) qed theorem projection_bounded_clinear: fixes M :: \('a::chilbert_space) set\ assumes a1: "closed_csubspace M" shows \bounded_clinear (projection M)\ \ \Theorem 2.7 in \<^cite>\conway2013course\\ using assms is_projection_on_iff_orthog orthog_proj_exists is_projection_on_bounded_clinear projection_eqI by blast proposition is_projection_on_idem: fixes M :: \('a::complex_inner) set\ assumes "is_projection_on \ M" shows "\ (\ x) = \ x" using is_projection_on_fixes_image is_projection_on_in_image assms by blast proposition projection_idem: fixes M :: "'a::chilbert_space set" assumes a1: "closed_csubspace M" shows "projection M (projection M x) = projection M x" by (metis assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex equals0D projection_fixes_image projection_in_image) proposition is_projection_on_kernel_is_orthogonal_complement: fixes M :: \'a::complex_inner set\ assumes a1: "is_projection_on \ M" and a2: "closed_csubspace M" shows "\ -` {0} = orthogonal_complement M" proof- have "x \ (\ -` {0})" if "x \ orthogonal_complement M" for x by (smt (verit, ccfv_SIG) a1 a2 closed_csubspace_def complex_vector.subspace_def complex_vector.subspace_diff is_projection_on_eqI orthogonal_complement_closed_subspace that vimage_singleton_eq) moreover have "x \ orthogonal_complement M" if s1: "x \ \ -` {0}" for x by (metis a1 a2 diff_zero is_projection_on_iff_orthog that vimage_singleton_eq) ultimately show ?thesis by blast qed \ \Theorem 2.7 in \<^cite>\conway2013course\\ proposition projection_kernel_is_orthogonal_complement: fixes M :: \'a::chilbert_space set\ assumes "closed_csubspace M" shows "(projection M) -` {0} = (orthogonal_complement M)" by (metis assms closed_csubspace_def complex_vector.subspace_def csubspace_is_convex insert_absorb insert_not_empty is_projection_on_kernel_is_orthogonal_complement projection_is_projection_on) lemma is_projection_on_id_minus: fixes M :: \'a::complex_inner set\ assumes is_proj: "is_projection_on \ M" and cc: "closed_csubspace M" shows "is_projection_on (id - \) (orthogonal_complement M)" using is_proj apply (simp add: cc is_projection_on_iff_orthog) using double_orthogonal_complement_increasing by blast text \Exercise 2 (section 2, chapter I) in \<^cite>\conway2013course\\ lemma projection_on_orthogonal_complement[simp]: fixes M :: "'a::chilbert_space set" assumes a1: "closed_csubspace M" shows "projection (orthogonal_complement M) = id - projection M" apply (auto intro!: ext) by (smt (verit, ccfv_SIG) add_diff_cancel_left' assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex diff_add_cancel double_orthogonal_complement_increasing insert_absorb insert_not_empty is_projection_on_iff_orthog orthogonal_complement_closed_subspace projection_eqI projection_is_projection_on subset_eq) lemma is_projection_on_zero: "is_projection_on (\_. 0) {0}" by (simp add: is_projection_on_def is_arg_min_def) lemma projection_zero[simp]: "projection {0} = (\_. 0)" using is_projection_on_zero by (metis (full_types) is_projection_on_in_image projection_def singletonD someI_ex) lemma is_projection_on_rank1: fixes t :: \'a::complex_inner\ shows \is_projection_on (\x. ((t \\<^sub>C x) / (t \\<^sub>C t)) *\<^sub>C t) (cspan {t})\ proof (cases \t = 0\) case True then show ?thesis by (simp add: is_projection_on_zero) next case False define P where \P x = ((t \\<^sub>C x) / (t \\<^sub>C t)) *\<^sub>C t\ for x define t' where \t' = t /\<^sub>C norm t\ with False have \norm t' = 1\ by (simp add: norm_inverse) have P_def': \P x = cinner t' x *\<^sub>C t'\ for x unfolding P_def t'_def apply auto by (metis divide_divide_eq_left divide_inverse mult.commute power2_eq_square power2_norm_eq_cinner) have spant': \cspan {t} = cspan {t'}\ by (simp add: False t'_def) have cc: \closed_csubspace (cspan {t})\ by (auto intro!: finite_cspan_closed closed_csubspace.intro) have ortho: \h - P h \ orthogonal_complement (cspan {t})\ for h unfolding orthogonal_complement_def P_def' spant' apply auto by (smt (verit, ccfv_threshold) \norm t' = 1\ add_cancel_right_left cinner_add_right cinner_commute' cinner_scaleC_right cnorm_eq_1 complex_vector.span_breakdown_eq complex_vector.span_empty diff_add_cancel mult_cancel_left1 singletonD) have inspan: \P h \ cspan {t}\ for h unfolding P_def' spant' by (simp add: complex_vector.span_base complex_vector.span_scale) show \is_projection_on P (cspan {t})\ apply (subst is_projection_on_iff_orthog) using cc ortho inspan by auto qed lemma projection_rank1: fixes t x :: \'a::complex_inner\ shows \projection (cspan {t}) x = ((t \\<^sub>C x) / (t \\<^sub>C t)) *\<^sub>C t\ apply (rule fun_cong, rule projection_eqI', simp) by (rule is_projection_on_rank1) subsection \More orthogonal complement\ text \The following lemmas logically fit into the "orthogonality" section but depend on projections for their proofs.\ text \Corollary 2.8 in \<^cite>\conway2013course\\ theorem double_orthogonal_complement_id[simp]: fixes M :: \'a::chilbert_space set\ assumes a1: "closed_csubspace M" shows "orthogonal_complement (orthogonal_complement M) = M" proof- have b2: "x \ (id - projection M) -` {0}" if c1: "x \ M" for x by (simp add: assms projection_fixes_image that) have b3: \x \ M\ if c1: \x \ (id - projection M) -` {0}\ for x by (metis assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex eq_id_iff equals0D fun_diff_def projection_in_image right_minus_eq that vimage_singleton_eq) have \x \ M \ x \ (id - projection M) -` {0}\ for x using b2 b3 by blast hence b4: \( id - (projection M) ) -` {0} = M\ by blast have b1: "orthogonal_complement (orthogonal_complement M) = (projection (orthogonal_complement M)) -` {0}" by (simp add: a1 projection_kernel_is_orthogonal_complement del: projection_on_orthogonal_complement) also have \... = ( id - (projection M) ) -` {0}\ by (simp add: a1) also have \... = M\ by (simp add: b4) finally show ?thesis by blast qed lemma orthogonal_complement_antimono_iff[simp]: fixes A B :: \('a::chilbert_space) set\ assumes \closed_csubspace A\ and \closed_csubspace B\ shows \orthogonal_complement A \ orthogonal_complement B \ A \ B\ proof (rule iffI) show \orthogonal_complement A \ orthogonal_complement B\ if \A \ B\ using that by auto assume \orthogonal_complement A \ orthogonal_complement B\ then have \orthogonal_complement (orthogonal_complement A) \ orthogonal_complement (orthogonal_complement B)\ by simp then show \A \ B\ using assms by auto qed lemma de_morgan_orthogonal_complement_plus: fixes A B::"('a::complex_inner) set" assumes \0 \ A\ and \0 \ B\ shows \orthogonal_complement (A +\<^sub>M B) = orthogonal_complement A \ orthogonal_complement B\ proof - have "x \ (orthogonal_complement A) \ (orthogonal_complement B)" if "x \ orthogonal_complement (A +\<^sub>M B)" for x proof - have \orthogonal_complement (A +\<^sub>M B) = orthogonal_complement (A + B)\ unfolding closed_sum_def by (subst orthogonal_complement_of_closure[symmetric], simp) hence \x \ orthogonal_complement (A + B)\ using that by blast hence t1: \\z \ (A + B). (z \\<^sub>C x) = 0\ by (simp add: orthogonal_complement_orthoI') have \A \ A + B\ using subset_iff add.commute set_zero_plus2 \0 \ B\ by fastforce hence \\z \ A. (z \\<^sub>C x) = 0\ using t1 by auto hence w1: \x \ (orthogonal_complement A)\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) have \B \ A + B\ using \0 \ A\ subset_iff set_zero_plus2 by blast hence \\ z \ B. (z \\<^sub>C x) = 0\ using t1 by auto hence \x \ (orthogonal_complement B)\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) thus ?thesis using w1 by auto qed moreover have "x \ (orthogonal_complement (A +\<^sub>M B))" if v1: "x \ (orthogonal_complement A) \ (orthogonal_complement B)" for x proof- have \x \ (orthogonal_complement A)\ using v1 by blast hence \\y\ A. (y \\<^sub>C x) = 0\ by (simp add: orthogonal_complement_orthoI') have \x \ (orthogonal_complement B)\ using v1 by blast hence \\ y\ B. (y \\<^sub>C x) = 0\ by (simp add: orthogonal_complement_orthoI') have \\ a\A. \ b\B. (a+b) \\<^sub>C x = 0\ by (simp add: \\y\A. y \\<^sub>C x = 0\ \\y\B. (y \\<^sub>C x) = 0\ cinner_add_left) hence \\ y \ (A + B). y \\<^sub>C x = 0\ using set_plus_elim by force hence \x \ (orthogonal_complement (A + B))\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) moreover have \(orthogonal_complement (A + B)) = (orthogonal_complement (A +\<^sub>M B))\ unfolding closed_sum_def by (subst orthogonal_complement_of_closure[symmetric], simp) ultimately have \x \ (orthogonal_complement (A +\<^sub>M B))\ by blast thus ?thesis by blast qed ultimately show ?thesis by blast qed lemma de_morgan_orthogonal_complement_inter: fixes A B::"'a::chilbert_space set" assumes a1: \closed_csubspace A\ and a2: \closed_csubspace B\ shows \orthogonal_complement (A \ B) = orthogonal_complement A +\<^sub>M orthogonal_complement B\ proof- have \orthogonal_complement A +\<^sub>M orthogonal_complement B = orthogonal_complement (orthogonal_complement (orthogonal_complement A +\<^sub>M orthogonal_complement B))\ by (simp add: closed_subspace_closed_sum) also have \\ = orthogonal_complement (orthogonal_complement (orthogonal_complement A) \ orthogonal_complement (orthogonal_complement B))\ by (simp add: de_morgan_orthogonal_complement_plus orthogonal_complementI) also have \\ = orthogonal_complement (A \ B)\ by (simp add: a1 a2) finally show ?thesis by simp qed lemma orthogonal_complement_of_cspan: \orthogonal_complement A = orthogonal_complement (cspan A)\ by (metis (no_types, opaque_lifting) closed_csubspace.subspace complex_vector.span_minimal complex_vector.span_superset double_orthogonal_complement_increasing orthogonal_complement_antimono orthogonal_complement_closed_subspace subset_antisym) lemma orthogonal_complement_orthogonal_complement_closure_cspan: \orthogonal_complement (orthogonal_complement S) = closure (cspan S)\ for S :: \'a::chilbert_space set\ proof - have \orthogonal_complement (orthogonal_complement S) = orthogonal_complement (orthogonal_complement (closure (cspan S)))\ by (simp flip: orthogonal_complement_of_closure orthogonal_complement_of_cspan) also have \\ = closure (cspan S)\ by simp finally show \orthogonal_complement (orthogonal_complement S) = closure (cspan S)\ by - qed instance ccsubspace :: (chilbert_space) complete_orthomodular_lattice proof fix X Y :: \'a ccsubspace\ show "inf X (- X) = bot" apply transfer by (simp add: closed_csubspace_def complex_vector.subspace_0 orthogonal_complement_zero_intersection) have \t \ M +\<^sub>M orthogonal_complement M\ if \closed_csubspace M\ for t::'a and M by (metis (no_types, lifting) UNIV_I closed_csubspace.subspace complex_vector.subspace_def de_morgan_orthogonal_complement_inter double_orthogonal_complement_id orthogonal_complement_closed_subspace orthogonal_complement_zero orthogonal_complement_zero_intersection that) hence b1: \M +\<^sub>M orthogonal_complement M = UNIV\ if \closed_csubspace M\ for M :: \'a set\ using that by blast show "sup X (- X) = top" apply transfer using b1 by auto show "- (- X) = X" apply transfer by simp show "- Y \ - X" if "X \ Y" using that apply transfer by simp have c1: "M +\<^sub>M orthogonal_complement M \ N \ N" if "closed_csubspace M" and "closed_csubspace N" and "M \ N" for M N :: "'a set" using that by (simp add: closed_sum_is_sup) have c2: \u \ M +\<^sub>M (orthogonal_complement M \ N)\ if a1: "closed_csubspace M" and a2: "closed_csubspace N" and a3: "M \ N" and x1: \u \ N\ for M :: "'a set" and N :: "'a set" and u proof - have d4: \(projection M) u \ M\ by (metis a1 closed_csubspace_def csubspace_is_convex equals0D orthog_proj_exists projection_in_image) hence d2: \(projection M) u \ N\ using a3 by auto have d1: \csubspace N\ by (simp add: a2) have \u - (projection M) u \ orthogonal_complement M\ by (simp add: a1 orthogonal_complementI projection_orthogonal) moreover have \u - (projection M) u \ N\ by (simp add: d1 d2 complex_vector.subspace_diff x1) ultimately have d3: \u - (projection M) u \ ((orthogonal_complement M) \ N)\ by simp hence \\ v \ ((orthogonal_complement M) \ N). u = (projection M) u + v\ by (metis d3 diff_add_cancel ordered_field_class.sign_simps(2)) then obtain v where \v \ ((orthogonal_complement M) \ N)\ and \u = (projection M) u + v\ by blast hence \u \ M + ((orthogonal_complement M) \ N)\ by (metis d4 set_plus_intro) thus ?thesis unfolding closed_sum_def using closure_subset by blast qed have c3: "N \ M +\<^sub>M ((orthogonal_complement M) \ N)" if "closed_csubspace M" and "closed_csubspace N" and "M \ N" for M N :: "'a set" using c2 that by auto show "sup X (inf (- X) Y) = Y" if "X \ Y" using that apply transfer using c1 c3 by (simp add: subset_antisym) show "X - Y = inf X (- Y)" apply transfer by simp qed subsection \Orthogonal spaces\ definition \orthogonal_spaces S T \ (\x\space_as_set S. \y\space_as_set T. is_orthogonal x y)\ lemma orthogonal_spaces_leq_compl: \orthogonal_spaces S T \ S \ -T\ unfolding orthogonal_spaces_def apply transfer by (auto simp: orthogonal_complement_def) lemma orthogonal_bot[simp]: \orthogonal_spaces S bot\ by (simp add: orthogonal_spaces_def) lemma orthogonal_spaces_sym: \orthogonal_spaces S T \ orthogonal_spaces T S\ unfolding orthogonal_spaces_def using is_orthogonal_sym by blast lemma orthogonal_sup: \orthogonal_spaces S T1 \ orthogonal_spaces S T2 \ orthogonal_spaces S (sup T1 T2)\ apply (rule orthogonal_spaces_sym) apply (simp add: orthogonal_spaces_leq_compl) using orthogonal_spaces_leq_compl orthogonal_spaces_sym by blast lemma orthogonal_sum: assumes \finite F\ and \\x. x\F \ orthogonal_spaces S (T x)\ shows \orthogonal_spaces S (sum T F)\ using assms apply induction by (auto intro!: orthogonal_sup) lemma orthogonal_spaces_ccspan: \(\x\S. \y\T. is_orthogonal x y) \ orthogonal_spaces (ccspan S) (ccspan T)\ by (meson ccspan_leq_ortho_ccspan ccspan_superset orthogonal_spaces_def orthogonal_spaces_leq_compl subset_iff) subsection \Orthonormal bases\ lemma ortho_basis_exists: fixes S :: \'a::chilbert_space set\ assumes \is_ortho_set S\ shows \\B. B \ S \ is_ortho_set B \ closure (cspan B) = UNIV\ proof - define on where \on B \ B \ S \ is_ortho_set B\ for B :: \'a set\ have \\B\Collect on. \B'\Collect on. B \ B' \ B' = B\ proof (rule subset_Zorn_nonempty; simp) show \\S. on S\ apply (rule exI[of _ S]) using assms on_def by fastforce next fix C :: \'a set set\ assume \C \ {}\ assume \subset.chain (Collect on) C\ then have C_on: \B \ C \ on B\ and C_order: \B \ C \ B' \ C \ B \ B' \ B' \ B\ for B B' by (auto simp: subset.chain_def) have \is_orthogonal x y\ if \x\\C\ \y\\C\ \x \ y\ for x y by (smt (verit) UnionE C_order C_on on_def is_ortho_set_def subsetD that(1) that(2) that(3)) moreover have \0 \ \ C\ by (meson UnionE C_on is_ortho_set_def on_def) moreover have \\C \ S\ using C_on \C \ {}\ on_def by blast ultimately show \on (\ C)\ unfolding on_def is_ortho_set_def by simp qed then obtain B where \on B\ and B_max: \B' \ B \ on B' \ B=B'\ for B' by auto have \\ = 0\ if \ortho: \\b\B. is_orthogonal \ b\ for \ :: 'a proof (rule ccontr) assume \\ \ 0\ define \ B' where \\ = \ /\<^sub>R norm \\ and \B' = B \ {\}\ have [simp]: \norm \ = 1\ using \\ \ 0\ by (auto simp: \_def) have \ortho: \is_orthogonal \ b\ if \b \ B\ for b using \ortho that \_def by auto have orthoB': \is_orthogonal x y\ if \x\B'\ \y\B'\ \x \ y\ for x y using that \on B\ \ortho \ortho[THEN is_orthogonal_sym[THEN iffD1]] by (auto simp: B'_def on_def is_ortho_set_def) have B'0: \0 \ B'\ using B'_def \norm \ = 1\ \on B\ is_ortho_set_def on_def by fastforce have \S \ B'\ using B'_def \on B\ on_def by auto from orthoB' B'0 \S \ B'\ have \on B'\ by (simp add: on_def is_ortho_set_def) with B_max have \B = B'\ by (metis B'_def Un_upper1) then have \\ \ B\ using B'_def by blast then have \is_orthogonal \ \\ using \ortho by blast then show False using B'0 \B = B'\ \\ \ B\ by fastforce qed then have \orthogonal_complement B = {0}\ by (auto simp: orthogonal_complement_def) then have \UNIV = orthogonal_complement (orthogonal_complement B)\ by simp also have \\ = orthogonal_complement (orthogonal_complement (closure (cspan B)))\ by (metis (mono_tags, opaque_lifting) \orthogonal_complement B = {0}\ cinner_zero_left complex_vector.span_superset empty_iff insert_iff orthogonal_complementI orthogonal_complement_antimono orthogonal_complement_of_closure subsetI subset_antisym) also have \\ = closure (cspan B)\ apply (rule double_orthogonal_complement_id) by simp finally have \closure (cspan B) = UNIV\ by simp with \on B\ show ?thesis by (auto simp: on_def) qed lemma orthonormal_basis_exists: fixes S :: \'a::chilbert_space set\ assumes \is_ortho_set S\ and \\x. x\S \ norm x = 1\ shows \\B. B \ S \ is_onb B\ proof - from \is_ortho_set S\ obtain B where \is_ortho_set B\ and \B \ S\ and \closure (cspan B) = UNIV\ using ortho_basis_exists by blast define B' where \B' = (\x. x /\<^sub>R norm x) ` B\ have \S = (\x. x /\<^sub>R norm x) ` S\ by (simp add: assms(2)) then have \B' \ S\ using B'_def \S \ B\ by blast moreover have \ccspan B' = top\ apply (transfer fixing: B') apply (simp add: B'_def scaleR_scaleC) apply (subst complex_vector.span_image_scale') using \is_ortho_set B\ \closure (cspan B) = UNIV\ is_ortho_set_def by auto moreover have \is_ortho_set B'\ using \is_ortho_set B\ by (auto simp: B'_def is_ortho_set_def) moreover have \\b\B'. norm b = 1\ using \is_ortho_set B\ apply (auto simp: B'_def is_ortho_set_def) by (metis field_class.field_inverse norm_eq_zero) ultimately show ?thesis by (auto simp: is_onb_def) qed definition some_chilbert_basis :: \'a::chilbert_space set\ where \some_chilbert_basis = (SOME B::'a set. is_onb B)\ lemma is_onb_some_chilbert_basis[simp]: \is_onb (some_chilbert_basis :: 'a::chilbert_space set)\ using orthonormal_basis_exists[OF is_ortho_set_empty] by (auto simp add: some_chilbert_basis_def intro: someI2) lemma is_ortho_set_some_chilbert_basis[simp]: \is_ortho_set some_chilbert_basis\ using is_onb_def is_onb_some_chilbert_basis by blast lemma is_normal_some_chilbert_basis: \\x. x \ some_chilbert_basis \ norm x = 1\ using is_onb_def is_onb_some_chilbert_basis by blast lemma ccspan_some_chilbert_basis[simp]: \ccspan some_chilbert_basis = top\ using is_onb_def is_onb_some_chilbert_basis by blast lemma span_some_chilbert_basis[simp]: \closure (cspan some_chilbert_basis) = UNIV\ by (metis ccspan.rep_eq ccspan_some_chilbert_basis top_ccsubspace.rep_eq) lemma cindependent_some_chilbert_basis[simp]: \cindependent some_chilbert_basis\ using is_ortho_set_cindependent is_ortho_set_some_chilbert_basis by blast lemma finite_some_chilbert_basis[simp]: \finite (some_chilbert_basis :: 'a :: {chilbert_space, cfinite_dim} set)\ apply (rule cindependent_cfinite_dim_finite) by simp lemma some_chilbert_basis_nonempty: \(some_chilbert_basis :: 'a::{chilbert_space, not_singleton} set) \ {}\ proof (rule ccontr, simp) define B :: \'a set\ where \B = some_chilbert_basis\ assume [simp]: \B = {}\ have \UNIV = closure (cspan B)\ using B_def span_some_chilbert_basis by blast also have \\ = {0}\ by simp also have \\ \ UNIV\ using Extra_General.UNIV_not_singleton by blast finally show False by simp qed lemma basis_projections_reconstruct_has_sum: assumes \is_ortho_set B\ and normB: \\b. b\B \ norm b = 1\ and \B: \\ \ space_as_set (ccspan B)\ shows \((\b. (b \\<^sub>C \) *\<^sub>C b) has_sum \) B\ proof (rule has_sumI_metric) fix e :: real assume \e > 0\ define e2 where \e2 = e/2\ have [simp]: \e2 > 0\ by (simp add: \0 < e\ e2_def) define bb where \bb \ b = (b \\<^sub>C \) *\<^sub>C b\ for \ and b :: 'a have linear_bb: \clinear (\\. bb \ b)\ for b by (simp add: bb_def cinner_add_right clinear_iff scaleC_left.add) from \B obtain \ where dist\\: \dist \ \ < e2\ and \B: \\ \ cspan B\ apply atomize_elim apply (simp add: ccspan.rep_eq closure_approachable) using \0 < e2\ by blast from \B obtain F where \finite F\ and \F \ B\ and \F: \\ \ cspan F\ by (meson vector_finitely_spanned) have \dist (sum (bb \) G) \ < e\ if \finite G\ and \F \ G\ and \G \ B\ for G proof - have sum\: \sum (bb \) G = \\ proof - from \F \F \ G\ have \G: \\ \ cspan G\ using complex_vector.span_mono by blast then obtain f where \sum: \\ = (\b\G. f b *\<^sub>C b)\ using complex_vector.span_finite[OF \finite G\] by auto have \sum (bb \) G = (\c\G. \b\G. bb (f b *\<^sub>C b) c)\ apply (simp add: \sum) apply (rule sum.cong, simp) apply (rule complex_vector.linear_sum[where f=\\x. bb x _\]) by (rule linear_bb) also have \\ = (\(c,b)\G\G. bb (f b *\<^sub>C b) c)\ by (simp add: sum.cartesian_product) also have \\ = (\b\G. f b *\<^sub>C b)\ apply (rule sym) apply (rule sum.reindex_bij_witness_not_neutral [where j=\\b. (b,b)\ and i=fst and T'=\G\G - (\b. (b,b)) ` G\ and S'=\{}\]) using \finite G\ apply (auto simp: bb_def) apply (metis (no_types, lifting) assms(1) imageI is_ortho_set_antimono is_ortho_set_def that(3)) using normB \G \ B\ cnorm_eq_1 by blast also have \\ = \\ by (simp flip: \sum) finally show ?thesis by - qed have \dist (sum (bb \) G) (sum (bb \) G) < e2\ proof - define \ where \\ = \ - \\ have \(dist (sum (bb \) G) (sum (bb \) G))\<^sup>2 = (norm (sum (bb \) G))\<^sup>2\ by (simp add: dist_norm \_def complex_vector.linear_diff[OF linear_bb] sum_subtractf) also have \\ = (norm (sum (bb \) G))\<^sup>2 + (norm (\ - sum (bb \) G))\<^sup>2 - (norm (\ - sum (bb \) G))\<^sup>2\ by simp also have \\ = (norm (sum (bb \) G + (\ - sum (bb \) G)))\<^sup>2 - (norm (\ - sum (bb \) G))\<^sup>2\ proof - have \(\b\G. bb \ b \\<^sub>C bb \ c) = bb \ c \\<^sub>C \\ if \c \ G\ for c apply (subst sum_single[where i=c]) using that apply (auto intro!: \finite G\ simp: bb_def) apply (metis \G \ B\ \is_ortho_set B\ is_ortho_set_antimono is_ortho_set_def) using \G \ B\ normB cnorm_eq_1 by blast then have \is_orthogonal (sum (bb \) G) (\ - sum (bb \) G)\ by (simp add: cinner_sum_left cinner_diff_right cinner_sum_right) then show ?thesis apply (rule_tac arg_cong[where f=\\x. x - _\]) by (rule pythagorean_theorem[symmetric]) qed also have \\ = (norm \)\<^sup>2 - (norm (\ - sum (bb \) G))\<^sup>2\ by simp also have \\ \ (norm \)\<^sup>2\ by simp also have \\ = (dist \ \)\<^sup>2\ by (simp add: \_def dist_norm) also have \\ < e2\<^sup>2\ using dist\\ apply (rule power_strict_mono) by auto finally show ?thesis by (smt (verit) \0 < e2\ power_mono) qed with sum\ dist\\ show \dist (sum (bb \) G) \ < e\ apply (rule_tac dist_triangle_lt[where z=\]) by (simp add: e2_def dist_commute) qed with \finite F\ and \F \ B\ show \\F. finite F \ F \ B \ (\G. finite G \ F \ G \ G \ B \ dist (sum (bb \) G) \ < e)\ by (auto intro!: exI[of _ F]) qed lemma basis_projections_reconstruct: assumes \is_ortho_set B\ and \\b. b\B \ norm b = 1\ and \\ \ space_as_set (ccspan B)\ shows \(\\<^sub>\b\B. (b \\<^sub>C \) *\<^sub>C b) = \\ using assms basis_projections_reconstruct_has_sum infsumI by blast lemma basis_projections_reconstruct_summable: assumes \is_ortho_set B\ and \\b. b\B \ norm b = 1\ and \\ \ space_as_set (ccspan B)\ shows \(\b. (b \\<^sub>C \) *\<^sub>C b) summable_on B\ by (simp add: assms basis_projections_reconstruct basis_projections_reconstruct_has_sum summable_iff_has_sum_infsum) lemma parseval_identity_has_sum: assumes \is_ortho_set B\ and normB: \\b. b\B \ norm b = 1\ and \\ \ space_as_set (ccspan B)\ shows \((\b. (norm (b \\<^sub>C \))\<^sup>2) has_sum (norm \)\<^sup>2) B\ proof - have *: \(\v. (norm v)\<^sup>2) (\b\F. (b \\<^sub>C \) *\<^sub>C b) = (\b\F. (norm (b \\<^sub>C \))\<^sup>2)\ if \finite F\ and \F \ B\ for F apply (subst pythagorean_theorem_sum) using \is_ortho_set B\ normB that apply (auto intro!: sum.cong[OF refl] simp: is_ortho_set_def) by blast from assms have \((\b. (b \\<^sub>C \) *\<^sub>C b) has_sum \) B\ by (simp add: basis_projections_reconstruct_has_sum) then have \((\F. \b\F. (b \\<^sub>C \) *\<^sub>C b) \ \) (finite_subsets_at_top B)\ by (simp add: has_sum_def) then have \((\F. (\v. (norm v)\<^sup>2) (\b\F. (b \\<^sub>C \) *\<^sub>C b)) \ (norm \)\<^sup>2) (finite_subsets_at_top B)\ apply (rule isCont_tendsto_compose[rotated]) by simp then have \((\F. (\b\F. (norm (b \\<^sub>C \))\<^sup>2)) \ (norm \)\<^sup>2) (finite_subsets_at_top B)\ apply (rule tendsto_cong[THEN iffD2, rotated]) apply (rule eventually_finite_subsets_at_top_weakI) by (simp add: *) then show \((\b. (norm (b \\<^sub>C \))\<^sup>2) has_sum (norm \)\<^sup>2) B\ by (simp add: has_sum_def) qed lemma parseval_identity_summable: assumes \is_ortho_set B\ and \\b. b\B \ norm b = 1\ and \\ \ space_as_set (ccspan B)\ shows \(\b. (norm (b \\<^sub>C \))\<^sup>2) summable_on B\ using parseval_identity_has_sum[OF assms] has_sum_imp_summable by blast lemma parseval_identity: assumes \is_ortho_set B\ and \\b. b\B \ norm b = 1\ and \\ \ space_as_set (ccspan B)\ shows \(\\<^sub>\b\B. (norm (b \\<^sub>C \))\<^sup>2) = (norm \)\<^sup>2\ using parseval_identity_has_sum[OF assms] using infsumI by blast subsection \Riesz-representation theorem\ lemma orthogonal_complement_kernel_functional: fixes f :: \'a::complex_inner \ complex\ assumes \bounded_clinear f\ shows \\x. orthogonal_complement (f -` {0}) = cspan {x}\ proof (cases \orthogonal_complement (f -` {0}) = {0}\) case True then show ?thesis apply (rule_tac x=0 in exI) by auto next case False then obtain x where xortho: \x \ orthogonal_complement (f -` {0})\ and xnon0: \x \ 0\ using complex_vector.subspace_def by fastforce from xnon0 xortho have r1: \f x \ 0\ by (metis cinner_eq_zero_iff orthogonal_complement_orthoI vimage_singleton_eq) have \\ k. y = k *\<^sub>C x\ if \y \ orthogonal_complement (f -` {0})\ for y proof (cases \y = 0\) case True then show ?thesis by auto next case False with that have \f y \ 0\ by (metis cinner_eq_zero_iff orthogonal_complement_orthoI vimage_singleton_eq) then obtain k where k_def: \f x = k * f y\ by (metis add.inverse_inverse minus_divide_eq_eq) with assms have \f x = f (k *\<^sub>C y)\ by (simp add: bounded_clinear.axioms(1) clinear.scaleC) hence \f x - f (k *\<^sub>C y) = 0\ by simp with assms have s1: \f (x - k *\<^sub>C y) = 0\ by (simp add: bounded_clinear.axioms(1) complex_vector.linear_diff) from that have \k *\<^sub>C y \ orthogonal_complement (f -` {0})\ by (simp add: complex_vector.subspace_scale) with xortho have s2: \x - (k *\<^sub>C y) \ orthogonal_complement (f -` {0})\ by (simp add: complex_vector.subspace_diff) have s3: \(x - (k *\<^sub>C y)) \ f -` {0}\ using s1 by simp moreover have \(f -` {0}) \ (orthogonal_complement (f -` {0})) = {0}\ by (meson assms closed_csubspace_def complex_vector.subspace_def kernel_is_closed_csubspace orthogonal_complement_zero_intersection) ultimately have \x - (k *\<^sub>C y) = 0\ using s2 by blast thus ?thesis by (metis ceq_vector_fraction_iff eq_iff_diff_eq_0 k_def r1 scaleC_scaleC) qed then have \orthogonal_complement (f -` {0}) \ cspan {x}\ using complex_vector.span_superset complex_vector.subspace_scale by blast moreover from xortho have \orthogonal_complement (f -` {0}) \ cspan {x}\ by (simp add: complex_vector.span_minimal) ultimately show ?thesis by auto qed -lemma riesz_frechet_representation_existence: +lemma riesz_representation_existence: \ \Theorem 3.4 in \<^cite>\conway2013course\\ fixes f::\'a::chilbert_space \ complex\ assumes a1: \bounded_clinear f\ shows \\t. \x. f x = t \\<^sub>C x\ proof(cases \\ x. f x = 0\) case True thus ?thesis by (metis cinner_zero_left) next case False obtain t where spant: \orthogonal_complement (f -` {0}) = cspan {t}\ using orthogonal_complement_kernel_functional using assms by blast have \projection (orthogonal_complement (f -` {0})) x = ((t \\<^sub>C x)/(t \\<^sub>C t)) *\<^sub>C t\ for x apply (subst spant) by (rule projection_rank1) hence \f (projection (orthogonal_complement (f -` {0})) x) = (((t \\<^sub>C x))/(t \\<^sub>C t)) * (f t)\ for x using a1 unfolding bounded_clinear_def by (simp add: complex_vector.linear_scale) hence l2: \f (projection (orthogonal_complement (f -` {0})) x) = ((cnj (f t)/(t \\<^sub>C t)) *\<^sub>C t) \\<^sub>C x\ for x using complex_cnj_divide by force have \f (projection (f -` {0}) x) = 0\ for x by (metis (no_types, lifting) assms bounded_clinear_def closed_csubspace.closed complex_vector.linear_subspace_vimage complex_vector.subspace_0 complex_vector.subspace_single_0 csubspace_is_convex insert_absorb insert_not_empty kernel_is_closed_csubspace projection_in_image vimage_singleton_eq) hence "\a b. f (projection (f -` {0}) a + b) = 0 + f b" using additive.add assms by (simp add: bounded_clinear_def complex_vector.linear_add) hence "\a. 0 + f (projection (orthogonal_complement (f -` {0})) a) = f a" apply (simp add: assms) by (metis add.commute diff_add_cancel) hence \f x = ((cnj (f t)/(t \\<^sub>C t)) *\<^sub>C t) \\<^sub>C x\ for x by (simp add: l2) thus ?thesis by blast qed -lemma riesz_frechet_representation_unique: +lemma riesz_representation_unique: \ \Theorem 3.4 in \<^cite>\conway2013course\\ fixes f::\'a::complex_inner \ complex\ assumes \\x. f x = (t \\<^sub>C x)\ assumes \\x. f x = (u \\<^sub>C x)\ shows \t = u\ by (metis add_diff_cancel_left' assms(1) assms(2) cinner_diff_left cinner_gt_zero_iff diff_add_cancel diff_zero) subsection \Adjoints\ definition "is_cadjoint F G \ (\x. \y. (F x \\<^sub>C y) = (x \\<^sub>C G y))" lemma is_adjoint_sym: \is_cadjoint F G \ is_cadjoint G F\ unfolding is_cadjoint_def apply auto by (metis cinner_commute') definition \cadjoint G = (SOME F. is_cadjoint F G)\ for G :: "'b::complex_inner \ 'a::complex_inner" lemma cadjoint_exists: fixes G :: "'b::chilbert_space \ 'a::complex_inner" assumes [simp]: \bounded_clinear G\ shows \\F. is_cadjoint F G\ proof - include notation_norm have [simp]: \clinear G\ using assms unfolding bounded_clinear_def by blast define g :: \'a \ 'b \ complex\ where \g x y = (x \\<^sub>C G y)\ for x y have \bounded_clinear (g x)\ for x proof - have \g x (a + b) = g x a + g x b\ for a b unfolding g_def using additive.add cinner_add_right clinear_def by (simp add: cinner_add_right complex_vector.linear_add) moreover have \g x (k *\<^sub>C a) = k *\<^sub>C (g x a)\ for a k unfolding g_def by (simp add: complex_vector.linear_scale) ultimately have \clinear (g x)\ by (simp add: clinearI) moreover have \\ M. \ y. \ G y \ \ \ y \ * M\ using \bounded_clinear G\ unfolding bounded_clinear_def bounded_clinear_axioms_def by blast then have \\M. \y. \ g x y \ \ \ y \ * M\ using g_def by (simp add: bounded_clinear.bounded bounded_clinear_cinner_right_comp) ultimately show ?thesis unfolding bounded_linear_def using bounded_clinear.intro using bounded_clinear_axioms_def by blast qed hence \\x. \t. \y. g x y = (t \\<^sub>C y)\ - using riesz_frechet_representation_existence by blast + using riesz_representation_existence by blast then obtain F where \\x. \y. g x y = (F x \\<^sub>C y)\ by metis then have \is_cadjoint F G\ unfolding is_cadjoint_def g_def by simp thus ?thesis by auto qed lemma cadjoint_is_cadjoint[simp]: fixes G :: "'b::chilbert_space \ 'a::complex_inner" assumes [simp]: \bounded_clinear G\ shows \is_cadjoint (cadjoint G) G\ by (metis assms cadjoint_def cadjoint_exists someI_ex) lemma is_cadjoint_unique: assumes \is_cadjoint F1 G\ assumes \is_cadjoint F2 G\ shows \F1 = F2\ - by (metis (full_types) assms(1) assms(2) ext is_cadjoint_def riesz_frechet_representation_unique) + by (metis (full_types) assms(1) assms(2) ext is_cadjoint_def riesz_representation_unique) lemma cadjoint_univ_prop: fixes G :: "'b::chilbert_space \ 'a::complex_inner" assumes a1: \bounded_clinear G\ shows \cadjoint G x \\<^sub>C y = x \\<^sub>C G y\ using assms cadjoint_is_cadjoint is_cadjoint_def by blast lemma cadjoint_univ_prop': fixes G :: "'b::chilbert_space \ 'a::complex_inner" assumes a1: \bounded_clinear G\ shows \x \\<^sub>C cadjoint G y = G x \\<^sub>C y\ by (metis cadjoint_univ_prop assms cinner_commute') notation cadjoint ("_\<^sup>\" [99] 100) lemma cadjoint_eqI: fixes G:: \'b::complex_inner \ 'a::complex_inner\ and F:: \'a \ 'b\ assumes \\x y. (F x \\<^sub>C y) = (x \\<^sub>C G y)\ shows \G\<^sup>\ = F\ by (metis assms cadjoint_def is_cadjoint_def is_cadjoint_unique someI_ex) lemma cadjoint_bounded_clinear: fixes A :: "'a::chilbert_space \ 'b::complex_inner" assumes a1: "bounded_clinear A" shows \bounded_clinear (A\<^sup>\)\ proof include notation_norm have b1: \((A\<^sup>\) x \\<^sub>C y) = (x \\<^sub>C A y)\ for x y using cadjoint_univ_prop a1 by auto have \is_orthogonal ((A\<^sup>\) (x1 + x2) - ((A\<^sup>\) x1 + (A\<^sup>\) x2)) y\ for x1 x2 y by (simp add: b1 cinner_diff_left cinner_add_left) hence b2: \(A\<^sup>\) (x1 + x2) - ((A\<^sup>\) x1 + (A\<^sup>\) x2) = 0\ for x1 x2 using cinner_eq_zero_iff by blast thus z1: \(A\<^sup>\) (x1 + x2) = (A\<^sup>\) x1 + (A\<^sup>\) x2\ for x1 x2 by (simp add: b2 eq_iff_diff_eq_0) have f1: \is_orthogonal ((A\<^sup>\) (r *\<^sub>C x) - (r *\<^sub>C (A\<^sup>\) x )) y\ for r x y by (simp add: b1 cinner_diff_left) thus z2: \(A\<^sup>\) (r *\<^sub>C x) = r *\<^sub>C (A\<^sup>\) x\ for r x using cinner_eq_zero_iff eq_iff_diff_eq_0 by blast have \\ (A\<^sup>\) x \^2 = ((A\<^sup>\) x \\<^sub>C (A\<^sup>\) x)\ for x by (metis cnorm_eq_square) moreover have \\ (A\<^sup>\) x \^2 \ 0\ for x by simp ultimately have \\ (A\<^sup>\) x \^2 = \ ((A\<^sup>\) x \\<^sub>C (A\<^sup>\) x) \\ for x by (metis abs_pos cinner_ge_zero) hence \\ (A\<^sup>\) x \^2 = \ (x \\<^sub>C A ((A\<^sup>\) x)) \\ for x by (simp add: b1) moreover have \\(x \\<^sub>C A ((A\<^sup>\) x))\ \ \x\ * \A ((A\<^sup>\) x)\\ for x by (simp add: abs_complex_def complex_inner_class.Cauchy_Schwarz_ineq2 less_eq_complex_def) ultimately have b5: \\ (A\<^sup>\) x \^2 \ \x\ * \A ((A\<^sup>\) x)\\ for x by (metis complex_of_real_mono_iff) have \\M. M \ 0 \ (\ x. \A ((A\<^sup>\) x)\ \ M * \(A\<^sup>\) x\)\ using a1 by (metis (mono_tags, opaque_lifting) bounded_clinear.bounded linear mult_nonneg_nonpos mult_zero_right norm_ge_zero order.trans semiring_normalization_rules(7)) then obtain M where q1: \M \ 0\ and q2: \\ x. \A ((A\<^sup>\) x)\ \ M * \(A\<^sup>\) x\\ by blast have \\ x::'b. \x\ \ 0\ by simp hence b6: \\x\ * \A ((A\<^sup>\) x)\ \ \x\ * M * \(A\<^sup>\) x\\ for x using q2 by (smt ordered_comm_semiring_class.comm_mult_left_mono vector_space_over_itself.scale_scale) have z3: \\ (A\<^sup>\) x \ \ \x\ * M\ for x proof(cases \\(A\<^sup>\) x\ = 0\) case True thus ?thesis by (simp add: \0 \ M\) next case False have \\ (A\<^sup>\) x \^2 \ \x\ * M * \(A\<^sup>\) x\\ by (smt b5 b6) thus ?thesis by (smt False mult_right_cancel mult_right_mono norm_ge_zero semiring_normalization_rules(29)) qed thus \\K. \x. \(A\<^sup>\) x\ \ \x\ * K\ by auto qed proposition double_cadjoint: fixes U :: \'a::chilbert_space \ 'b::complex_inner\ assumes a1: "bounded_clinear U" shows "U\<^sup>\\<^sup>\ = U" by (metis assms cadjoint_def cadjoint_is_cadjoint is_adjoint_sym is_cadjoint_unique someI_ex) lemma cadjoint_id[simp]: \id\<^sup>\ = id\ by (simp add: cadjoint_eqI id_def) lemma scaleC_cadjoint: fixes A::"'a::chilbert_space \ 'b::complex_inner" assumes "bounded_clinear A" shows \(\t. a *\<^sub>C A t)\<^sup>\ = (\s. cnj a *\<^sub>C (A\<^sup>\) s)\ proof - have b3: \((\ s. (cnj a) *\<^sub>C ((A\<^sup>\) s)) x \\<^sub>C y) = (x \\<^sub>C (\ t. a *\<^sub>C (A t)) y)\ for x y by (simp add: assms cadjoint_univ_prop) have "((\t. a *\<^sub>C A t)\<^sup>\) b = cnj a *\<^sub>C (A\<^sup>\) b" for b::'b proof- have "bounded_clinear (\t. a *\<^sub>C A t)" by (simp add: assms bounded_clinear_const_scaleC) thus ?thesis by (metis (no_types) cadjoint_eqI b3) qed thus ?thesis by blast qed lemma is_projection_on_is_cadjoint: fixes M :: \'a::complex_inner set\ assumes a1: \is_projection_on \ M\ and a2: \closed_csubspace M\ shows \is_cadjoint \ \\ by (smt (verit, ccfv_threshold) a1 a2 cinner_diff_left cinner_eq_flip is_cadjoint_def is_projection_on_iff_orthog orthogonal_complement_orthoI right_minus_eq) lemma is_projection_on_cadjoint: fixes M :: \'a::complex_inner set\ assumes \is_projection_on \ M\ and \closed_csubspace M\ shows \\\<^sup>\ = \\ using assms is_projection_on_is_cadjoint cadjoint_eqI is_cadjoint_def by blast lemma projection_cadjoint: fixes M :: \'a::chilbert_space set\ assumes \closed_csubspace M\ shows \(projection M)\<^sup>\ = projection M\ using is_projection_on_cadjoint assms by (metis closed_csubspace.closed closed_csubspace.subspace csubspace_is_convex empty_iff orthog_proj_exists projection_is_projection_on) subsection \More projections\ text \These lemmas logically belong in the "projections" section above but depend on lemmas developed later.\ lemma is_projection_on_plus: assumes "\x y. x \ A \ y \ B \ is_orthogonal x y" assumes \closed_csubspace A\ assumes \closed_csubspace B\ assumes \is_projection_on \A A\ assumes \is_projection_on \B B\ shows \is_projection_on (\x. \A x + \B x) (A +\<^sub>M B)\ proof (rule is_projection_on_iff_orthog[THEN iffD2, rule_format]) show clAB: \closed_csubspace (A +\<^sub>M B)\ by (simp add: assms(2) assms(3) closed_subspace_closed_sum) fix h have 1: \\A h + \B h \ A +\<^sub>M B\ by (meson clAB assms(2) assms(3) assms(4) assms(5) closed_csubspace_def closed_sum_left_subset closed_sum_right_subset complex_vector.subspace_def in_mono is_projection_on_in_image) have \\A (\B h) = 0\ by (smt (verit, del_insts) assms(1) assms(2) assms(4) assms(5) cinner_eq_zero_iff is_cadjoint_def is_projection_on_in_image is_projection_on_is_cadjoint) then have \h - (\A h + \B h) = (h - \B h) - \A (h - \B h)\ by (smt (verit) add.right_neutral add_diff_cancel_left' assms(2) assms(4) closed_csubspace.subspace complex_vector.subspace_diff diff_add_eq_diff_diff_swap diff_diff_add is_projection_on_iff_orthog orthog_proj_unique orthogonal_complement_closed_subspace) also have \\ \ orthogonal_complement A\ using assms(2) assms(4) is_projection_on_iff_orthog by blast finally have orthoA: \h - (\A h + \B h) \ orthogonal_complement A\ by - have \\B (\A h) = 0\ by (smt (verit, del_insts) assms(1) assms(3) assms(4) assms(5) cinner_eq_zero_iff is_cadjoint_def is_projection_on_in_image is_projection_on_is_cadjoint) then have \h - (\A h + \B h) = (h - \A h) - \B (h - \A h)\ by (smt (verit) add.right_neutral add_diff_cancel assms(3) assms(5) closed_csubspace.subspace complex_vector.subspace_diff diff_add_eq_diff_diff_swap diff_diff_add is_projection_on_iff_orthog orthog_proj_unique orthogonal_complement_closed_subspace) also have \\ \ orthogonal_complement B\ using assms(3) assms(5) is_projection_on_iff_orthog by blast finally have orthoB: \h - (\A h + \B h) \ orthogonal_complement B\ by - from orthoA orthoB have 2: \h - (\A h + \B h) \ orthogonal_complement (A +\<^sub>M B)\ by (metis IntI assms(2) assms(3) closed_csubspace_def complex_vector.subspace_def de_morgan_orthogonal_complement_plus) from 1 2 show \h - (\A h + \B h) \ orthogonal_complement (A +\<^sub>M B) \ \A h + \B h \ A +\<^sub>M B\ by simp qed lemma projection_plus: fixes A B :: "'a::chilbert_space set" assumes "\x y. x:A \ y:B \ is_orthogonal x y" assumes \closed_csubspace A\ assumes \closed_csubspace B\ shows \projection (A +\<^sub>M B) = (\x. projection A x + projection B x)\ proof - have \is_projection_on (\x. projection A x + projection B x) (A +\<^sub>M B)\ apply (rule is_projection_on_plus) using assms by auto then show ?thesis by (meson assms(2) assms(3) closed_csubspace.subspace closed_subspace_closed_sum csubspace_is_convex projection_eqI') qed lemma is_projection_on_insert: assumes ortho: "\s. s \ S \ is_orthogonal a s" assumes \is_projection_on \ (closure (cspan S))\ assumes \is_projection_on \a (cspan {a})\ shows "is_projection_on (\x. \a x + \ x) (closure (cspan (insert a S)))" proof - from ortho have \x \ cspan {a} \ y \ closure (cspan S) \ is_orthogonal x y\ for x y using is_orthogonal_cspan is_orthogonal_closure is_orthogonal_sym by (smt (verit, ccfv_threshold) empty_iff insert_iff) then have \is_projection_on (\x. \a x + \ x) (cspan {a} +\<^sub>M closure (cspan S))\ apply (rule is_projection_on_plus) using assms by (auto simp add: closed_csubspace.intro) also have \\ = closure (cspan (insert a S))\ using closed_sum_cspan[where X=\{a}\] by simp finally show ?thesis by - qed lemma projection_insert: fixes a :: \'a::chilbert_space\ assumes a1: "\s. s \ S \ is_orthogonal a s" shows "projection (closure (cspan (insert a S))) u = projection (cspan {a}) u + projection (closure (cspan S)) u" using is_projection_on_insert[where S=S, OF a1] by (metis (no_types, lifting) closed_closure closed_csubspace.intro closure_is_csubspace complex_vector.subspace_span csubspace_is_convex finite.intros(1) finite.intros(2) finite_cspan_closed_csubspace projection_eqI' projection_is_projection_on') lemma projection_insert_finite: fixes S :: \'a::chilbert_space set\ assumes a1: "\s. s \ S \ is_orthogonal a s" and a2: "finite S" shows "projection (cspan (insert a S)) u = projection (cspan {a}) u + projection (cspan S) u" using projection_insert by (metis a1 a2 closure_finite_cspan finite.insertI) subsection \Canonical basis (\onb_enum\)\ setup \Sign.add_const_constraint (\<^const_name>\is_ortho_set\, SOME \<^typ>\'a set \ bool\)\ class onb_enum = basis_enum + complex_inner + assumes is_orthonormal: "is_ortho_set (set canonical_basis)" and is_normal: "\x. x \ (set canonical_basis) \ norm x = 1" setup \Sign.add_const_constraint (\<^const_name>\is_ortho_set\, SOME \<^typ>\'a::complex_inner set \ bool\)\ lemma cinner_canonical_basis: assumes \i < length (canonical_basis :: 'a::onb_enum list)\ assumes \j < length (canonical_basis :: 'a::onb_enum list)\ shows \cinner (canonical_basis!i :: 'a) (canonical_basis!j) = (if i=j then 1 else 0)\ by (metis assms(1) assms(2) distinct_canonical_basis is_normal is_ortho_set_def is_orthonormal nth_eq_iff_index_eq nth_mem of_real_1 power2_norm_eq_cinner power_one) lemma canonical_basis_is_onb[simp]: \is_onb (set canonical_basis :: 'a::onb_enum set)\ by (simp add: is_normal is_onb_def is_orthonormal) instance onb_enum \ chilbert_space proof have \complete (UNIV :: 'a set)\ using finite_cspan_complete[where B=\set canonical_basis\] by simp then show "convergent X" if "Cauchy X" for X :: "nat \ 'a" by (simp add: complete_def convergent_def that) qed subsection \Conjugate space\ instantiation conjugate_space :: (complex_inner) complex_inner begin lift_definition cinner_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ complex" is \\x y. cinner y x\. instance apply (intro_classes; transfer) apply (simp_all add: ) apply (simp add: cinner_add_right) using cinner_ge_zero norm_eq_sqrt_cinner by auto end instance conjugate_space :: (chilbert_space) chilbert_space.. subsection \Misc (ctd.)\ lemma separating_dense_span: assumes \\F G :: 'a::chilbert_space \ 'b::{complex_normed_vector,not_singleton}. bounded_clinear F \ bounded_clinear G \ (\x\S. F x = G x) \ F = G\ shows \closure (cspan S) = UNIV\ proof - have \\ = 0\ if \\ \ orthogonal_complement S\ for \ proof - obtain \ :: 'b where \\ \ 0\ by fastforce have \(\x. cinner \ x *\<^sub>C \) = (\_. 0)\ apply (rule assms[rule_format]) using orthogonal_complement_orthoI that by (auto simp add: bounded_clinear_cinner_right bounded_clinear_scaleC_const) then have \cinner \ \ = 0\ by (meson \\ \ 0\ scaleC_eq_0_iff) then show \\ = 0\ by auto qed then have \orthogonal_complement (orthogonal_complement S) = UNIV\ by (metis UNIV_eq_I cinner_zero_right orthogonal_complementI) then show \closure (cspan S) = UNIV\ by (simp add: orthogonal_complement_orthogonal_complement_closure_cspan) qed 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,1751 +1,1750 @@ 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 f = sqrt (\\<^sub>\x. norm (f x)^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 (f::'a::finite\_)" unfolding has_ell2_norm_def by simp lemma ell2_norm_finite: "ell2_norm (f::'a::finite\complex) = sqrt (\x\UNIV. (norm (f x))^2)" 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_norm_square: \(ell2_norm x)\<^sup>2 = (\\<^sub>\i. (cmod (x i))\<^sup>2)\ unfolding ell2_norm_def apply (subst real_sqrt_pow2) by (simp_all add: infsum_nonneg) lemma ell2_ket: fixes a defines \f \ (\i. of_bool (a = i))\ 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 = \{f::'a\complex. has_ell2_norm f}\ 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 \\f g. \\<^sub>\x. (cnj (f x) * g x)\ . 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 intro_classes 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 lemma sum_ell2_transfer[transfer_rule]: includes lifting_syntax shows \(((=) ===> pcr_ell2 (=)) ===> rel_set (=) ===> pcr_ell2 (=)) (\f X x. sum (\y. f y x) X) sum\ proof (intro rel_funI, rename_tac f f' X X') fix f and f' :: \'a \ 'b ell2\ assume [transfer_rule]: \((=) ===> pcr_ell2 (=)) f f'\ fix X X' :: \'a set\ assume \rel_set (=) X X'\ then have [simp]: \X' = X\ by (simp add: rel_set_eq) show \pcr_ell2 (=) (\x. \y\X. f y x) (sum f' X')\ unfolding \X' = X\ proof (induction X rule: infinite_finite_induct) case (infinite X) show ?case apply (simp add: infinite) by transfer_prover next case empty show ?case apply (simp add: empty) by transfer_prover next case (insert x F) note [transfer_rule] = insert.IH show ?case apply (simp add: insert) by transfer_prover qed qed lemma clinear_Rep_ell2[simp]: \clinear (\\. Rep_ell2 \ i)\ by (simp add: clinearI plus_ell2.rep_eq scaleC_ell2.rep_eq) lemma Abs_ell2_inverse_finite[simp]: \Rep_ell2 (Abs_ell2 \) = \\ for \ :: \_::finite \ complex\ by (simp add: Abs_ell2_inverse) 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 trunc_ell2_UNIV[simp]: \trunc_ell2 UNIV \ = \\ 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) \\<^sub>C (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 lemma trunc_ell2_norm_mono: \M \ N \ norm (trunc_ell2 M \) \ norm (trunc_ell2 N \)\ proof (rule power2_le_imp_le[rotated], force, transfer) fix M N :: \'a set\ and \ :: \'a \ complex\ assume \M \ N\ and \has_ell2_norm \\ have \(ell2_norm (\i. if i \ M then \ i else 0))\<^sup>2 = (\\<^sub>\i\M. (cmod (\ i))\<^sup>2)\ unfolding ell2_norm_square apply (rule infsum_cong_neutral) by auto also have \\ \ (\\<^sub>\i\N. (cmod (\ i))\<^sup>2)\ apply (rule infsum_mono2) using \has_ell2_norm \\ \M \ N\ by (auto simp add: ell2_norm_square has_ell2_norm_def simp flip: norm_power intro: summable_on_subset_banach) also have \\ = (ell2_norm (\i. if i \ N then \ i else 0))\<^sup>2\ unfolding ell2_norm_square apply (rule infsum_cong_neutral) by auto finally show \(ell2_norm (\i. if i \ M then \ i else 0))\<^sup>2 \ (ell2_norm (\i. if i \ N then \ i else 0))\<^sup>2\ by - qed lemma trunc_ell2_reduces_norm: \norm (trunc_ell2 M \) \ norm \\ by (metis subset_UNIV trunc_ell2_UNIV trunc_ell2_norm_mono) lemma trunc_ell2_twice[simp]: \trunc_ell2 M (trunc_ell2 N \) = trunc_ell2 (M\N) \\ apply transfer by auto lemma trunc_ell2_union: \trunc_ell2 (M \ N) \ = trunc_ell2 M \ + trunc_ell2 N \ - trunc_ell2 (M\N) \\ apply transfer by auto lemma trunc_ell2_union_disjoint: \M \ N = {} \ trunc_ell2 (M \ N) \ = trunc_ell2 M \ + trunc_ell2 N \\ by (simp add: trunc_ell2_union) lemma trunc_ell2_union_Diff: \M \ N \ trunc_ell2 (N-M) \ = trunc_ell2 N \ - trunc_ell2 M \\ using trunc_ell2_union_disjoint[where M=\N-M\ and N=M and \=\] by (simp add: Un_commute inf.commute le_iff_sup) lemma trunc_ell2_add: \trunc_ell2 M (\ + \) = trunc_ell2 M \ + trunc_ell2 M \\ apply transfer by auto lemma trunc_ell2_scaleC: \trunc_ell2 M (c *\<^sub>C \) = c *\<^sub>C trunc_ell2 M \\ apply transfer by auto lemma bounded_clinear_trunc_ell2[bounded_clinear]: \bounded_clinear (trunc_ell2 M)\ by (auto intro!: bounded_clinearI[where K=1] trunc_ell2_reduces_norm simp: trunc_ell2_add trunc_ell2_scaleC) lemma trunc_ell2_lim: \((\S. trunc_ell2 S \) \ trunc_ell2 M \) (finite_subsets_at_top M)\ proof - have \((\S. trunc_ell2 S (trunc_ell2 M \)) \ trunc_ell2 M \) (finite_subsets_at_top UNIV)\ using trunc_ell2_lim_at_UNIV by blast then have \((\S. trunc_ell2 (S\M) \) \ trunc_ell2 M \) (finite_subsets_at_top UNIV)\ by simp then show \((\S. trunc_ell2 S \) \ trunc_ell2 M \) (finite_subsets_at_top M)\ unfolding filterlim_def apply (subst (asm) filtermap_filtermap[where g=\\S. S\M\, symmetric]) apply (subst (asm) finite_subsets_at_top_inter[where A=M and B=UNIV]) by auto qed lemma trunc_ell2_lim_general: assumes big: \\G. finite G \ G \ M \ (\\<^sub>F H in F. H \ G)\ assumes small: \\\<^sub>F H in F. H \ M\ shows \((\S. trunc_ell2 S \) \ trunc_ell2 M \) F\ proof (rule tendstoI) fix e :: real assume \e > 0\ from trunc_ell2_lim[THEN tendsto_iff[THEN iffD1], rule_format, OF \e > 0\, where M=M and \=\] obtain G where \finite G\ and \G \ M\ and close: \dist (trunc_ell2 G \) (trunc_ell2 M \) < e\ apply atomize_elim unfolding eventually_finite_subsets_at_top by blast from \finite G\ \G \ M\ and big have \\\<^sub>F H in F. H \ G\ by - with small have \\\<^sub>F H in F. H \ M \ H \ G\ by (simp add: eventually_conj_iff) then show \\\<^sub>F H in F. dist (trunc_ell2 H \) (trunc_ell2 M \) < e\ proof (rule eventually_mono) fix H assume GHM: \H \ M \ H \ G\ have \dist (trunc_ell2 H \) (trunc_ell2 M \) = norm (trunc_ell2 (M-H) \)\ by (simp add: GHM dist_ell2_def norm_minus_commute trunc_ell2_union_Diff) also have \\ \ norm (trunc_ell2 (M-G) \)\ by (simp add: Diff_mono GHM trunc_ell2_norm_mono) also have \\ = dist (trunc_ell2 G \) (trunc_ell2 M \)\ by (simp add: \G \ M\ dist_ell2_def norm_minus_commute trunc_ell2_union_Diff) also have \\ < e\ using close by simp finally show \dist (trunc_ell2 H \) (trunc_ell2 M \) < e\ by - qed qed lemma norm_ell2_bound_trunc: assumes \\M. finite M \ norm (trunc_ell2 M \) \ B\ shows \norm \ \ B\ proof - note trunc_ell2_lim_at_UNIV[of \] then have \((\S. norm (trunc_ell2 S \)) \ norm \) (finite_subsets_at_top UNIV)\ using tendsto_norm by auto then show \norm \ \ B\ apply (rule tendsto_upperbound) using assms apply (rule eventually_finite_subsets_at_top_weakI) by auto qed lemma trunc_ell2_uminus: \trunc_ell2 (-M) \ = \ - trunc_ell2 M \\ by (metis Int_UNIV_left boolean_algebra_class.diff_eq subset_UNIV trunc_ell2_UNIV trunc_ell2_union_Diff) subsection \Kets and bras\ lift_definition ket :: \'a \ 'a ell2\ is \\x y. of_bool (x=y)\ 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. of_bool ((undefined::'a) = y)) \ (\_. 0)\ by (metis (mono_tags) of_bool_eq(2) zero_neq_one) qed thus \\x y::'a ell2. x \ y\ by blast qed lemma cinner_ket_left: \ket i \\<^sub>C \ = Rep_ell2 \ i\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto lemma cinner_ket_right: \(\ \\<^sub>C ket i) = cnj (Rep_ell2 \ i)\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto (* Doesn't really belong in this subsection but uses lemmas from this subsection for its proof. *) lemma bounded_clinear_Rep_ell2[simp, bounded_clinear]: \bounded_clinear (\\. Rep_ell2 \ x)\ apply (subst asm_rl[of \(\\. Rep_ell2 \ x) = (\\. ket x \\<^sub>C \)\]) apply (auto simp: cinner_ket_left) by (simp add: bounded_clinear_cinner_right) lemma cinner_ket_eqI: assumes \\i. ket i \\<^sub>C \ = ket i \\<^sub>C \\ 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 \\<^sub>C ket i) = 1\ proof- have \norm (ket i) = 1\ by simp hence \sqrt (cmod (ket i \\<^sub>C ket i)) = 1\ by (metis norm_eq_sqrt_cinner) hence \cmod (ket i \\<^sub>C ket i) = 1\ using real_sqrt_eq_1_iff by blast moreover have \(ket i \\<^sub>C ket i) = cmod (ket i \\<^sub>C ket i)\ proof- have \(ket i \\<^sub>C 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 of_bool_def) lemma cinner_ket: \(ket i \\<^sub>C ket j) = of_bool (i=j)\ 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 of_bool_def) -lemma inj_ket[simp]: \inj ket\ +lemma inj_ket[simp]: \inj_on ket M\ 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" definition \canonical_basis_length_ell2 (_ :: 'a ell2 itself) = length (Enum.enum :: 'a list)\ 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 show \canonical_basis_length TYPE('a ell2) = length (canonical_basis :: 'a ell2 list)\ by (simp add: canonical_basis_length_ell2_def canonical_basis_ell2_def) 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] card_UNIV_length_enum by metis 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. A *\<^sub>V ket x = B *\<^sub>V 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) \\<^sub>C ket j = ket i \\<^sub>C (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_closure[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[simp]: "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 lemma bounded_clinear_equal_ket: fixes f g :: \'a ell2 \ _\ assumes \bounded_clinear f\ assumes \bounded_clinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule bounded_clinear_eq_on_closure[of f g \range ket\]) using assms by auto lemma bounded_antilinear_equal_ket: fixes f g :: \'a ell2 \ _\ assumes \bounded_antilinear f\ assumes \bounded_antilinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule bounded_antilinear_eq_on[of f g \range ket\]) using assms by auto lemma is_onb_ket[simp]: \is_onb (range ket)\ by (auto simp: is_onb_def) lemma ell2_sum_ket: \\ = (\i\UNIV. Rep_ell2 \ i *\<^sub>C ket i)\ for \ :: \_::finite ell2\ apply transfer apply (rule ext) apply (subst sum_single) by auto lemma trunc_ell2_singleton: \trunc_ell2 {x} \ = Rep_ell2 \ x *\<^sub>C ket x\ apply transfer by auto lemma trunc_ell2_insert: \trunc_ell2 (insert x M) \ = Rep_ell2 \ x *\<^sub>C ket x + trunc_ell2 M \\ if \x \ M\ using trunc_ell2_union_disjoint[where M=\{x}\ and N=M] using that by (auto simp: trunc_ell2_singleton) lemma trunc_ell2_finite_sum: \trunc_ell2 M \ = (\i\M. Rep_ell2 \ i *\<^sub>C ket i)\ if \finite M\ using that apply induction by (auto simp: trunc_ell2_insert) lemma is_orthogonal_trunc_ell2: \is_orthogonal (trunc_ell2 M \) (trunc_ell2 N \)\ if \M \ N = {}\ proof - have *: \cnj (if i \ M then a else 0) * (if i \ N then b else 0) = 0\ for a b i using that by auto show ?thesis apply (transfer fixing: M N) by (simp add: * ) qed 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 lemma ell2_decompose_has_sum: \((\x. Rep_ell2 \ x *\<^sub>C ket x) has_sum \) UNIV\ proof (unfold has_sum_def) have *: \trunc_ell2 M \ = (\x\M. Rep_ell2 \ x *\<^sub>C ket x)\ if \finite M\ for M using that apply induction by (auto simp: trunc_ell2_insert) show \(sum (\x. Rep_ell2 \ x *\<^sub>C ket x) \ \) (finite_subsets_at_top UNIV)\ apply (rule Lim_transform_eventually) apply (rule trunc_ell2_lim_at_UNIV) using * by (rule eventually_finite_subsets_at_top_weakI) qed lemma ell2_decompose_infsum: \\ = (\\<^sub>\x. Rep_ell2 \ x *\<^sub>C ket x)\ by (metis ell2_decompose_has_sum infsumI) lemma ell2_decompose_summable: \(\x. Rep_ell2 \ x *\<^sub>C ket x) summable_on UNIV\ using ell2_decompose_has_sum summable_on_def by blast lemma Rep_ell2_cblinfun_apply_sum: \Rep_ell2 (A *\<^sub>V \) y = (\\<^sub>\x. Rep_ell2 \ x * Rep_ell2 (A *\<^sub>V ket x) y)\ proof - have 1: \bounded_linear (\z. Rep_ell2 (A *\<^sub>V z) y)\ by (auto intro!: bounded_clinear_compose[unfolded o_def, OF bounded_clinear_Rep_ell2] cblinfun.bounded_clinear_right bounded_clinear.bounded_linear) have 2: \(\x. Rep_ell2 \ x *\<^sub>C ket x) summable_on UNIV\ by (simp add: ell2_decompose_summable) have \Rep_ell2 (A *\<^sub>V \) y = Rep_ell2 (A *\<^sub>V (\\<^sub>\x. Rep_ell2 \ x *\<^sub>C ket x)) y\ by (simp flip: ell2_decompose_infsum) also have \\ = (\\<^sub>\x. Rep_ell2 (A *\<^sub>V (Rep_ell2 \ x *\<^sub>C ket x)) y)\ apply (subst infsum_bounded_linear[symmetric, where f=\\z. Rep_ell2 (A *\<^sub>V z) y\]) using 1 2 by (auto simp: o_def) also have \\ = (\\<^sub>\x. Rep_ell2 \ x * Rep_ell2 (A *\<^sub>V ket x) y)\ by (simp add: cblinfun.scaleC_right scaleC_ell2.rep_eq) finally show ?thesis by - qed subsection \One-dimensional spaces\ instantiation ell2 :: (CARD_1) one begin lift_definition one_ell2 :: "'a ell2" is "\_. 1" by simp instance.. end lemma ket_CARD_1_is_1: \ket x = 1\ for x :: \'a::CARD_1\ apply transfer by simp instantiation ell2 :: (CARD_1) times begin lift_definition times_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x * b x" by simp instance.. end instantiation ell2 :: (CARD_1) divide begin lift_definition divide_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x / b x" by simp instance.. end instantiation ell2 :: (CARD_1) inverse begin lift_definition inverse_ell2 :: "'a ell2 \ 'a ell2" is "\a x. inverse (a x)" by simp instance.. end 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 intro_classes 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" apply transfer 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 \Explicit bounded operators\ definition explicit_cblinfun :: \('a \ 'b \ complex) \ ('b ell2, 'a ell2) cblinfun\ where - \explicit_cblinfun m = cblinfun_extension (range ket) (\a. Abs_ell2 (\j. m j (inv ket a)))\ + \explicit_cblinfun M = cblinfun_extension (range ket) (\a. Abs_ell2 (\j. M j (inv ket a)))\ definition explicit_cblinfun_exists :: \('a \ 'b \ complex) \ bool\ where - \explicit_cblinfun_exists m \ - (\a. has_ell2_norm (\j. m j a)) \ - cblinfun_extension_exists (range ket) (\a. Abs_ell2 (\j. m j (inv ket a)))\ + \explicit_cblinfun_exists M \ + (\a. has_ell2_norm (\j. M j a)) \ + cblinfun_extension_exists (range ket) (\a. Abs_ell2 (\j. M j (inv ket a)))\ lemma explicit_cblinfun_exists_bounded: - fixes B :: real assumes \\S T \. finite S \ finite T \ (\a. a\T \ \ a = 0) \ - (\b\S. (cmod (\a\T. \ a *\<^sub>C m b a))\<^sup>2) \ B * (\a\T. (cmod (\ a))\<^sup>2)\ - shows \explicit_cblinfun_exists m\ + (\b\S. (cmod (\a\T. \ a *\<^sub>C M b a))\<^sup>2) \ B * (\a\T. (cmod (\ a))\<^sup>2)\ + shows \explicit_cblinfun_exists M\ proof - define F f where \F = complex_vector.construct (range ket) f\ - and \f = (\a. Abs_ell2 (\j. m j (inv ket a)))\ + and \f = (\a. Abs_ell2 (\j. M j (inv ket a)))\ from assms[where S=\{}\ and T=\{undefined}\ and \=\\x. of_bool (x=undefined)\] have \B \ 0\ by auto - have has_norm: \has_ell2_norm (\b. m b a)\ for a + have has_norm: \has_ell2_norm (\b. M b a)\ for a proof (unfold has_ell2_norm_def, intro nonneg_bdd_above_summable_on bdd_aboveI) - show \0 \ cmod ((m x a)\<^sup>2)\ for x + show \0 \ cmod ((M x a)\<^sup>2)\ for x by simp fix B' - assume \B' \ sum (\x. cmod ((m x a)\<^sup>2)) ` {F. F \ UNIV \ finite F}\ - then obtain S where [simp]: \finite S\ and B'_def: \B' = (\x\S. cmod ((m x a)\<^sup>2))\ + assume \B' \ sum (\x. cmod ((M x a)\<^sup>2)) ` {F. F \ UNIV \ finite F}\ + then obtain S where [simp]: \finite S\ and B'_def: \B' = (\x\S. cmod ((M x a)\<^sup>2))\ by blast from assms[where S=S and T=\{a}\ and \=\\x. of_bool (x=a)\] show \B' \ B\ by (simp add: norm_power B'_def) qed have \clinear F\ by (auto intro!: complex_vector.linear_construct simp: F_def) have F_B: \norm (F \) \ (sqrt B) * norm \\ if \_range_ket: \\ \ cspan (range ket)\ for \ proof - from that obtain T' where \finite T'\ and \T' \ range ket\ and \T': \\ \ cspan T'\ by (meson vector_finitely_spanned) (* from that obtain T' r where \finite T'\ and \T' \ range ket\ and \T': \\ = (\v\T'. r v *\<^sub>C v)\ unfolding complex_vector.span_explicit by blast *) then obtain T where T'_def: \T' = ket ` T\ by (meson subset_image_iff) have \finite T\ by (metis T'_def \finite T'\ finite_image_iff inj_ket inj_on_subset subset_UNIV) have \T: \\ \ cspan (ket ` T)\ using T'_def \T' by blast have Rep_\: \Rep_ell2 \ x = 0\ if \x \ T\ for x using _ _ \T apply (rule complex_vector.linear_eq_0_on_span) apply auto by (metis ket.rep_eq that of_bool_def) have \norm (trunc_ell2 S (F \)) \ sqrt B * norm \\ if \finite S\ for S proof - have *: \cconstruct (range ket) f \ = (\x\T. Rep_ell2 \ x *\<^sub>C f (ket x))\ proof (rule complex_vector.linear_eq_on[where x=\ and B=\ket ` T\]) show \clinear (cconstruct (range ket) f)\ using F_def \clinear F\ by blast show \clinear (\a. \x\T. Rep_ell2 a x *\<^sub>C f (ket x))\ by (auto intro!: clinear_compose[unfolded o_def, OF clinear_Rep_ell2] complex_vector.linear_compose_sum) show \\ \ cspan (ket ` T)\ by (simp add: \T) have \f b = (\x\T. Rep_ell2 b x *\<^sub>C f (ket x))\ if \b \ ket ` T\ for b proof - define b' where \b' = inv ket b\ have bb': \b = ket b'\ using b'_def that by force show ?thesis apply (subst sum_single[where i=b']) using that by (auto simp add: \finite T\ bb' ket.rep_eq) qed then show \cconstruct (range ket) f b = (\x\T. Rep_ell2 b x *\<^sub>C f (ket x))\ if \b \ ket ` T\ for b apply (subst complex_vector.construct_basis) using that by auto qed have \(norm (trunc_ell2 S (F \)))\<^sup>2 = (norm (trunc_ell2 S (\x\T. Rep_ell2 \ x *\<^sub>C f (ket x))))\<^sup>2\ apply (rule arg_cong[where f=\\x. (norm (trunc_ell2 _ x))\<^sup>2\]) by (simp add: F_def * ) - also have \\ = (norm (trunc_ell2 S (\x\T. Rep_ell2 \ x *\<^sub>C Abs_ell2 (\b. m b x))))\<^sup>2\ + also have \\ = (norm (trunc_ell2 S (\x\T. Rep_ell2 \ x *\<^sub>C Abs_ell2 (\b. M b x))))\<^sup>2\ by (simp add: f_def) - also have \\ = (\i\S. (cmod (Rep_ell2 (\x\T. Rep_ell2 \ x *\<^sub>C Abs_ell2 (\b. m b x)) i))\<^sup>2)\ + also have \\ = (\i\S. (cmod (Rep_ell2 (\x\T. Rep_ell2 \ x *\<^sub>C Abs_ell2 (\b. M b x)) i))\<^sup>2)\ by (simp add: that norm_trunc_ell2_finite real_sqrt_pow2 sum_nonneg) - also have \\ = (\i\S. (cmod (\x\T. Rep_ell2 \ x *\<^sub>C Rep_ell2 (Abs_ell2 (\b. m b x)) i))\<^sup>2)\ + also have \\ = (\i\S. (cmod (\x\T. Rep_ell2 \ x *\<^sub>C Rep_ell2 (Abs_ell2 (\b. M b x)) i))\<^sup>2)\ by (simp add: complex_vector.linear_sum[OF clinear_Rep_ell2] clinear.scaleC[OF clinear_Rep_ell2]) - also have \\ = (\i\S. (cmod (\x\T. Rep_ell2 \ x *\<^sub>C m i x))\<^sup>2)\ + also have \\ = (\i\S. (cmod (\x\T. Rep_ell2 \ x *\<^sub>C M i x))\<^sup>2)\ using has_norm by (simp add: Abs_ell2_inverse) also have \\ \ B * (\x\T. (cmod (Rep_ell2 \ x))\<^sup>2)\ using \finite S\ \finite T\ Rep_\ by (rule assms) also have \\ = B * ((norm (trunc_ell2 T \))\<^sup>2)\ by (simp add: \finite T\ norm_trunc_ell2_finite sum_nonneg) also have \\ \ B * (norm \)\<^sup>2\ by (simp add: mult_left_mono \B \ 0\ trunc_ell2_reduces_norm) finally show ?thesis apply (rule_tac power2_le_imp_le) by (simp_all add: \0 \ B\ power_mult_distrib) qed then show ?thesis by (rule norm_ell2_bound_trunc) qed then have \cblinfun_extension_exists (cspan (range ket)) F\ apply (rule cblinfun_extension_exists_hilbert[rotated -1]) by (auto intro: \clinear F\ complex_vector.linear_add complex_vector.linear_scale) then have ex: \cblinfun_extension_exists (range ket) f\ apply (rule cblinfun_extension_exists_restrict[rotated -1]) by (simp_all add: F_def complex_vector.span_superset complex_vector.construct_basis) from ex has_norm show ?thesis using explicit_cblinfun_exists_def f_def by blast qed -lemma explicit_cblinfun_exists_finite_dim[simp]: \explicit_cblinfun_exists m\ for m :: "_::finite \ _ :: finite \ _" +lemma explicit_cblinfun_exists_finite_dim[simp]: \explicit_cblinfun_exists m\ for m :: "_::finite \ _::finite \ _" by (auto simp: explicit_cblinfun_exists_def cblinfun_extension_exists_finite_dim) -lemma explicit_cblinfun_ket: \explicit_cblinfun m *\<^sub>V ket a = Abs_ell2 (\b. m b a)\ if \explicit_cblinfun_exists m\ +lemma explicit_cblinfun_ket: \explicit_cblinfun M *\<^sub>V ket a = Abs_ell2 (\b. M b a)\ if \explicit_cblinfun_exists M\ using that by (auto simp: explicit_cblinfun_exists_def explicit_cblinfun_def cblinfun_extension_apply) -lemma Rep_ell2_explicit_cblinfun_ket[simp]: \Rep_ell2 (explicit_cblinfun m *\<^sub>V ket a) b = m b a\ if \explicit_cblinfun_exists m\ +lemma Rep_ell2_explicit_cblinfun_ket[simp]: \Rep_ell2 (explicit_cblinfun M *\<^sub>V ket a) b = M b a\ if \explicit_cblinfun_exists M\ using that apply (simp add: explicit_cblinfun_ket) by (simp add: Abs_ell2_inverse explicit_cblinfun_exists_def) 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 assumes "inj_map \" shows classical_operator_exists_inj: "classical_operator_exists \" and classical_operator_norm_inj: \norm (classical_operator \) \ 1\ proof - have \is_orthogonal (case \ x of None \ 0 | Some x' \ ket x') (case \ y of None \ 0 | Some y' \ ket y')\ if \x \ y\ for x y apply (cases \\ x\; cases \\ y\) using that assms by (auto simp add: inj_map_def) then have 1: \is_orthogonal (case \ (inv ket x) of None \ 0 | Some x' \ ket x') (case \ (inv ket y) of None \ 0 | Some y' \ ket y')\ if \x \ range ket\ and \y \ range ket\ and \x \ y\ for x y using that by auto have \norm (case \ x of None \ 0 | Some x \ ket x) \ 1 * norm (ket x)\ for x apply (cases \\ x\) by auto then have 2: \norm (case \ (inv ket x) of None \ 0 | Some x \ ket x) \ 1 * norm x\ if \x \ range ket\ for x using that by auto show \classical_operator_exists \\ unfolding classical_operator_exists_def using _ _ 1 2 apply (rule cblinfun_extension_exists_ortho) by simp_all show \norm (classical_operator \) \ 1\ unfolding classical_operator_def Let_def using _ _ 1 2 apply (rule cblinfun_extension_exists_ortho_norm) by simp_all 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) \\<^sub>C ket j = ket i \\<^sub>C (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) \\<^sub>C ket j = (classical_operator (inv_map \) *\<^sub>V ket i) \\<^sub>C ket j" unfolding F_def by blast also have "\ = ((case inv_map \ i of Some k \ ket k | None \ 0) \\<^sub>C ket j)" using w1 by simp also have "\ = (ket i \\<^sub>C (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) \\<^sub>C ket j = ket i \\<^sub>C (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 \\<^sub>C ket j) = (ket i \\<^sub>C 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) \\<^sub>C ket j = ket i \\<^sub>C (case Some c of None \ 0 | Some a \ ket a)" by simp thus "(case inv_map \ i of None \ 0 | Some a \ ket a) \\<^sub>C ket j = ket i \\<^sub>C (case \ j of None \ 0 | Some a \ ket a)" by (simp add: Some.hyps s1) qed qed also have "\ = ket i \\<^sub>C (classical_operator \ *\<^sub>V ket j)" by (simp add: w2) also have "\ = ket i \\<^sub>C (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/extra/Extra_Jordan_Normal_Form.thy b/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy @@ -1,419 +1,419 @@ section \\Extra_Jordan_Normal_Form\ -- Additional results for \<^session>\Jordan_Normal_Form\\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Extra_Jordan_Normal_Form imports Jordan_Normal_Form.Matrix Jordan_Normal_Form.Schur_Decomposition begin text \We define bundles to activate/deactivate the notation from \<^session>\Jordan_Normal_Form\. Reactivate the notation locally via "@{theory_text \includes jnf_notation\}" in a lemma statement. (Or sandwich a declaration using that notation between "@{theory_text \unbundle jnf_notation ... unbundle no_jnf_notation\}.) \ bundle jnf_notation begin notation transpose_mat ("(_\<^sup>T)" [1000]) notation cscalar_prod (infix "\c" 70) notation vec_index (infixl "$" 100) notation smult_vec (infixl "\\<^sub>v" 70) notation scalar_prod (infix "\" 70) notation index_mat (infixl "$$" 100) notation smult_mat (infixl "\\<^sub>m" 70) notation mult_mat_vec (infixl "*\<^sub>v" 70) notation pow_mat (infixr "^\<^sub>m" 75) notation append_vec (infixr "@\<^sub>v" 65) notation append_rows (infixr "@\<^sub>r" 65) end bundle no_jnf_notation begin no_notation transpose_mat ("(_\<^sup>T)" [1000]) no_notation cscalar_prod (infix "\c" 70) no_notation vec_index (infixl "$" 100) no_notation smult_vec (infixl "\\<^sub>v" 70) no_notation scalar_prod (infix "\" 70) no_notation index_mat (infixl "$$" 100) no_notation smult_mat (infixl "\\<^sub>m" 70) no_notation mult_mat_vec (infixl "*\<^sub>v" 70) no_notation pow_mat (infixr "^\<^sub>m" 75) no_notation append_vec (infixr "@\<^sub>v" 65) no_notation append_rows (infixr "@\<^sub>r" 65) end unbundle jnf_notation lemma mat_entry_explicit: fixes M :: "'a::field mat" assumes "M \ carrier_mat m n" and "i < m" and "j < n" shows "vec_index (M *\<^sub>v unit_vec n j) i = M $$ (i,j)" using assms by auto lemma mat_adjoint_def': "mat_adjoint M = transpose_mat (map_mat conjugate M)" apply (rule mat_eq_iff[THEN iffD2]) apply (auto simp: mat_adjoint_def transpose_mat_def) apply (subst mat_of_rows_index) by auto lemma mat_adjoint_swap: fixes M ::"complex mat" assumes "M \ carrier_mat nB nA" and "iA < dim_row M" and "iB < dim_col M" shows "(mat_adjoint M)$$(iB,iA) = cnj (M$$(iA,iB))" unfolding transpose_mat_def map_mat_def by (simp add: assms(2) assms(3) mat_adjoint_def') lemma cscalar_prod_adjoint: fixes M:: "complex mat" assumes "M \ carrier_mat nB nA" and "dim_vec v = nA" and "dim_vec u = nB" shows "v \c ((mat_adjoint M) *\<^sub>v u) = (M *\<^sub>v v) \c u" unfolding mat_adjoint_def using assms(1) assms(2,3)[symmetric] apply (simp add: scalar_prod_def sum_distrib_left field_simps) by (intro sum.swap) lemma scaleC_minus1_left_vec: "-1 \\<^sub>v v = - v" for v :: "_::ring_1 vec" unfolding smult_vec_def uminus_vec_def by auto lemma square_nneg_complex: fixes x :: complex assumes "x \ \" shows "x^2 \ 0" apply (cases x) using assms unfolding Reals_def less_eq_complex_def by auto definition "vec_is_zero n v = (\i vec_is_zero n v \ v = 0\<^sub>v n" unfolding vec_is_zero_def apply auto by (metis index_zero_vec(1)) fun gram_schmidt_sub0 where "gram_schmidt_sub0 n us [] = us" | "gram_schmidt_sub0 n us (w # ws) = (let w' = adjuster n w us + w in if vec_is_zero n w' then gram_schmidt_sub0 n us ws else gram_schmidt_sub0 n (w' # us) ws)" lemma (in cof_vec_space) adjuster_already_in_span: assumes "w \ carrier_vec n" assumes us_carrier: "set us \ carrier_vec n" assumes "corthogonal us" assumes "w \ span (set us)" shows "adjuster n w us + w = 0\<^sub>v n" proof - define v U where "v = adjuster n w us + w" and "U = set us" have span: "v \ span U" unfolding v_def U_def apply (rule adjust_preserves_span[THEN iffD1]) using assms corthogonal_distinct by simp_all have v_carrier: "v \ carrier_vec n" by (simp add: v_def assms corthogonal_distinct) have "v \c us!i = 0" if "i < length us" for i unfolding v_def apply (rule adjust_zero) using that assms by simp_all hence "v \c u = 0" if "u \ U" for u by (metis assms(3) U_def corthogonal_distinct distinct_Ex1 that) hence ortho: "u \c v = 0" if "u \ U" for u apply (subst conjugate_zero_iff[symmetric]) apply (subst conjugate_vec_sprod_comm) using that us_carrier v_carrier apply (auto simp: U_def)[2] apply (subst conjugate_conjugate_sprod) using that us_carrier v_carrier by (auto simp: U_def) from span obtain a where v: "lincomb a U = v" apply atomize_elim apply (rule finite_in_span[simplified]) unfolding U_def using us_carrier by auto have "v \c v = (\u\U. (a u \\<^sub>v u) \c v)" apply (subst v[symmetric]) unfolding lincomb_def apply (subst finsum_scalar_prod_sum) using U_def span us_carrier by auto also have "\ = (\u\U. a u * (u \c v))" using U_def assms(1) in_mono us_carrier v_def by fastforce also have "\ = (\u\U. a u * conjugate 0)" apply (rule sum.cong, simp) using span span_closed U_def us_carrier ortho by auto also have "\ = 0" by auto finally have "v \c v = 0" by - thus "v = 0\<^sub>v n" using U_def conjugate_square_eq_0_vec span span_closed us_carrier by blast qed lemma (in cof_vec_space) gram_schmidt_sub0_result: assumes "gram_schmidt_sub0 n us ws = us'" and "set ws \ carrier_vec n" and "set us \ carrier_vec n" and "distinct us" and "~ lin_dep (set us)" and "corthogonal us" shows "set us' \ carrier_vec n \ distinct us' \ corthogonal us' \ span (set (us @ ws)) = span (set us')" using assms proof (induct ws arbitrary: us us') case (Cons w ws) show ?case proof (cases "w \ span (set us)") case False let ?v = "adjuster n w us" have wW[simp]: "set (w#ws) \ carrier_vec n" using Cons by simp hence W[simp]: "set ws \ carrier_vec n" and w[simp]: "w : carrier_vec n" by auto have U[simp]: "set us \ carrier_vec n" using Cons by simp have UW: "set (us@ws) \ carrier_vec n" by simp have wU: "set (w#us) \ carrier_vec n" by simp have dist_U: "distinct us" using Cons by simp have w_U: "w \ set us" using False using span_mem by auto have ind_U: "~ lin_dep (set us)" using Cons by simp have ind_wU: "~ lin_dep (insert w (set us))" apply (subst lin_dep_iff_in_span[simplified, symmetric]) using w_U ind_U False by auto thm lin_dep_iff_in_span[simplified, symmetric] have corth: "corthogonal us" using Cons by simp have "?v + w \ 0\<^sub>v n" by (simp add: False adjust_nonzero dist_U) hence "\ vec_is_zero n (?v + w)" by (simp add: vec_is_zero) hence U'def: "gram_schmidt_sub0 n ((?v + w)#us) ws = us'" using Cons by simp have v: "?v : carrier_vec n" using dist_U by auto hence vw: "?v + w : carrier_vec n" by auto hence vwU: "set ((?v + w) # us) \ carrier_vec n" by auto have vsU: "?v : span (set us)" apply (rule adjuster_in_span[OF w]) using Cons by simp_all hence vsUW: "?v : span (set (us @ ws))" using span_is_monotone[of "set us" "set (us@ws)"] by auto have wsU: "w \ span (set us)" using lin_dep_iff_in_span[OF U ind_U w w_U] ind_wU by auto hence vwU: "?v + w \ span (set us)" using adjust_not_in_span[OF w U dist_U] by auto have span: "?v + w \ span (set us)" apply (subst span_add[symmetric]) by (simp_all add: False vsU) hence vwUS: "?v + w \ set us" using span_mem by auto have vwU: "set ((?v + w) # us) \ carrier_vec n" using U w vw by simp have dist2: "distinct (((?v + w) # us))" using vwUS by (simp add: dist_U) have orth2: "corthogonal ((adjuster n w us + w) # us)" using adjust_orthogonal[OF U corth w wsU]. have ind_vwU: "~ lin_dep (set ((adjuster n w us + w) # us))" apply simp apply (subst lin_dep_iff_in_span[simplified, symmetric]) by (simp_all add: ind_U vw vwUS span) have span_UwW_U': "span (set (us @ w # ws)) = span (set us')" using Cons(1)[OF U'def W vwU dist2 ind_vwU orth2] using span_Un[OF vwU wU gram_schmidt_sub_span[OF w U dist_U] W W refl] by simp show ?thesis apply (intro conjI) using Cons(1)[OF U'def W vwU dist2 ind_vwU orth2] span_UwW_U' by simp_all next case True let ?v = "adjuster n w us" have "?v + w = 0\<^sub>v n" apply (rule adjuster_already_in_span) using True Cons by auto hence "vec_is_zero n (?v + w)" by (simp add: vec_is_zero) hence U'_def: "us' = gram_schmidt_sub0 n us ws" using Cons by simp have span: "span (set (us @ w # ws)) = span (set us')" proof - have wU_U: "span (set (w # us)) = span (set us)" apply (subst already_in_span[OF _ True, simplified]) using Cons by auto have "span (set (us @ w # ws)) = span (set (w # us) \ set ws)" by simp also have "\ = span (set us \ set ws)" apply (rule span_Un) using wU_U Cons by auto also have "\ = local.span (set us')" using Cons U'_def by auto finally show ?thesis by - qed moreover have "set us' \ carrier_vec n \ distinct us' \ corthogonal us'" unfolding U'_def using Cons by simp ultimately show ?thesis by auto qed qed simp text \This is a variant of \<^term>\Gram_Schmidt.gram_schmidt\ that does not require the input vectors - \<^term>\ws\ to be distinct or orthogonal. (In comparison to \<^term>\Gram_Schmidt.gram_schmidt\, + \<^term>\ws\ to be distinct or linearly independent. (In comparison to \<^term>\Gram_Schmidt.gram_schmidt\, our version also returns the result in reversed order.)\ definition "gram_schmidt0 n ws = gram_schmidt_sub0 n [] ws" lemma (in cof_vec_space) gram_schmidt0_result: fixes ws defines "us' \ gram_schmidt0 n ws" assumes ws: "set ws \ carrier_vec n" shows "set us' \ carrier_vec n" (is ?thesis1) and "distinct us'" (is ?thesis2) and "corthogonal us'" (is ?thesis3) and "span (set ws) = span (set us')" (is ?thesis4) proof - have carrier_empty: "set [] \ carrier_vec n" by auto have distinct_empty: "distinct []" by simp have indep_empty: "lin_indpt (set [])" using basis_def subset_li_is_li unit_vecs_basis by auto have ortho_empty: "corthogonal []" by auto note gram_schmidt_sub0_result' = gram_schmidt_sub0_result [OF us'_def[symmetric, THEN meta_eq_to_obj_eq, unfolded gram_schmidt0_def] ws carrier_empty distinct_empty indep_empty ortho_empty] thus ?thesis1 ?thesis2 ?thesis3 ?thesis4 by auto qed locale complex_vec_space = cof_vec_space n "TYPE(complex)" for n :: nat lemma gram_schmidt0_corthogonal: assumes a1: "corthogonal R" and a2: "\x. x \ set R \ dim_vec x = d" shows "gram_schmidt0 d R = rev R" proof - have "gram_schmidt_sub0 d U R = rev R @ U" if "corthogonal ((rev U) @ R)" and "\x. x \ set U \ set R \ dim_vec x = d" for U proof (insert that, induction R arbitrary: U) case Nil show ?case by auto next case (Cons a R) have "a \ set (rev U @ a # R)" by simp moreover have uar: "corthogonal (rev U @ a # R)" by (simp add: Cons.prems(1)) ultimately have \a \ 0\<^sub>v d\ unfolding corthogonal_def by (metis conjugate_zero_vec in_set_conv_nth scalar_prod_right_zero zero_carrier_vec) then have nonzero_a: "\ vec_is_zero d a" by (simp add: Cons.prems(2) vec_is_zero) define T where "T = rev U @ a # R" have "T ! length (rev U) = a" unfolding T_def by (meson nth_append_length) moreover have "(T ! i \c T ! j = 0) = (i \ j)" if "ic T ! j = 0) = (length (rev U) \ j)" if "jc T ! j = 0" if "j length (rev U)" for j using that by blast hence "a \c T ! j = 0" if "j < length (rev U)" for j using \T ! length (rev U) = a\ that(1) \length (rev U) < length T\ dual_order.strict_trans by blast moreover have "T ! j = (rev U) ! j" if "j < length (rev U)" for j by (smt T_def \length (rev U) < length T\ dual_order.strict_trans list_update_append1 list_update_id nth_list_update_eq that) ultimately have "a \c u = 0" if "u \ set (rev U)" for u by (metis in_set_conv_nth that) hence "a \c u = 0" if "u \ set U" for u by (simp add: that) moreover have "\x. x \ set U \ dim_vec x = d" by (simp add: Cons.prems(2)) ultimately have "adjuster d a U = 0\<^sub>v d" proof(induction U) case Nil then show ?case by simp next case (Cons u U) moreover have "0 \\<^sub>v u + 0\<^sub>v d = 0\<^sub>v d" proof- have "dim_vec u = d" by (simp add: calculation(3)) thus ?thesis by auto qed ultimately show ?case by auto qed hence adjuster_a: "adjuster d a U + a = a" by (simp add: Cons.prems(2) carrier_vecI) have "gram_schmidt_sub0 d U (a # R) = gram_schmidt_sub0 d (a # U) R" by (simp add: adjuster_a nonzero_a) also have "\ = rev (a # R) @ U" apply (subst Cons.IH) using Cons.prems by simp_all finally show ?case by - qed from this[where U="[]"] show ?thesis unfolding gram_schmidt0_def using assms by auto qed lemma adjuster_carrier': (* Like adjuster_carrier but with one assm less *) assumes w: "(w :: 'a::conjugatable_field vec) : carrier_vec n" and us: "set (us :: 'a vec list) \ carrier_vec n" shows "adjuster n w us \ carrier_vec n" by (insert us, induction us, auto) lemma eq_mat_on_vecI: fixes M N :: \'a::field mat\ assumes eq: \\v. v\carrier_vec nA \ M *\<^sub>v v = N *\<^sub>v v\ assumes [simp]: \M \ carrier_mat nB nA\ \N \ carrier_mat nB nA\ shows \M = N\ proof (rule eq_matI) show [simp]: \dim_row M = dim_row N\ \dim_col M = dim_col N\ using assms(2) assms(3) by blast+ fix i j assume [simp]: \i < dim_row N\ \j < dim_col N\ show \M $$ (i, j) = N $$ (i, j)\ thm mat_entry_explicit[where M=M] apply (subst mat_entry_explicit[symmetric]) using assms apply auto[3] apply (subst mat_entry_explicit[symmetric]) using assms apply auto[3] apply (subst eq) apply auto using assms(3) unit_vec_carrier by blast qed lemma list_of_vec_plus: fixes v1 v2 :: \complex vec\ assumes \dim_vec v1 = dim_vec v2\ shows \list_of_vec (v1 + v2) = map2 (+) (list_of_vec v1) (list_of_vec v2)\ proof- have \i < dim_vec v1 \ (list_of_vec (v1 + v2)) ! i = (map2 (+) (list_of_vec v1) (list_of_vec v2)) ! i\ for i by (simp add: assms) thus ?thesis by (metis assms index_add_vec(2) length_list_of_vec length_map map_fst_zip nth_equalityI) qed lemma list_of_vec_mult: fixes v :: \complex vec\ shows \list_of_vec (c \\<^sub>v v) = map ((*) c) (list_of_vec v)\ by (metis (mono_tags, lifting) index_smult_vec(1) index_smult_vec(2) length_list_of_vec length_map nth_equalityI nth_list_of_vec nth_map) lemma map_map_vec_cols: \map (map_vec f) (cols m) = cols (map_mat f m)\ by (simp add: cols_def) lemma map_vec_conjugate: \map_vec conjugate v = conjugate v\ by fastforce unbundle no_jnf_notation end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Vector_Spaces.thy b/thys/Complex_Bounded_Operators/extra/Extra_Vector_Spaces.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Vector_Spaces.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Vector_Spaces.thy @@ -1,245 +1,239 @@ section \\Extra_Vector_Spaces\ -- Additional facts about vector spaces\ theory Extra_Vector_Spaces imports "HOL-Analysis.Inner_Product" "HOL-Analysis.Euclidean_Space" "HOL-Library.Indicator_Function" "HOL-Analysis.Topology_Euclidean_Space" "HOL-Analysis.Line_Segment" "HOL-Analysis.Bounded_Linear_Function" Extra_General begin subsection \Euclidean spaces\ typedef 'a euclidean_space = "UNIV :: ('a \ real) set" .. setup_lifting type_definition_euclidean_space instantiation euclidean_space :: (type) real_vector begin lift_definition plus_euclidean_space :: "'a euclidean_space \ 'a euclidean_space \ 'a euclidean_space" is "\f g x. f x + g x" . lift_definition zero_euclidean_space :: "'a euclidean_space" is "\_. 0" . lift_definition uminus_euclidean_space :: "'a euclidean_space \ 'a euclidean_space" is "\f x. - f x" . lift_definition minus_euclidean_space :: "'a euclidean_space \ 'a euclidean_space \ 'a euclidean_space" is "\f g x. f x - g x". lift_definition scaleR_euclidean_space :: "real \ 'a euclidean_space \ 'a euclidean_space" is "\c f x. c * f x" . instance apply intro_classes by (transfer; auto intro: distrib_left distrib_right)+ end instantiation euclidean_space :: (finite) real_inner begin lift_definition inner_euclidean_space :: "'a euclidean_space \ 'a euclidean_space \ real" is "\f g. \x\UNIV. f x * g x :: real" . definition "norm_euclidean_space (x::'a euclidean_space) = sqrt (inner x x)" definition "dist_euclidean_space (x::'a euclidean_space) y = norm (x-y)" definition "sgn x = x /\<^sub>R norm x" for x::"'a euclidean_space" definition "uniformity = (INF e\{0<..}. principal {(x::'a euclidean_space, y). dist x y < e})" definition "open U = (\x\U. \\<^sub>F (x'::'a euclidean_space, y) in uniformity. x' = x \ y \ U)" instance proof intro_classes fix x :: "'a euclidean_space" and y :: "'a euclidean_space" and z :: "'a euclidean_space" show "dist (x::'a euclidean_space) y = norm (x - y)" and "sgn (x::'a euclidean_space) = x /\<^sub>R norm x" and "uniformity = (INF e\{0<..}. principal {(x, y). dist (x::'a euclidean_space) y < e})" and "open U = (\x\U. \\<^sub>F (x', y) in uniformity. (x'::'a euclidean_space) = x \ y \ U)" and "norm x = sqrt (inner x x)" for U unfolding dist_euclidean_space_def norm_euclidean_space_def sgn_euclidean_space_def uniformity_euclidean_space_def open_euclidean_space_def by simp_all show "inner x y = inner y x" apply transfer by (simp add: mult.commute) show "inner (x + y) z = inner x z + inner y z" proof transfer fix x y z::"'a \ real" have "(\i\UNIV. (x i + y i) * z i) = (\i\UNIV. x i * z i + y i * z i)" by (simp add: distrib_left mult.commute) thus "(\i\UNIV. (x i + y i) * z i) = (\j\UNIV. x j * z j) + (\k\UNIV. y k * z k)" by (subst sum.distrib[symmetric]) qed show "inner (r *\<^sub>R x) y = r * (inner x y)" for r proof transfer fix r and x y::"'a\real" have "(\i\UNIV. r * x i * y i) = (\i\UNIV. r * (x i * y i))" by (simp add: mult.assoc) thus "(\i\UNIV. r * x i * y i) = r * (\j\UNIV. x j * y j)" by (subst sum_distrib_left) qed show "0 \ inner x x" apply transfer by (simp add: sum_nonneg) show "(inner x x = 0) = (x = 0)" proof (transfer, rule) fix f :: "'a \ real" assume "(\i\UNIV. f i * f i) = 0" hence "f x * f x = 0" for x apply (rule_tac sum_nonneg_eq_0_iff[THEN iffD1, rule_format, where A=UNIV and x=x]) by auto thus "f = (\_. 0)" by auto qed auto qed end instantiation euclidean_space :: (finite) euclidean_space begin lift_definition euclidean_space_basis_vector :: "'a \ 'a euclidean_space" is "\x. indicator {x}" . definition "Basis_euclidean_space == (euclidean_space_basis_vector ` UNIV)" instance proof intro_classes fix u :: "'a euclidean_space" and v :: "'a euclidean_space" show "(Basis::'a euclidean_space set) \ {}" unfolding Basis_euclidean_space_def by simp show "finite (Basis::'a euclidean_space set)" unfolding Basis_euclidean_space_def by simp show "inner u v = (if u = v then 1 else 0)" if "u \ Basis" and "v \ Basis" using that unfolding Basis_euclidean_space_def apply transfer apply auto by (auto simp: indicator_def) show "(\v\Basis. inner u v = 0) = (u = 0)" unfolding Basis_euclidean_space_def apply transfer by auto qed end (* euclidean_space :: (finite) euclidean_space *) subsection \Misc\ lemma closure_bounded_linear_image_subset_eq: assumes f: "bounded_linear f" shows "closure (f ` closure S) = closure (f ` S)" by (meson closed_closure closure_bounded_linear_image_subset closure_minimal closure_mono closure_subset f image_mono subset_antisym) lemma not_singleton_real_normed_is_perfect_space[simp]: \class.perfect_space (open :: 'a::{not_singleton,real_normed_vector} set \ bool)\ apply standard by (metis UNIV_not_singleton clopen closed_singleton empty_not_insert) lemma infsum_bounded_linear: assumes \bounded_linear f\ assumes \g summable_on S\ shows \infsum (f \ g) S = f (infsum g S)\ apply (rule infsum_comm_additive) using assms blinfun_apply_induct blinfun.additive_right by (auto simp: linear_continuous_within) lemma has_sum_bounded_linear: assumes \bounded_linear f\ assumes \(g has_sum x) S\ shows \((f o g) has_sum (f x)) S\ apply (rule has_sum_comm_additive) using assms blinfun_apply_induct blinfun.additive_right apply auto using isCont_def linear_continuous_at by fastforce -lemma summable_on_bounded_linear: - assumes \bounded_linear f\ - assumes \g summable_on S\ - shows \(f o g) summable_on S\ - by (metis assms(1) assms(2) has_sum_bounded_linear infsum_bounded_linear summable_iff_has_sum_infsum) - lemma abs_summable_on_bounded_linear: assumes \bounded_linear f\ assumes \g abs_summable_on S\ shows \(f o g) abs_summable_on S\ proof - have bound: \norm (f (g x)) \ onorm f * norm (g x)\ for x apply (rule onorm) by (simp add: assms(1)) from assms(2) have \(\x. onorm f *\<^sub>R g x) abs_summable_on S\ by (auto intro!: summable_on_cmult_right) then have \(\x. f (g x)) abs_summable_on S\ apply (rule abs_summable_on_comparison_test) using bound by (auto simp: assms(1) onorm_pos_le) then show ?thesis by auto qed lemma norm_plus_leq_norm_prod: \norm (a + b) \ sqrt 2 * norm (a, b)\ proof - have \(norm (a + b))\<^sup>2 \ (norm a + norm b)\<^sup>2\ using norm_triangle_ineq by auto also have \\ \ 2 * ((norm a)\<^sup>2 + (norm b)\<^sup>2)\ by (smt (verit, best) power2_sum sum_squares_bound) also have \\ \ (sqrt 2 * norm (a, b))\<^sup>2\ by (auto simp: power_mult_distrib norm_prod_def simp del: power_mono_iff) finally show ?thesis by auto qed lemma ex_norm1: assumes \(UNIV::'a::real_normed_vector set) \ {0}\ shows \\x::'a. norm x = 1\ proof- have \\x::'a. x \ 0\ using assms by fastforce then obtain x::'a where \x \ 0\ by blast hence \norm x \ 0\ by simp hence \(norm x) / (norm x) = 1\ by simp moreover have \(norm x) / (norm x) = norm (x /\<^sub>R (norm x))\ by simp ultimately have \norm (x /\<^sub>R (norm x)) = 1\ by simp thus ?thesis by blast qed lemma bdd_above_norm_f: assumes "bounded_linear f" shows \bdd_above {norm (f x) |x. norm x = 1}\ proof- have \\M. \x. norm x = 1 \ norm (f x) \ M\ using assms by (metis bounded_linear.axioms(2) bounded_linear_axioms_def) thus ?thesis by auto qed lemma any_norm_exists: assumes \n \ 0\ shows \\\::'a::{real_normed_vector,not_singleton}. norm \ = n\ proof - obtain \ :: 'a where \\ \ 0\ using not_singleton_card by force then have \norm (n *\<^sub>R sgn \) = n\ using assms by (auto simp: sgn_div_norm abs_mult) then show ?thesis by blast qed lemma abs_summable_on_scaleR_left [intro]: fixes c :: \'a :: real_normed_vector\ assumes "c \ 0 \ f abs_summable_on A" shows "(\x. f x *\<^sub>R c) abs_summable_on A" apply (cases \c = 0\) apply simp by (auto intro!: summable_on_cmult_left assms simp flip: real_norm_def) lemma abs_summable_on_scaleR_right [intro]: fixes f :: \'a \ 'b :: real_normed_vector\ assumes "c \ 0 \ f abs_summable_on A" shows "(\x. c *\<^sub>R f x) abs_summable_on A" apply (cases \c = 0\) apply simp by (auto intro!: summable_on_cmult_right assms) end