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,1573 +1,1563 @@ 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\). +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_enum\ 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)) 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" 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" 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) 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" 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 lemma norm_ell2_vec_of_basis_enum: "norm \ = (let \' = vec_of_basis_enum \ in sqrt (\ i \ {0 ..< dim_vec \'}. let z = vec_index \' i in (Re z)\<^sup>2 + (Im z)\<^sup>2))" (is "_ = ?rhs") 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 "\ = ?rhs" unfolding Let_def norm_of_real 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 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 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 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::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::enum ell2 \\<^sub>C\<^sub>L 'b::enum ell2)) = CARD('a)\ by (simp add: mat_of_cblinfun_def) 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 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\ 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) 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 \Matrix operations\ +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_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 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)) \ is_subspace_of_vec_list (length (canonical_basis :: 'a list)) (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: "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[of _ _ mat_of_cblinfun]) proof (insert ortho_Snorm, insert \distinct Snorm\, induction Snorm) case Nil show ?case by simp next case (Cons a Snorm) from Cons.prems have [simp]: "a \ set Snorm" by simp have "sum proj (set (a # Snorm)) = proj a + sum proj (set Snorm)" by auto also have "\ = proj a + Proj (ccspan (set Snorm))" apply (subst Cons.IH) using Cons.prems apply auto by (meson Cons.prems(1) is_ortho_set_antimono set_subset_Cons) also have "\ = Proj (ccspan ({a} \ set Snorm))" apply (rule Proj_orthog_ccspan_union[symmetric]) by (metis Cons.prems(1) \a \ set Snorm\ is_ortho_set_def list.set_intros(1) list.set_intros(2) singleton_iff) finally show ?case by simp qed also have "\ = mat_of_cblinfun (Proj (ccspan (set S)))" unfolding Span_Snorm by simp finally show ?thesis by - qed 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)))" 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 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,4576 +1,4622 @@ 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 \ = \" apply transfer by simp lemma apply_id_cblinfun[simp]: \(*\<^sub>V) id_cblinfun = id\ by auto lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" apply transfer by (simp add: clinear_continuous_at) declare cblinfun.scaleC_left[simp] lemma cblinfun_apply_clinear[simp]: \clinear (cblinfun_apply A)\ using bounded_clinear.axioms(1) cblinfun_apply by blast lemma cblinfun_cinner_eqI: fixes A B :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ assumes \\\. cinner \ (A *\<^sub>V \) = cinner \ (B *\<^sub>V \)\ shows \A = B\ proof - define C where \C = A - B\ have C0[simp]: \cinner \ (C \) = 0\ for \ by (simp add: C_def assms cblinfun.diff_left cinner_diff_right) { fix f g \ have \0 = cinner (f + \ *\<^sub>C g) (C *\<^sub>V (f + \ *\<^sub>C g))\ by (simp add: cinner_diff_right minus_cblinfun.rep_eq) also have \\ = \ *\<^sub>C cinner f (C g) + cnj \ *\<^sub>C cinner g (C f)\ by (smt (z3) C0 add.commute add.right_neutral cblinfun.add_right cblinfun.scaleC_right cblinfun_cinner_right.rep_eq cinner_add_left cinner_scaleC_left complex_scaleC_def) finally have \\ *\<^sub>C cinner f (C g) = - cnj \ *\<^sub>C cinner g (C f)\ by (simp add: eq_neg_iff_add_eq_0) } then have \cinner f (C g) = 0\ for f g by (metis complex_cnj_i complex_cnj_one complex_vector.scale_cancel_right complex_vector.scale_left_imp_eq equation_minus_iff i_squared mult_eq_0_iff one_neq_neg_one) then have \C g = 0\ for g using cinner_eq_zero_iff by blast then have \C = 0\ by (simp add: cblinfun_eqI) then show \A = B\ using C_def by auto qed lemma id_cblinfun_not_0[simp]: \(id_cblinfun :: 'a::{complex_normed_vector, not_singleton} \\<^sub>C\<^sub>L _) \ 0\ by (metis (full_types) Extra_General.UNIV_not_singleton cblinfun.zero_left cblinfun_id_cblinfun_apply ex_norm1 norm_zero one_neq_zero) lemma cblinfun_norm_geqI: assumes \norm (f *\<^sub>V x) / norm x \ K\ shows \norm f \ K\ using assms apply transfer by (smt (z3) bounded_clinear.bounded_linear le_onorm) (* This lemma is proven in Complex_Bounded_Linear_Function0 but we add the [simp] only here because we try to keep Complex_Bounded_Linear_Function0 as close to Bounded_Linear_Function as possible. *) declare scaleC_conv_of_complex[simp] lemma cblinfun_eq_0_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = 0" shows \F *\<^sub>V x = 0\ apply (rule complex_vector.linear_eq_0_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = G *\<^sub>V s" shows \F *\<^sub>V x = G *\<^sub>V x\ apply (rule complex_vector.linear_eq_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_0_on_UNIV_span: fixes basis::\'a::complex_normed_vector set\ assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = 0" shows \F = 0\ by (metis cblinfun_eq_0_on_span UNIV_I assms cblinfun.zero_left cblinfun_eqI) lemma cblinfun_eq_on_UNIV_span: fixes basis::"'a::complex_normed_vector set" and \::"'a \ 'b::complex_normed_vector" assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = G *\<^sub>V s" shows \F = G\ proof- have "F - G = 0" apply (rule cblinfun_eq_0_on_UNIV_span[where basis=basis]) using assms by (auto simp add: cblinfun.diff_left) thus ?thesis by simp qed lemma cblinfun_eq_on_canonical_basis: fixes f g::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = g *\<^sub>V u" shows "f = g" apply (rule cblinfun_eq_on_UNIV_span[where basis=basis]) using assms is_generator_set is_cindependent_set by auto lemma cblinfun_eq_0_on_canonical_basis: fixes f ::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = 0" shows "f = 0" by (simp add: assms cblinfun_eq_on_canonical_basis) lemma cinner_canonical_basis_eq_0: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ 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 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 apply auto by (simp add: ex_norm1) next case False then have \(1 - \) * norm A > 0\ using assms by fastforce then obtain \ where geq: \norm (A *\<^sub>V \) \ norm A - ((1 - \) * norm A)\ and \norm \ = 1\ using cblinfun_norm_approx_witness by blast have \norm A * \ = norm A - (1 - \) * norm A\ by (simp add: mult.commute right_diff_distrib') also have \\ \ norm (A *\<^sub>V \)\ by (rule geq) finally show ?thesis using \norm \ = 1\ by auto qed lemma cblinfun_to_CARD_1_0[simp]: \(A :: _ \\<^sub>C\<^sub>L _::CARD_1) = 0\ apply (rule cblinfun_eqI) by auto lemma cblinfun_from_CARD_1_0[simp]: \(A :: _::CARD_1 \\<^sub>C\<^sub>L _) = 0\ apply (rule cblinfun_eqI) apply (subst CARD_1_vec_0) by auto lemma cblinfun_cspan_UNIV: fixes basis :: \('a::{complex_normed_vector,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_normed_vector) set\ and basisA :: \'a set\ and basisB :: \'b set\ assumes \cspan basisA = UNIV\ and \cspan basisB = UNIV\ assumes basis: \\a b. a\basisA \ b\basisB \ \F\basis. \a'\basisA. F *\<^sub>V a' = (if a'=a then b else 0)\ shows \cspan basis = UNIV\ proof - obtain basisA' where \basisA' \ basisA\ and \cindependent basisA'\ and \cspan basisA' = UNIV\ by (metis assms(1) complex_vector.maximal_independent_subset complex_vector.span_eq top_greatest) then have [simp]: \finite basisA'\ by (simp add: cindependent_cfinite_dim_finite) have basis': \\a b. a\basisA' \ b\basisB \ \F\basis. \a'\basisA'. F *\<^sub>V a' = (if a'=a then b else 0)\ using basis \basisA' \ basisA\ by fastforce obtain F where F: \F a b \ basis \ F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using basis' by metis then have F_apply: \F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' using that by auto have F_basis: \F a b \ basis\ if \a\basisA'\ \b\basisB\ for a b using that F by auto have b_span: \\G\cspan {F a b|b. b\basisB}. \a'\basisA'. G *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ for a b proof - from \cspan basisB = UNIV\ obtain r t where \finite t\ and \t \ basisB\ and b_lincom: \b = (\a\t. r a *\<^sub>C a)\ unfolding complex_vector.span_alt apply atomize_elim by blast define G where \G = (\i\t. r i *\<^sub>C F a i)\ have \G \ cspan {F a b|b. b\basisB}\ using \finite t\ \t \ basisB\ unfolding G_def by (smt (verit, ccfv_threshold) complex_vector.span_base complex_vector.span_scale complex_vector.span_sum mem_Collect_eq subset_eq) moreover have \G *\<^sub>V a' = (if a'=a then b else 0)\ if \a'\basisA'\ for a' apply (cases \a'=a\) using \t \ basisB\ \a\basisA'\ \a'\basisA'\ by (auto simp: b_lincom G_def cblinfun.sum_left F_apply intro!: sum.neutral sum.cong) ultimately show ?thesis by blast qed have a_span: \cspan (\a\basisA'. cspan {F a b|b. b\basisB}) = UNIV\ proof (intro equalityI subset_UNIV subsetI, rename_tac H) fix H obtain G where G: \G a b \ cspan {F a b|b. b\basisB} \ G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using b_span by blast then have G_cspan: \G a b \ cspan {F a b|b. b\basisB}\ if \a\basisA'\ for a b using that by auto from G have G: \G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' using that by auto define H' where \H' = (\a\basisA'. G a (H *\<^sub>V a))\ have \H' \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ unfolding H'_def using G_cspan by (smt (verit, del_insts) UN_iff complex_vector.span_clauses(1) complex_vector.span_sum) moreover have \H' = H\ using \cspan basisA' = UNIV\ apply (rule cblinfun_eq_on_UNIV_span) apply (auto simp: H'_def cblinfun.sum_left) apply (subst sum_single) by (auto simp: G) ultimately show \H \ cspan (\a\basisA'. cspan {F a b |b. b \ basisB})\ by simp qed moreover have \cspan basis \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ using F_basis by (smt (z3) UN_subset_iff complex_vector.span_alt complex_vector.span_minimal complex_vector.subspace_span mem_Collect_eq subset_iff) ultimately show \cspan basis = UNIV\ by auto qed instance cblinfun :: (\{cfinite_dim,complex_normed_vector}\, \{cfinite_dim,complex_normed_vector}\) cfinite_dim proof intro_classes obtain basisA :: \'a set\ where [simp]: \cspan basisA = UNIV\ \cindependent basisA\ \finite basisA\ using finite_basis by blast obtain basisB :: \'b set\ where [simp]: \cspan basisB = UNIV\ \cindependent basisB\ \finite basisB\ using finite_basis by blast define f where \f a b = cconstruct basisA (\x. if x=a then b else 0)\ for a :: 'a and b :: 'b have f_a: \f a b a = b\ if \a : basisA\ for a b by (simp add: complex_vector.construct_basis f_def that) have f_not_a: \f a b c = 0\ if \a : basisA\ and \c : basisA\ and \a \ c\for a b c using that by (simp add: complex_vector.construct_basis f_def) define F where \F a b = CBlinfun (f a b)\ for a b have \clinear (f a b)\ for a b by (auto intro: complex_vector.linear_construct simp: f_def) then have \bounded_clinear (f a b)\ for a b by auto then have F_apply: \cblinfun_apply (F a b) = f a b\ for a b by (simp add: F_def bounded_clinear_CBlinfun_apply) define basis where \basis = {F a b| a b. a\basisA \ b\basisB}\ have \cspan basis = UNIV\ apply (rule cblinfun_cspan_UNIV[where basisA=basisA and basisB=basisB]) apply (auto simp: basis_def) by (metis F_apply f_a f_not_a) moreover have \finite basis\ unfolding basis_def apply (rule finite_image_set2) by auto ultimately show \\S :: ('a \\<^sub>C\<^sub>L 'b) set. finite S \ cspan S = UNIV\ by auto qed 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))\ apply (intro continuous_on_norm linear_continuous_on) by (simp add: Complex_Vector_Spaces.bounded_clinear.bounded_linear cblinfun.bounded_clinear_right) 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 apply (rule on_closure_leI[where x=x and S=S]) using S bound 1 2 by auto then show \norm f \ b\ apply (rule_tac norm_cblinfun_bound) using \0 \ b\ by auto qed lemma dense_span_separating: \closure (cspan S) = UNIV \ bounded_clinear F \ bounded_clinear G \ (\x\S. F x = G x) \ F = G\ proof - fix F G :: \'a \ 'b\ assume dense: \closure (cspan S) = UNIV\ assume [simp]: \bounded_clinear F\ \bounded_clinear G\ assume \\x\S. F x = G x\ then have \F x = G x\ if \x \ cspan S\ for x apply (rule_tac complex_vector.linear_eq_on[of F G _ S]) using that by (auto simp: bounded_clinear.clinear) then show \F = G\ apply (rule_tac ext) apply (rule on_closure_eqI[of \cspan S\ F G]) using dense by (auto intro!: continuous_at_imp_continuous_on clinear_continuous_at) 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)\ apply (rule infsum_bounded_linear[unfolded o_def, of \cblinfun_apply A\]) using assms by (auto simp add: bounded_clinear.bounded_linear bounded_cbilinear.bounded_clinear_right bounded_cbilinear_cblinfun_apply) lemma has_sum_cblinfun_apply: assumes \has_sum g S x\ shows \has_sum (\x. A *\<^sub>V g x) S (A *\<^sub>V x)\ apply (rule has_sum_bounded_linear[unfolded o_def, of \cblinfun_apply A\]) using assms by (auto simp add: bounded_clinear.bounded_linear cblinfun.bounded_clinear_right) 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]) 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 \has_sum f I x\ shows \has_sum (\i. cinner a (f i)) I (cinner a x)\ 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 \has_sum f I x\ shows \has_sum (\i. f i \\<^sub>C a) I (x \\<^sub>C a)\ apply (rule has_sum_bounded_linear[where f=\\x. x \\<^sub>C a\, unfolded o_def]) using assms by (simp_all add: bounded_antilinear.bounded_linear bounded_antilinear_cinner_left) 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\ apply (auto simp: F1_def F2_def image_def) by force+ have \(\(x,y)\F. norm (cinner (a x) (b y))) \ (\(x,y)\F1\F2. norm (cinner (a x) (b y)))\ apply (rule sum_mono2) by auto also from pos have \\ = norm (\(x,y)\F1\F2. cinner (a x) (b y))\ apply (auto intro!: of_real_eq_iff[THEN iffD1] simp: case_prod_beta) apply (subst abs_complex_def[unfolded o_def, symmetric, THEN fun_cong])+ apply (subst (2) abs_pos) apply (rule sum_nonneg) apply (metis Compl_eq_Diff_UNIV Diff_iff SigmaE \F1 \ - X\ \F2 \ - Y\ fst_conv prod.sel(2) subsetD) apply (rule sum.cong) apply simp by (metis Compl_iff SigmaE \F1 \ - X\ \F2 \ - Y\ abs_pos fst_conv prod.sel(2) subset_eq) 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)\ apply (rule_tac summable_on_product_finite_left) by auto 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\ apply (rule_tac summable_on_product_finite_right) by auto have 4: \(\(x, y). cinner (a x) (b y)) summable_on X\Y\ by (simp add: \finite X\ \finite Y\) show ?thesis apply (subst asm_rl[of \UNIV = (-X)\(-Y) \ X\(-Y) \ (-X)\Y \ X\Y\]) using 1 2 3 4 by (auto 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.\ 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))\ apply (subst infsum_cinner_right) apply (rule asum) apply (subst infsum_cinner_left) apply (rule bsum) by simp also have \\ = (\\<^sub>\(k,l). cinner (a k) (b l))\ apply (subst infsum_Sigma'_banach) using absum by auto 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))\ apply (subst infsum_Sigma'_banach) using absum by (auto simp: sigma Cauchy_cinner_product_summable') also have \\ = (\\<^sub>\k. \i\k. a i \\<^sub>C b (k-i))\ apply (subst infsum_finite[symmetric]) by (auto 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 by blast lemma cblinfun_eq_gen_eqI: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \ccspan G = \\ shows "A = B" apply (rule cblinfun_eqI) apply (rule cblinfun_eq_on[where G=G]) using assms apply auto by (metis ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) +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 apply transfer by (simp add: scaleR_scaleC) thus "((*\<^sub>R) r::'a \\<^sub>L 'b \ _) = (*\<^sub>C) (complex_of_real r)" for r by auto show "a *\<^sub>C (x + y) = a *\<^sub>C x + a *\<^sub>C y" for a :: complex and x y :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_right) show "(a + b) *\<^sub>C x = a *\<^sub>C x + b *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_left) show "a *\<^sub>C b *\<^sub>C x = (a * b) *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by simp have \1 *\<^sub>C f x = f x\ for f :: \'a\'b\ and x by auto thus "1 *\<^sub>C x = x" for x :: "'a \\<^sub>L 'b" by (simp add: scaleC_blinfun.rep_eq blinfun_eqI) have \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ if \bounded_linear f\ for f :: \'a \ 'b\ and a :: complex proof- have \cmod a \ 0\ by simp have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \\ x. (cmod a) *(\ ereal ((norm (f x)) / (norm x)) \) \ (cmod a) * K\ using \cmod a \ 0\ by (metis abs_ereal.simps(1) abs_ereal_pos abs_pos ereal_mult_left_mono times_ereal.simps(1)) hence \\ x. (\ ereal ((cmod a) * (norm (f x)) / (norm x)) \) \ (cmod a) * K\ by simp hence \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have p1: \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ using \\ x. \ ereal (cmod a * (norm (f x)) / (norm x)) \ \ cmod a * K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) have p2: \\i. i \ UNIV \ 0 \ ereal (cmod a * norm (f i) / norm i)\ by simp hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ (SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\)\ using \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I p2 abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ cmod a * K\ using \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ by simp hence \\ ( SUP i\UNIV::'a set. ereal ((\ x. (cmod a) * (norm (f x)) / norm x) i)) \ \ \\ by auto hence w2: \( SUP i\UNIV::'a set. ereal ((\ x. cmod a * (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. cmod a * (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) have \(UNIV::('a set)) \ {}\ by simp moreover have \\ i. i \ (UNIV::('a set)) \ (\ x. (norm (f x)) / norm x :: ereal) i \ 0\ by simp moreover have \cmod a \ 0\ by simp ultimately have \(SUP i\(UNIV::('a set)). ((cmod a)::ereal) * (\ x. (norm (f x)) / norm x :: ereal) i ) = ((cmod a)::ereal) * ( SUP i\(UNIV::('a set)). (\ x. (norm (f x)) / norm x :: ereal) i )\ by (simp add: Sup_ereal_mult_left') hence \(SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp hence z1: \real_of_ereal ( (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) ) = real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) )\ by simp have z2: \real_of_ereal (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = (SUP x. cmod a * (norm (f x) / norm x))\ using w2 by auto have \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = (cmod a) * real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp moreover have \real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) ) = ( SUP x. ((norm (f x)) / norm x) )\ proof- have \\ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) \ \ \\ proof- have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ using \\ x. \ ereal ((norm (f x)) / (norm x)) \ \ K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ (SUP x. \ereal ((norm (f x)) / (norm x))\)\ using \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I \\i. i \ UNIV \ 0 \ ereal (norm (f i) / norm i)\ abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ K\ using \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ by simp thus ?thesis by auto qed hence \ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) thus ?thesis by simp qed have z3: \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = cmod a * (SUP x. norm (f x) / norm x)\ by (simp add: \real_of_ereal (SUP x. ereal (norm (f x) / norm x)) = (SUP x. norm (f x) / norm x)\) hence w1: \(SUP x. cmod a * (norm (f x) / norm x)) = cmod a * (SUP x. norm (f x) / norm x)\ using z1 z2 by linarith have v1: \onorm (\x. a *\<^sub>C f x) = (SUP x. norm (a *\<^sub>C f x) / norm x)\ by (simp add: onorm_def) have v2: \(SUP x. norm (a *\<^sub>C f x) / norm x) = (SUP x. ((cmod a) * norm (f x)) / norm x)\ by simp have v3: \(SUP x. ((cmod a) * norm (f x)) / norm x) = (SUP x. (cmod a) * ((norm (f x)) / norm x))\ by simp have v4: \(SUP x. (cmod a) * ((norm (f x)) / norm x)) = (cmod a) * (SUP x. ((norm (f x)) / norm x))\ using w1 by blast show \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ using v1 v2 v3 v4 by (metis (mono_tags, lifting) onorm_def) qed thus \norm (a *\<^sub>C x) = cmod a * norm x\ for a::complex and x::\('a, 'b) blinfun\ apply transfer by blast qed end (* We do not have clinear_blinfun_compose_right *) lemma clinear_blinfun_compose_left: \clinear (\x. blinfun_compose x y)\ by (auto intro!: clinearI simp: blinfun_eqI scaleC_blinfun.rep_eq bounded_bilinear.add_left bounded_bilinear_blinfun_compose) instantiation blinfun :: (real_normed_vector, cbanach) "cbanach" begin instance.. end lemma blinfun_compose_assoc: "(A o\<^sub>L B) o\<^sub>L C = A o\<^sub>L (B o\<^sub>L C)" by (simp add: blinfun_eqI) lift_definition blinfun_of_cblinfun::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a \\<^sub>L 'b\ is "id" apply transfer by (simp add: bounded_clinear.bounded_linear) lift_definition blinfun_cblinfun_eq :: \'a \\<^sub>L 'b \ 'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ bool\ is "(=)" . lemma blinfun_cblinfun_eq_bi_unique[transfer_rule]: \bi_unique blinfun_cblinfun_eq\ unfolding bi_unique_def apply transfer by auto lemma blinfun_cblinfun_eq_right_total[transfer_rule]: \right_total blinfun_cblinfun_eq\ unfolding right_total_def apply transfer by (simp add: bounded_clinear.bounded_linear) named_theorems cblinfun_blinfun_transfer lemma cblinfun_blinfun_transfer_0[cblinfun_blinfun_transfer]: "blinfun_cblinfun_eq (0::(_,_) blinfun) (0::(_,_) cblinfun)" apply transfer by simp lemma cblinfun_blinfun_transfer_plus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (+) (+)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_minus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (-) (-)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_uminus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (uminus) (uminus)" unfolding rel_fun_def apply transfer by auto definition "real_complex_eq r c \ complex_of_real r = c" lemma bi_unique_real_complex_eq[transfer_rule]: \bi_unique real_complex_eq\ unfolding real_complex_eq_def bi_unique_def by auto lemma left_total_real_complex_eq[transfer_rule]: \left_total real_complex_eq\ unfolding real_complex_eq_def left_total_def by auto lemma cblinfun_blinfun_transfer_scaleC[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(real_complex_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (scaleR) (scaleC)" unfolding rel_fun_def apply transfer by (simp add: real_complex_eq_def scaleR_scaleC) lemma cblinfun_blinfun_transfer_CBlinfun[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(eq_onp bounded_clinear ===> blinfun_cblinfun_eq) Blinfun CBlinfun" unfolding rel_fun_def blinfun_cblinfun_eq.rep_eq eq_onp_def by (auto simp: CBlinfun_inverse Blinfun_inverse bounded_clinear.bounded_linear) lemma cblinfun_blinfun_transfer_norm[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=)) norm norm" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_dist[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> (=)) dist dist" unfolding rel_fun_def dist_norm apply transfer by auto lemma cblinfun_blinfun_transfer_sgn[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) sgn sgn" unfolding rel_fun_def sgn_blinfun_def sgn_cblinfun_def apply transfer by (auto simp: scaleR_scaleC) lemma cblinfun_blinfun_transfer_Cauchy[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> (=)) Cauchy Cauchy" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding Cauchy_def by transfer_prover qed lemma cblinfun_blinfun_transfer_tendsto[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> blinfun_cblinfun_eq ===> (=) ===> (=)) tendsto tendsto" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding tendsto_iff by transfer_prover qed lemma cblinfun_blinfun_transfer_compose[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (o\<^sub>L) (o\<^sub>C\<^sub>L)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_apply[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=) ===> (=)) blinfun_apply cblinfun_apply" unfolding rel_fun_def apply transfer by auto lemma blinfun_of_cblinfun_inj: \blinfun_of_cblinfun f = blinfun_of_cblinfun g \ f = g\ by (metis cblinfun_apply_inject blinfun_of_cblinfun.rep_eq) lemma blinfun_of_cblinfun_inv: assumes "\c. \x. f *\<^sub>v (c *\<^sub>C x) = c *\<^sub>C (f *\<^sub>v x)" shows "\g. blinfun_of_cblinfun g = f" using assms proof transfer show "\g\Collect bounded_clinear. id g = f" if "bounded_linear f" and "\c x. f (c *\<^sub>C x) = c *\<^sub>C f x" for f :: "'a \ 'b" using that bounded_linear_bounded_clinear by auto qed lemma blinfun_of_cblinfun_zero: \blinfun_of_cblinfun 0 = 0\ apply transfer by simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ apply transfer by auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_norm: fixes f::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ shows \norm f = norm (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_cblinfun_compose: fixes f::\'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector\ and g::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b\ shows \blinfun_of_cblinfun (f o\<^sub>C\<^sub>L g) = (blinfun_of_cblinfun f) o\<^sub>L (blinfun_of_cblinfun g)\ apply transfer by auto 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)\ apply transfer by (auto intro!: ext bounded_clinear.clinear complex_vector.linear_scale) lemma cblinfun_compose_scaleR_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>R B) = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_id_right[simp]: shows "U o\<^sub>C\<^sub>L id_cblinfun = U" apply transfer by auto lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" apply transfer by auto lemma cblinfun_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)\ apply (induction S rule:infinite_finite_induct) by (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)\ apply (induction S rule:infinite_finite_induct) by (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) subsection \Adjoint\ lift_definition adj :: "'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ 'b \\<^sub>C\<^sub>L 'a" ("_*" [99] 100) is cadjoint by (fact cadjoint_bounded_clinear) lemma id_cblinfun_adjoint[simp]: "id_cblinfun* = id_cblinfun" apply transfer using cadjoint_id by (metis eq_id_iff) lemma double_adj[simp]: "(A*)* = A" apply transfer using double_cadjoint by blast lemma adj_cblinfun_compose[simp]: fixes B::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and A::\'b \\<^sub>C\<^sub>L 'c::complex_inner\ shows "(A o\<^sub>C\<^sub>L B)* = (B*) o\<^sub>C\<^sub>L (A*)" proof transfer fix A :: \'b \ 'c\ and B :: \'a \ 'b\ assume \bounded_clinear A\ and \bounded_clinear B\ hence \bounded_clinear (A \ B)\ by (simp add: comp_bounded_clinear) have \((A \ B) u \\<^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*)" apply transfer by (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)\ apply intro_classes by simp from True have c2: \class.not_singleton TYPE('c)\ apply intro_classes by simp have normA: \norm A = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c1) have normAadj: \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A* *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c2) have \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ unfolding normAadj apply (subst cinner_adj_right) apply (subst cinner_commute) apply (subst complex_mod_cnj) by rule also have \\ = Sup ((\(\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \)) ` prod.swap ` UNIV)\ by auto also have \\ = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (subst image_image) by auto also have \\ = norm A\ unfolding normA by (simp add: mult.commute) finally show ?thesis by - next case False then consider (b) \\x::'b. x = 0\ | (c) \\x::'c. x = 0\ by auto then have \A = 0\ apply (cases; transfer) apply (metis (full_types) bounded_clinear_def complex_vector.linear_0) by auto then show \norm (A*) = norm A\ by simp qed lemma antilinear_adj[simp]: \antilinear adj\ apply (rule antilinearI) by (auto simp add: adj_plus) lemma bounded_antilinear_adj[bounded_antilinear, simp]: \bounded_antilinear adj\ by (auto intro!: antilinearI exI[of _ 1] simp: bounded_antilinear_def bounded_antilinear_axioms_def adj_plus) lemma adjoint_eqI: fixes G:: \'b::chilbert_space \\<^sub>C\<^sub>L 'a::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*)\ apply (rule adjoint_eqI[symmetric]) by (simp add: cblinfun.minus_left cinner_adj_left) lemma cinner_real_hermiteanI: \ \Prop. II.2.12 in @{cite conway2013course}\ assumes \\\. cinner \ (A *\<^sub>V \) \ \\ shows \A = A*\ proof - { fix g h :: 'a { fix \ :: complex have \cinner h (A h) + cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) + (abs \)\<^sup>2 * cinner g (A g) = cinner (h + \ *\<^sub>C g) (A *\<^sub>V (h + \ *\<^sub>C g))\ (is \?sum4 = _\) apply (auto simp: cinner_add_right cinner_add_left cblinfun.add_right cblinfun.scaleC_right ring_class.ring_distribs) by (metis cnj_x_x mult.commute) also have \\ \ \\ using assms by auto finally have \?sum4 = cnj ?sum4\ using Reals_cnj_iff by fastforce then have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner (A h) g + cnj \ *\<^sub>C cinner (A g) h\ using Reals_cnj_iff abs_complex_real assms by force also have \\ = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by (simp add: cinner_adj_right) finally have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by - } from this[where \2=1] this[where \2=\] have 1: \cinner g (A h) + cinner h (A g) = cinner h (A* *\<^sub>V g) + cinner g (A* *\<^sub>V h)\ and i: \- \ * cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) - \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by auto from arg_cong2[OF 1 arg_cong[OF i, where f=\(*) (-\)\], where f=plus] have \cinner h (A g) = cinner h (A* *\<^sub>V g)\ by (auto simp: ring_class.ring_distribs) } then show "A = A*" by (simp add: adjoint_eqI cinner_adj_right) qed lemma norm_AAadj[simp]: \norm (A o\<^sub>C\<^sub>L A*) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof (cases \class.not_singleton TYPE('b)\) case True then have [simp]: \class.not_singleton TYPE('b)\ by - have 1: \(norm A)\<^sup>2 * \ \ norm (A o\<^sub>C\<^sub>L A*)\ if \\ < 1\ and \\ \ 0\ for \ proof - obtain \ where \: \norm ((A*) *\<^sub>V \) \ norm (A*) * sqrt \\ and [simp]: \norm \ = 1\ apply atomize_elim apply (rule cblinfun_norm_approx_witness_mult[internalize_sort' 'a]) using \\ < 1\ by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) have \complex_of_real ((norm A)\<^sup>2 * \) = (norm (A*) * sqrt \)\<^sup>2\ by (simp add: ordered_field_class.sign_simps(23) that(2)) also have \\ \ (norm ((A* *\<^sub>V \)))\<^sup>2\ apply (rule complex_of_real_mono) using \ apply (rule power_mono) using \\ \ 0\ by auto also have \\ \ cinner (A* *\<^sub>V \) (A* *\<^sub>V \)\ by (auto simp flip: power2_norm_eq_cinner) also have \\ = cinner \ (A *\<^sub>V A* *\<^sub>V \)\ by (simp add: cinner_adj_left) also have \\ = cinner \ ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ by auto also have \\ \ norm (A o\<^sub>C\<^sub>L A*)\ using \norm \ = 1\ by (smt (verit, best) Im_complex_of_real Re_complex_of_real \(A* *\<^sub>V \) \\<^sub>C (A* *\<^sub>V \) = \ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \)\ \\ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \) = \ \\<^sub>C ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ cdot_square_norm cinner_ge_zero cmod_Re complex_inner_class.Cauchy_Schwarz_ineq2 less_eq_complex_def mult_cancel_left1 mult_cancel_right1 norm_cblinfun) finally show ?thesis by (auto simp: less_eq_complex_def) qed then have 1: \(norm A)\<^sup>2 \ norm (A o\<^sub>C\<^sub>L A*)\ by (metis field_le_mult_one_interval less_eq_real_def ordered_field_class.sign_simps(5)) have 2: \norm (A o\<^sub>C\<^sub>L A*) \ (norm A)\<^sup>2\ proof (rule norm_cblinfun_bound) show \0 \ (norm A)\<^sup>2\ by simp fix \ have \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) = norm (A *\<^sub>V A* *\<^sub>V \)\ by auto also have \\ \ norm A * norm (A* *\<^sub>V \)\ by (simp add: norm_cblinfun) also have \\ \ norm A * norm (A*) * norm \\ by (metis mult.assoc norm_cblinfun norm_imp_pos_and_ge ordered_comm_semiring_class.comm_mult_left_mono) also have \\ = (norm A)\<^sup>2 * norm \\ by (simp add: power2_eq_square) finally show \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) \ (norm A)\<^sup>2 * norm \\ by - qed from 1 2 show ?thesis by simp next case False then have [simp]: \class.CARD_1 TYPE('b)\ by (rule not_singleton_vs_CARD_1) have \A = 0\ apply (rule cblinfun_to_CARD_1_0[internalize_sort' 'b]) by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) then show ?thesis by auto qed lemma sum_adj: \(sum a F)* = sum (\i. (a i)*) F\ apply (induction rule:infinite_finite_induct) by (auto simp add: adj_plus) lemma has_sum_adj: assumes \has_sum f I x\ shows \has_sum (\x. adj (f x)) I (adj x)\ 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) 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\ apply transfer by auto lemma cblinfun_power_Suc': \cblinfun_power A (Suc n) = A o\<^sub>C\<^sub>L cblinfun_power A n\ apply transfer by 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\]) apply simp by (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 unitary_twosided_isometry: "unitary U \ isometry U \ isometry (U*)" unfolding unitary_def isometry_def by simp lemma isometryD[simp]: "isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding isometry_def by simp (* Not [simp] because isometryD[simp] + unitary_isometry[simp] already have the same effect *) lemma unitaryD1: "unitary U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding unitary_def by simp lemma unitaryD2[simp]: "unitary U \ U o\<^sub>C\<^sub>L U* = id_cblinfun" unfolding unitary_def by simp lemma unitary_isometry[simp]: "unitary U \ isometry U" unfolding unitary_def isometry_def by simp lemma unitary_adj[simp]: "unitary (U*) = unitary U" unfolding unitary_def by auto lemma isometry_cblinfun_compose[simp]: assumes "isometry A" and "isometry B" shows "isometry (A o\<^sub>C\<^sub>L B)" proof- have "B* o\<^sub>C\<^sub>L A* o\<^sub>C\<^sub>L (A o\<^sub>C\<^sub>L B) = id_cblinfun" if "A* o\<^sub>C\<^sub>L A = id_cblinfun" and "B* o\<^sub>C\<^sub>L B = id_cblinfun" using that by (smt (verit, del_insts) adjoint_eqI cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply) thus ?thesis using assms unfolding isometry_def by simp qed lemma unitary_cblinfun_compose[simp]: "unitary (A o\<^sub>C\<^sub>L B)" if "unitary A" and "unitary B" using that by (smt (z3) adj_cblinfun_compose cblinfun_compose_assoc cblinfun_compose_id_right double_adj isometryD isometry_cblinfun_compose unitary_def unitary_isometry) lemma unitary_surj: assumes "unitary U" shows "surj (cblinfun_apply U)" apply (rule surjI[where f=\cblinfun_apply (U*)\]) using assms unfolding unitary_def apply transfer using comp_eq_dest_lhs by force lemma unitary_id[simp]: "unitary id_cblinfun" by (simp add: unitary_def) lemma orthogonal_on_basis_is_isometry: assumes spanB: \ccspan B = \\ assumes orthoU: \\b c. b\B \ c\B \ cinner (U *\<^sub>V b) (U *\<^sub>V c) = cinner b c\ shows \isometry U\ proof - have [simp]: \b \ closure (cspan B)\ for b using spanB apply transfer by simp have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ and \\\B\ for \ \ by (simp add: cinner_adj_left orthoU that(1) that(2)) have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ for \ \ apply (rule bounded_clinear_eq_on[where t=\ and G=B]) using bounded_clinear_cinner_right *[OF that] by auto have \U* *\<^sub>V U *\<^sub>V \ = \\ if \\\B\ for \ apply (rule cinner_extensionality) apply (subst cinner_eq_flip) by (simp add: * that) then have \U* o\<^sub>C\<^sub>L U = id_cblinfun\ by (metis cblinfun_apply_cblinfun_compose cblinfun_eq_gen_eqI cblinfun_id_cblinfun_apply spanB) then show \isometry U\ using isometry_def by blast qed 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) 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 apply transfer by simp lemma isometry_cblinfun_right[simp]: \isometry cblinfun_right\ apply (rule orthogonal_on_basis_is_isometry[of some_chilbert_basis]) apply simp apply transfer by 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) apply transfer by 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) apply transfer by 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)\ apply transfer by 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,\)\ apply transfer by 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\ apply transfer by simp 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 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 apply transfer by (simp add: bounded_clinear_def complex_vector.linear_0) lemma cblinfun_image_bot[simp]: "U *\<^sub>S bot = bot" using cblinfun_image_0 by auto lemma cblinfun_image_sup[simp]: fixes A B :: \'a::chilbert_space ccsubspace\ and U :: "'a \\<^sub>C\<^sub>L'b::chilbert_space" shows \U *\<^sub>S (sup A B) = sup (U *\<^sub>S A) (U *\<^sub>S B)\ apply transfer using bounded_clinear.bounded_linear closure_image_closed_sum by blast lemma scaleC_cblinfun_image[simp]: fixes A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b :: chilbert_space\ and S :: \'a ccsubspace\ and \ :: complex shows \(\ *\<^sub>C A) *\<^sub>S S = \ *\<^sub>C (A *\<^sub>S S)\ proof- have \closure ( ( ((*\<^sub>C) \) \ (cblinfun_apply A) ) ` space_as_set S) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis closure_scaleC image_comp) hence \(closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis (mono_tags, lifting) comp_apply image_cong scaleC_cblinfun.rep_eq) hence \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by (metis space_as_set_inverse cblinfun_image.rep_eq scaleC_ccsubspace.rep_eq) have x1: "Abs_clinear_space (closure ((*\<^sub>V) (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure ((*\<^sub>V) A ` space_as_set S))" using \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by blast show ?thesis unfolding cblinfun_image_def using x1 by force qed lemma cblinfun_image_id[simp]: "id_cblinfun *\<^sub>S \ = \" apply transfer by (simp add: closed_csubspace.closed) lemma cblinfun_compose_image: \(A o\<^sub>C\<^sub>L B) *\<^sub>S S = A *\<^sub>S (B *\<^sub>S S)\ apply transfer unfolding image_comp[symmetric] apply (rule closure_bounded_linear_image_subset_eq[symmetric]) by (simp add: bounded_clinear.bounded_linear) lemmas cblinfun_assoc_left = cblinfun_compose_assoc[symmetric] cblinfun_compose_image[symmetric] add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space", symmetric] lemmas cblinfun_assoc_right = cblinfun_compose_assoc cblinfun_compose_image add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space"] lemma cblinfun_image_INF_leq[simp]: fixes U :: "'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::cbanach" and V :: "'a \ 'b ccsubspace" shows \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ apply transfer by (simp add: INT_greatest Inter_lower closure_mono image_mono) lemma isometry_cblinfun_image_inf_distrib': fixes U::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::cbanach\ and B C::"'a ccsubspace" shows "U *\<^sub>S (inf B C) \ inf (U *\<^sub>S B) (U *\<^sub>S C)" proof - define V where \V b = (if b then B else C)\ for b have \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ by auto then show ?thesis unfolding V_def by (metis (mono_tags, lifting) INF_UNIV_bool_expand) qed lemma cblinfun_image_eq: fixes S :: "'a::cbanach ccsubspace" and A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::cbanach" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and "ccspan G \ S" shows "A *\<^sub>S S = B *\<^sub>S S" proof (use assms in transfer) fix G :: "'a set" and A :: "'a \ 'b" and B :: "'a \ 'b" and S :: "'a set" assume a1: "bounded_clinear A" assume a2: "bounded_clinear B" assume a3: "\x. x \ G \ A x = B x" assume a4: "S \ closure (cspan G)" have "A ` closure S = B ` closure S" by (smt (verit, best) UnCI a1 a2 a3 a4 bounded_clinear_eq_on closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis bounded_clinear.bounded_linear a1 a2 closure_bounded_linear_image_subset_eq) qed lemma cblinfun_fixes_range: assumes "A o\<^sub>C\<^sub>L B = B" and "\ \ space_as_set (B *\<^sub>S top)" shows "A *\<^sub>V \ = \" proof- define rangeB rangeB' where "rangeB = space_as_set (B *\<^sub>S top)" and "rangeB' = range (cblinfun_apply B)" from assms have "\ \ closure rangeB'" by (simp add: cblinfun_image.rep_eq rangeB'_def top_ccsubspace.rep_eq) then obtain \i where \i_lim: "\i \ \" and \i_B: "\i i \ rangeB'" for i using closure_sequential by blast have A_invariant: "A *\<^sub>V \i i = \i i" for i proof- from \i_B obtain \ where \: "\i i = B *\<^sub>V \" using rangeB'_def by blast hence "A *\<^sub>V \i i = (A o\<^sub>C\<^sub>L B) *\<^sub>V \" by (simp add: cblinfun_compose.rep_eq) also have "\ = B *\<^sub>V \" by (simp add: assms) also have "\ = \i i" by (simp add: \) finally show ?thesis. qed from \i_lim have "(\i. A *\<^sub>V (\i i)) \ A *\<^sub>V \" by (rule isCont_tendsto_compose[rotated], simp) with A_invariant have "(\i. \i i) \ A *\<^sub>V \" by auto with \i_lim show "A *\<^sub>V \ = \" using LIMSEQ_unique by blast qed lemma zero_cblinfun_image[simp]: "0 *\<^sub>S S = (0::_ ccsubspace)" apply transfer by (simp add: complex_vector.subspace_0 image_constant[where x=0]) lemma cblinfun_image_INF_eq_general: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L'c::chilbert_space" and Uinv :: "'c \\<^sub>C\<^sub>L'b" assumes UinvUUinv: "Uinv o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L Uinv = Uinv" and UUinvU: "U o\<^sub>C\<^sub>L Uinv o\<^sub>C\<^sub>L U = U" \ \Meaning: \<^term>\Uinv\ is a Pseudoinverse of \<^term>\U\\ and V: "\i. V i \ Uinv *\<^sub>S top" shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof (rule antisym) show "U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S V i)" by (rule cblinfun_image_INF_leq) next define rangeU rangeUinv where "rangeU = U *\<^sub>S top" and "rangeUinv = Uinv *\<^sub>S top" define INFUV INFV where INFUV_def: "INFUV = (INF i. U *\<^sub>S V i)" and INFV_def: "INFV = (INF i. V i)" from assms have "V i \ rangeUinv" for i unfolding rangeUinv_def by simp moreover have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set rangeUinv" for \ using UinvUUinv cblinfun_fixes_range rangeUinv_def that by fastforce ultimately have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set (V i)" for \ i using less_eq_ccsubspace.rep_eq that by blast hence d1: "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>S (V i) = (V i)" for i proof transfer show "closure ((Uinv \ U) ` V i) = V i" if "pred_fun \ closed_csubspace V" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ (Uinv \ U) \ = \" for V :: "'a \ 'b set" and Uinv :: "'c \ 'b" and U :: "'b \ 'c" and i :: 'a using that proof auto show "x \ V i" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ closure (V i)" for x :: 'b using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure (V i)" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ V i" for x :: 'b using that using setdist_eq_0_sing_1 setdist_sing_in_set by blast qed qed have "U *\<^sub>S V i \ rangeU" for i by (simp add: cblinfun_image_mono rangeU_def) hence "INFUV \ rangeU" unfolding INFUV_def by (meson INF_lower UNIV_I order_trans) moreover have "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set rangeU" for \ using UUinvU cblinfun_fixes_range rangeU_def that by fastforce ultimately have x: "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set INFUV" for \ by (simp add: in_mono less_eq_ccsubspace.rep_eq that) have "closure ((U \ Uinv) ` INFUV) = INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ (U \ Uinv) \ = \" for INFUV :: "'c set" and U :: "'b \ 'c" and Uinv :: "'c \ 'b" using that proof auto show "x \ INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ closure INFUV" for x :: 'c using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ INFUV" for x :: 'c using that using setdist_eq_0_sing_1 setdist_sing_in_set by (simp add: closed_csubspace.closed) qed hence "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>S INFUV = INFUV" by (metis (mono_tags, opaque_lifting) x cblinfun_image.rep_eq cblinfun_image_id id_cblinfun_apply image_cong space_as_set_inject) hence "INFUV = U *\<^sub>S Uinv *\<^sub>S INFUV" by (simp add: cblinfun_compose_image) also have "\ \ U *\<^sub>S (INF i. Uinv *\<^sub>S U *\<^sub>S V i)" unfolding INFUV_def by (metis cblinfun_image_mono cblinfun_image_INF_leq) also have "\ = U *\<^sub>S INFV" using d1 by (metis (no_types, lifting) INFV_def cblinfun_assoc_left(2) image_cong) finally show "INFUV \ U *\<^sub>S INFV". qed lemma unitary_range[simp]: assumes "unitary U" shows "U *\<^sub>S top = top" using assms unfolding unitary_def apply transfer by (metis closure_UNIV comp_apply surj_def) lemma range_adjoint_isometry: assumes "isometry U" shows "U* *\<^sub>S top = top" proof- from assms have "top = U* *\<^sub>S U *\<^sub>S top" by (simp add: cblinfun_assoc_left(2)) also have "\ \ U* *\<^sub>S top" by (simp add: cblinfun_image_mono) finally show ?thesis using top.extremum_unique by blast qed lemma cblinfun_image_INF_eq[simp]: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L 'c::chilbert_space" assumes \isometry U\ shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof - from \isometry U\ have "U* o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L U* = U*" unfolding isometry_def by simp moreover from \isometry U\ have "U o\<^sub>C\<^sub>L U* o\<^sub>C\<^sub>L U = U" unfolding isometry_def by (simp add: cblinfun_compose_assoc) moreover have "V i \ U* *\<^sub>S top" for i by (simp add: range_adjoint_isometry assms) ultimately show ?thesis by (rule cblinfun_image_INF_eq_general) qed lemma isometry_cblinfun_image_inf_distrib[simp]: fixes U::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and X Y::"'a ccsubspace" assumes "isometry U" shows "U *\<^sub>S (inf X Y) = inf (U *\<^sub>S X) (U *\<^sub>S Y)" using cblinfun_image_INF_eq[where V="\b. if b then X else Y" and U=U] unfolding INF_UNIV_bool_expand using assms by auto lemma cblinfun_image_ccspan: shows "A *\<^sub>S ccspan G = ccspan ((*\<^sub>V) A ` G)" apply transfer by (simp add: bounded_clinear.bounded_linear bounded_clinear_def closure_bounded_linear_image_subset_eq complex_vector.linear_span_image) lemma cblinfun_apply_in_image[simp]: "A *\<^sub>V \ \ space_as_set (A *\<^sub>S \)" by (metis cblinfun_image.rep_eq closure_subset in_mono range_eqI top_ccsubspace.rep_eq) lemma cblinfun_plus_image_distr: \(A + B) *\<^sub>S S \ A *\<^sub>S S \ B *\<^sub>S S\ apply transfer by (smt (verit, ccfv_threshold) closed_closure closed_sum_def closure_minimal closure_subset image_subset_iff set_plus_intro subset_eq) lemma cblinfun_sum_image_distr: \(\i\I. A i) *\<^sub>S S \ (SUP i\I. A i *\<^sub>S S)\ proof (cases \finite I\) case True then show ?thesis proof induction case empty then show ?case by auto next case (insert x F) then show ?case apply auto by (smt (z3) cblinfun_plus_image_distr inf_sup_aci(6) le_iff_sup) qed next case False then show ?thesis by auto qed 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: 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) subsection \Sandwiches\ lift_definition sandwich :: \('a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner) \ (('a \\<^sub>C\<^sub>L 'a) \\<^sub>C\<^sub>L ('b \\<^sub>C\<^sub>L 'b))\ is \\(A::'a\\<^sub>C\<^sub>L'b) B. A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ proof fix A :: \'a \\<^sub>C\<^sub>L 'b\ and B B1 B2 :: \'a \\<^sub>C\<^sub>L 'a\ and c :: complex show \A o\<^sub>C\<^sub>L (B1 + B2) o\<^sub>C\<^sub>L A* = (A o\<^sub>C\<^sub>L B1 o\<^sub>C\<^sub>L A*) + (A o\<^sub>C\<^sub>L B2 o\<^sub>C\<^sub>L A*)\ by (simp add: cblinfun_compose_add_left cblinfun_compose_add_right) show \A o\<^sub>C\<^sub>L (c *\<^sub>C B) o\<^sub>C\<^sub>L A* = c *\<^sub>C (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*)\ by auto show \\K. \B. norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * K\ proof (rule exI[of _ \norm A * norm (A*)\], rule allI) fix B have \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by blast also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) finally show \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * (norm A * norm (A*))\ by (simp add: mult.assoc vector_space_over_itself.scale_left_commute) qed qed lemma sandwich_0[simp]: \sandwich 0 = 0\ by (simp add: cblinfun_eqI sandwich.rep_eq) lemma sandwich_apply: \sandwich A *\<^sub>V B = A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ apply (transfer fixing: A B) by auto lemma norm_sandwich: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'a::{chilbert_space} \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof - have main: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'c::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'd::{complex_inner}\ proof (rule norm_cblinfun_eqI) show \(norm A)\<^sup>2 \ norm (sandwich A *\<^sub>V id_cblinfun) / norm (id_cblinfun :: 'c \\<^sub>C\<^sub>L _)\ apply (auto simp: sandwich_apply) by - fix B have \norm (sandwich A *\<^sub>V B) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by (auto simp: sandwich_apply simp del: norm_adj) also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) also have \\ \ (norm A)\<^sup>2 * norm B\ by (simp add: power2_eq_square mult.assoc vector_space_over_itself.scale_left_commute) finally show \norm (sandwich A *\<^sub>V B) \ (norm A)\<^sup>2 * norm B\ by - show \0 \ (norm A)\<^sup>2\ by auto qed show ?thesis proof (cases \class.not_singleton TYPE('a)\) case True show ?thesis apply (rule main[internalize_sort' 'c2]) apply standard[1] using True by simp next case False have \A = 0\ apply (rule cblinfun_from_CARD_1_0[internalize_sort' 'a]) apply (rule not_singleton_vs_CARD_1) apply (rule False) by standard then show ?thesis by simp qed qed lemma sandwich_apply_adj: \sandwich A (B*) = (sandwich A B)*\ by (simp add: cblinfun_assoc_left(1) sandwich_apply) lemma sandwich_id[simp]: "sandwich id_cblinfun = id_cblinfun" apply (rule cblinfun_eqI) by (auto simp: sandwich_apply) subsection \Projectors\ lift_definition Proj :: "('a::chilbert_space) ccsubspace \ 'a \\<^sub>C\<^sub>L'a" is \projection\ by (rule projection_bounded_clinear) lemma Proj_range[simp]: "Proj S *\<^sub>S top = S" proof transfer fix S :: \'a set\ assume \closed_csubspace S\ then have "closure (range (projection S)) \ S" by (metis closed_csubspace.closed closed_csubspace.subspace closure_closed complex_vector.subspace_0 csubspace_is_convex dual_order.eq_iff insert_absorb insert_not_empty projection_image) moreover have "S \ closure (range (projection S))" using \closed_csubspace S\ by (metis closed_csubspace_def closure_subset csubspace_is_convex equals0D projection_image subset_iff) ultimately show \closure (range (projection S)) = S\ by auto qed lemma adj_Proj: \(Proj M)* = Proj M\ apply transfer by (simp add: projection_cadjoint) lemma Proj_idempotent[simp]: \Proj M o\<^sub>C\<^sub>L Proj M = Proj M\ proof - have u1: \(cblinfun_apply (Proj M)) = projection (space_as_set M)\ apply transfer by blast have \closed_csubspace (space_as_set M)\ using space_as_set by auto hence u2: \(projection (space_as_set M))\(projection (space_as_set M)) = (projection (space_as_set M))\ using projection_idem by fastforce have \(cblinfun_apply (Proj M)) \ (cblinfun_apply (Proj M)) = cblinfun_apply (Proj M)\ using u1 u2 by simp hence \cblinfun_apply ((Proj M) o\<^sub>C\<^sub>L (Proj M)) = cblinfun_apply (Proj M)\ by (simp add: cblinfun_compose.rep_eq) thus ?thesis using cblinfun_apply_inject by auto qed lift_definition is_Proj :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'a \ bool\ is \\P. \M. is_projection_on P M\ . 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\] apply transfer by (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\ apply transfer by 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 fastforce qed have x1: \x \ (space_as_set B +\<^sub>M space_as_set C)\ if "x \ space_as_set A" for x proof - have f1: "x \ closure {a + b |a b. a \ space_as_set B \ b \ space_as_set C}" by (simp add: q1 that) have "{a + b |a b. a \ space_as_set B \ b \ space_as_set C} = {a. \p. p \ space_as_set B \ (\q. q \ space_as_set C \ a = p + q)}" by blast hence "x \ closure {a. \b\space_as_set B. \c\space_as_set C. a = b + c}" using f1 by (simp add: Bex_def_raw) thus ?thesis using that unfolding closed_sum_def set_plus_def by blast qed hence \x \ space_as_set (Abs_clinear_space (space_as_set B +\<^sub>M space_as_set C))\ if "x \ space_as_set A" for x using that by (metis space_as_set_inverse sup_ccsubspace.rep_eq) thus ?thesis by (simp add: x1 less_eq_ccsubspace.rep_eq subset_eq sup_ccsubspace.rep_eq) qed lemma is_Proj_idempotent: assumes "is_Proj P" shows "P o\<^sub>C\<^sub>L P = P" using assms 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\ apply transfer by (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 apply transfer by (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 apply blast by (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 subsection \Kernel / eigenspaces\ lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector \ 'a ccsubspace" is "\ f. f -` {0}" by (metis kernel_is_closed_csubspace) definition eigenspace :: "complex \ 'a::complex_normed_vector \\<^sub>C\<^sub>L'a \ 'a ccsubspace" where "eigenspace a A = kernel (A - a *\<^sub>C id_cblinfun)" lemma kernel_scaleC[simp]: "a\0 \ kernel (a *\<^sub>C A) = kernel A" for a :: complex and A :: "(_,_) cblinfun" apply transfer using complex_vector.scale_eq_0_iff by blast lemma kernel_0[simp]: "kernel 0 = top" apply transfer by auto lemma kernel_id[simp]: "kernel id_cblinfun = 0" apply transfer by simp lemma eigenspace_scaleC[simp]: assumes a1: "a \ 0" shows "eigenspace b (a *\<^sub>C A) = eigenspace (b/a) A" proof - have "b *\<^sub>C (id_cblinfun::('a, _) cblinfun) = a *\<^sub>C (b / a) *\<^sub>C id_cblinfun" using a1 by (metis ceq_vector_fraction_iff) hence "kernel (a *\<^sub>C A - b *\<^sub>C id_cblinfun) = kernel (A - (b / a) *\<^sub>C id_cblinfun)" using a1 by (metis (no_types) complex_vector.scale_right_diff_distrib kernel_scaleC) thus ?thesis unfolding eigenspace_def by blast qed lemma eigenspace_memberD: assumes "x \ space_as_set (eigenspace e A)" shows "A *\<^sub>V x = e *\<^sub>C x" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms apply transfer by auto lemma eigenspace_memberI: assumes "A *\<^sub>V x = e *\<^sub>C x" shows "x \ space_as_set (eigenspace e A)" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms apply transfer by auto 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) apply simp by (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\ apply transfer by 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))\ apply transfer by (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 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) apply assumption by (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 cblinfun_inv :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ ('b,'a) cblinfun\ where \cblinfun_inv A = (SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun)\ lemma assumes \iso_cblinfun A\ shows cblinfun_inv_left: \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and cblinfun_inv_right: \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ proof - from assms obtain B where AB: \A o\<^sub>C\<^sub>L B = id_cblinfun\ and BA: \B o\<^sub>C\<^sub>L A = id_cblinfun\ using iso_cblinfun_def by blast from BA have \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ by (metis (mono_tags, lifting) cblinfun_inv_def someI_ex) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB BA show \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ by auto qed lemma cblinfun_inv_uniq: assumes "A o\<^sub>C\<^sub>L B = id_cblinfun" and "B o\<^sub>C\<^sub>L A = id_cblinfun" shows "cblinfun_inv A = B" using assms by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) subsection \One-dimensional spaces\ instantiation cblinfun :: (one_dim, one_dim) complex_inner begin text \Once we have a theory for the trace, we could instead define the Hilbert-Schmidt inner product and relax the \<^class>\one_dim\-sort constraint to (\<^class>\cfinite_dim\,\<^class>\complex_normed_vector\) or similar\ definition "cinner_cblinfun (A::'a \\<^sub>C\<^sub>L 'b) (B::'a \\<^sub>C\<^sub>L 'b) = cnj (one_dim_iso (A *\<^sub>V 1)) * one_dim_iso (B *\<^sub>V 1)" instance proof intro_classes fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "(A \\<^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 apply transfer by (simp add: norm_mult abs_complex_def one_dim_onorm' cnj_x_x power2_eq_square bounded_clinear.clinear) qed end instantiation cblinfun :: (one_dim, one_dim) one_dim begin lift_definition one_cblinfun :: "'a \\<^sub>C\<^sub>L 'b" is "one_dim_iso" by (rule bounded_clinear_one_dim_iso) lift_definition times_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f g. f o one_dim_iso o g" by (simp add: comp_bounded_clinear) lift_definition inverse_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f. ((*) (one_dim_iso (inverse (f 1)))) o one_dim_iso" by (auto intro!: comp_bounded_clinear bounded_clinear_mult_right) definition divide_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" where "divide_cblinfun A B = A * inverse B" definition "canonical_basis_cblinfun = [1 :: 'a \\<^sub>C\<^sub>L 'b]" instance proof intro_classes let ?basis = "canonical_basis :: ('a \\<^sub>C\<^sub>L 'b) list" fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "distinct ?basis" unfolding canonical_basis_cblinfun_def by simp have "(1::'a \\<^sub>C\<^sub>L 'b) \ (0::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_one zero_neq_one) thus "cindependent (set ?basis)" unfolding canonical_basis_cblinfun_def by simp have "A \ cspan (set ?basis)" for A proof - define c :: complex where "c = one_dim_iso (A *\<^sub>V 1)" have "A x = one_dim_iso (A 1) *\<^sub>C one_dim_iso x" for x by (smt (z3) cblinfun.scaleC_right complex_vector.scale_left_commute one_dim_iso_idem one_dim_scaleC_1) hence "A = one_dim_iso (A *\<^sub>V 1) *\<^sub>C 1" apply transfer by metis thus "A \ cspan (set ?basis)" unfolding canonical_basis_cblinfun_def by (smt complex_vector.span_base complex_vector.span_scale list.set_intros(1)) qed thus "cspan (set ?basis) = UNIV" by auto have "A = (1::'a \\<^sub>C\<^sub>L 'b) \ norm (1::'a \\<^sub>C\<^sub>L 'b) = (1::real)" apply transfer by simp thus "A \ set ?basis \ norm A = 1" unfolding canonical_basis_cblinfun_def by simp show "?basis = [1]" unfolding canonical_basis_cblinfun_def by simp show "c *\<^sub>C 1 * c' *\<^sub>C 1 = (c * c') *\<^sub>C (1::'a\\<^sub>C\<^sub>L'b)" apply transfer by auto have "(1::'a \\<^sub>C\<^sub>L 'b) = (0::'a \\<^sub>C\<^sub>L 'b) \ False" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_zero' zero_neq_neg_one) thus "is_ortho_set (set ?basis)" unfolding is_ortho_set_def canonical_basis_cblinfun_def by auto show "A div B = A * inverse B" by (simp add: divide_cblinfun_def) show "inverse (c *\<^sub>C 1) = (1::'a\\<^sub>C\<^sub>L'b) /\<^sub>C c" apply transfer by (simp add: o_def one_dim_inverse) qed end lemma id_cblinfun_eq_1[simp]: \id_cblinfun = 1\ apply transfer by auto lemma one_dim_apply_is_times[simp]: fixes A :: "'a::one_dim \\<^sub>C\<^sub>L 'a" and B :: "'a \\<^sub>C\<^sub>L 'a" shows "A o\<^sub>C\<^sub>L B = A * B" apply transfer by simp lemma one_comp_one_cblinfun[simp]: "1 o\<^sub>C\<^sub>L 1 = 1" apply transfer unfolding o_def by simp lemma one_cblinfun_adj[simp]: "1* = 1" apply transfer by simp lemma scaleC_1_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\ apply transfer by auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ apply transfer by auto lemma one_dim_iso_id_cblinfun: \one_dim_iso id_cblinfun = id_cblinfun\ by simp lemma one_dim_iso_id_cblinfun_eq_1: \one_dim_iso id_cblinfun = 1\ by simp lemma one_dim_iso_comp_distr[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a o\<^sub>C\<^sub>L one_dim_iso b\ by (smt (z3) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC) lemma one_dim_iso_comp_distr_times[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a * one_dim_iso b\ by (smt (verit, del_insts) mult.left_neutral mult_scaleC_left one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left) lemma one_dim_iso_adjoint[simp]: \one_dim_iso (A*) = (one_dim_iso A)*\ by (smt (z3) one_cblinfun_adj one_cinner_a_scaleC_one one_dim_iso_of_one one_dim_iso_scaleC scaleC_adj) lemma one_dim_iso_adjoint_complex[simp]: \one_dim_iso (A*) = cnj (one_dim_iso A)\ by (metis (mono_tags, lifting) one_cblinfun_adj one_dim_iso_idem one_dim_scaleC_1 scaleC_adj) lemma one_dim_cblinfun_compose_commute: \a o\<^sub>C\<^sub>L b = b o\<^sub>C\<^sub>L a\ for a b :: \('a::one_dim,'a) cblinfun\ by (simp add: one_dim_iso_inj) lemma one_cblinfun_apply_one[simp]: \1 *\<^sub>V 1 = 1\ by (simp add: one_cblinfun.rep_eq) 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" apply transfer by auto definition "heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'b::chilbert_space itself) \ unitary (heterogenous_cblinfun_id :: 'a \\<^sub>C\<^sub>L 'b) \ unitary (heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a)" lemma heterogenous_same_type_cblinfun[simp]: \heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'a::chilbert_space itself)\ unfolding heterogenous_same_type_cblinfun_def by auto instantiation cblinfun :: (chilbert_space, chilbert_space) ord begin definition less_eq_cblinfun :: \('a \\<^sub>C\<^sub>L 'b) \ ('a \\<^sub>C\<^sub>L 'b) \ bool\ where less_eq_cblinfun_def_heterogenous: \less_eq_cblinfun A B = (if heterogenous_same_type_cblinfun TYPE('a) TYPE('b) then \\::'b. cinner \ ((B-A) *\<^sub>V heterogenous_cblinfun_id *\<^sub>V \) \ 0 else (A=B))\ definition \less_cblinfun (A :: 'a \\<^sub>C\<^sub>L 'b) B \ A \ B \ \ B \ A\ instance.. end lemma less_eq_cblinfun_def: \A \ B \ (\\. cinner \ (A *\<^sub>V \) \ cinner \ (B *\<^sub>V \))\ unfolding less_eq_cblinfun_def_heterogenous by (auto simp del: less_eq_complex_def simp: cblinfun.diff_left cinner_diff_right) instantiation cblinfun :: (chilbert_space, chilbert_space) ordered_complex_vector begin instance proof intro_classes fix x y z :: \'a \\<^sub>C\<^sub>L 'b\ fix a b :: complex define pos where \pos X \ (\\. cinner \ (X *\<^sub>V \) \ 0)\ for X :: \'b \\<^sub>C\<^sub>L 'b\ consider (unitary) \heterogenous_same_type_cblinfun TYPE('a) TYPE('b)\ \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B = pos ((B-A) o\<^sub>C\<^sub>L (heterogenous_cblinfun_id :: 'b\\<^sub>C\<^sub>L'a))\ | (trivial) \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B \ A = B\ apply atomize_elim by (auto simp: pos_def less_eq_cblinfun_def_heterogenous) note cases = this have [simp]: \pos 0\ unfolding pos_def by auto have pos_nondeg: \X = 0\ if \pos X\ and \pos (-X)\ for X apply (rule cblinfun_cinner_eqI, simp) using that by (metis (no_types, lifting) cblinfun.minus_left cinner_minus_right dual_order.antisym equation_minus_iff neg_le_0_iff_le pos_def) have pos_add: \pos (X+Y)\ if \pos X\ and \pos Y\ for X Y by (smt (z3) pos_def cblinfun.diff_left cinner_minus_right cinner_simps(3) diff_ge_0_iff_ge diff_minus_eq_add neg_le_0_iff_le order_trans that(1) that(2) uminus_cblinfun.rep_eq) have pos_scaleC: \pos (a *\<^sub>C X)\ if \a\0\ and \pos X\ for X a using that unfolding pos_def by (auto simp: cblinfun.scaleC_left) let ?id = \heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a\ show \x \ x\ apply (cases rule:cases) by auto show \(x < y) \ (x \ y \ \ y \ x)\ unfolding less_cblinfun_def by simp show \x \ z\ if \x \ y\ and \y \ z\ proof (cases rule:cases) case unitary define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ and \b = (z-y) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ with unitary that have \pos a\ and \pos b\ by auto then have \pos (a + b)\ by (rule pos_add) moreover have \a + b = (z - x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ unfolding a_def b_def by (metis (no_types, lifting) bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel ordered_field_class.sign_simps(2) ordered_field_class.sign_simps(8)) ultimately show ?thesis using unitary by auto next case trivial with that show ?thesis by auto qed show \x = y\ if \x \ y\ and \y \ x\ proof (cases rule:cases) case unitary then have \unitary ?id\ by (auto simp: heterogenous_same_type_cblinfun_def) define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L ?id\ and \b = (x-y) o\<^sub>C\<^sub>L ?id\ with unitary that have \pos a\ and \pos b\ by auto then have \a = 0\ apply (rule_tac pos_nondeg) apply (auto simp: a_def b_def) by (smt (verit, best) add.commute bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose cblinfun_compose_zero_left diff_0 diff_add_cancel group_cancel.rule0 group_cancel.sub1) then show ?thesis unfolding a_def using \unitary ?id\ by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_compose_zero_left eq_iff_diff_eq_0 unitaryD2) next case trivial with that show ?thesis by simp qed show \x + y \ x + z\ if \y \ z\ proof (cases rule:cases) case unitary with that show ?thesis by auto next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ a *\<^sub>C y\ if \x \ y\ and \0 \ a\ proof (cases rule:cases) case unitary with that pos_scaleC show ?thesis by (metis cblinfun_compose_scaleC_left complex_vector.scale_right_diff_distrib) next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ b *\<^sub>C x\ if \a \ b\ and \0 \ x\ proof (cases rule:cases) case unitary with that show ?thesis by (auto intro!: pos_scaleC simp flip: scaleC_diff_left) next case trivial with that show ?thesis by auto qed qed end lemma positive_id_cblinfun[simp]: "id_cblinfun \ 0" unfolding less_eq_cblinfun_def using cinner_ge_zero by auto lemma positive_hermitianI: \A = A*\ if \A \ 0\ apply (rule cinner_real_hermiteanI) using that by (auto simp: complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma 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 (* 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\ apply transfer by (simp add: scaleC_add_right) lemma norm_vector_to_cblinfun[simp]: "norm (vector_to_cblinfun x) = norm x" proof transfer have "bounded_clinear (one_dim_iso::'a \ complex)" by simp moreover have "onorm (one_dim_iso::'a \ complex) * norm x = norm x" for x :: 'b by simp ultimately show "onorm (\\. one_dim_iso (\::'a) *\<^sub>C x) = norm x" for x :: 'b by (subst onorm_scaleC_left) qed lemma bounded_clinear_vector_to_cblinfun[bounded_clinear]: "bounded_clinear vector_to_cblinfun" apply (rule bounded_clinearI[where K=1]) apply (transfer, simp add: scaleC_add_right) apply (transfer, simp add: mult.commute) by simp lemma vector_to_cblinfun_scaleC[simp]: "vector_to_cblinfun (a *\<^sub>C \) = a *\<^sub>C vector_to_cblinfun \" for a::complex 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 \" apply transfer by (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 apply auto by (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 - -subsection \Bifunctionals\ - -lift_definition bifunctional :: \'a::complex_normed_vector \\<^sub>C\<^sub>L (('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex)\ - is \\x f. f *\<^sub>V x\ - by (simp add: cblinfun.flip) - -lemma bifunctional_apply[simp]: \(bifunctional *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ - by (transfer fixing: x f, simp) - -lemma bifunctional_isometric[simp]: \norm (bifunctional *\<^sub>V x) = norm x\ for x :: \'a::complex_inner\ -proof - - define f :: \'a \\<^sub>C\<^sub>L complex\ where \f = CBlinfun (\y. cinner x y)\ - then have [simp]: \f *\<^sub>V y = cinner x y\ for y - by (simp add: bounded_clinear_CBlinfun_apply bounded_clinear_cinner_right) - then have [simp]: \norm f = norm x\ - apply (auto intro!: norm_cblinfun_eqI[where x=x] simp: power2_norm_eq_cinner[symmetric]) - apply (smt (verit, best) norm_eq_sqrt_cinner norm_ge_zero power2_norm_eq_cinner real_div_sqrt) - using Cauchy_Schwarz_ineq2 by blast - show ?thesis - apply (auto intro!: norm_cblinfun_eqI[where x=f]) - apply (metis norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) - by (metis norm_cblinfun ordered_field_class.sign_simps(33)) -qed - -lemma norm_bifunctional[simp]: \norm (bifunctional :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ -proof - - obtain x :: 'a where [simp]: \norm x = 1\ - by (meson UNIV_not_singleton ex_norm1) - show ?thesis - by (auto intro!: norm_cblinfun_eqI[where x=x]) -qed +(* Probably true without restricting to chilbert_space using a different proof. *) +lemma rank1_compose_left: \rank1 (a o\<^sub>C\<^sub>L b)\ if \rank1 b\ and \a o\<^sub>C\<^sub>L b \ 0\ + for b :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and a :: \'b \\<^sub>C\<^sub>L 'c::chilbert_space\ + using that apply (auto simp: rank1_iff_butterfly cblinfun_comp_butterfly) + by blast + +(* Probably true without restricting to chilbert_space using a different proof. *) +lemma rank1_compose_right: \rank1 (a o\<^sub>C\<^sub>L b)\ if \rank1 a\ and \a o\<^sub>C\<^sub>L b \ 0\ + for b :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and a :: \'b \\<^sub>C\<^sub>L 'c::chilbert_space\ + using that apply (auto simp: rank1_iff_butterfly butterfly_comp_cblinfun) + by blast + +(* Probably true without restricting to chilbert_space using a different proof. *) +lemma rank1_scaleC: \rank1 (c *\<^sub>C a)\ if \rank1 a\ and \c \ 0\ + for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ + using that apply (auto simp: rank1_iff_butterfly) + by (metis butterfly_scaleC_left) + +(* Probably true without restricting to chilbert_space using a different proof. *) +lemma rank1_uminus: \rank1 (-a)\ if \rank1 a\ + for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ + using that rank1_scaleC[where c=\-1\ and a=a] + by simp + +(* Probably true without restricting to chilbert_space using a different proof. *) +lemma rank1_uminus_iff[simp]: \rank1 (-a) \ rank1 a\ + for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ + 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) subsection \Banach-Steinhaus\ theorem cbanach_steinhaus: fixes F :: \'c \ 'a::cbanach \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\x. \M. \n. norm ((F n) *\<^sub>V x) \ M\ shows \\M. \ n. norm (F n) \ M\ - using cblinfun_blinfun_transfer[transfer_rule] apply (rule TrueI)? (* Deletes current facts *) + using cblinfun_blinfun_transfer[transfer_rule] apply fail? (* Deletes current facts *) proof (use assms in transfer) fix F :: \'c \ 'a \\<^sub>L 'b\ assume \(\x. \M. \n. norm (F n *\<^sub>v x) \ M)\ hence \\x. bounded (range (\n. blinfun_apply (F n) x))\ by (metis (no_types, lifting) boundedI rangeE) hence \bounded (range F)\ by (simp add: banach_steinhaus) thus \\M. \n. norm (F n) \ M\ by (simp add: bounded_iff) qed subsection \Riesz-representation theorem\ theorem riesz_frechet_representation_cblinfun_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ shows \\t. \x. f *\<^sub>V x = (t \\<^sub>C x)\ apply transfer by (rule riesz_frechet_representation_existence) lemma riesz_frechet_representation_cblinfun_unique: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::complex_inner \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = (t \\<^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: 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 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) 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 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]) 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) -definition the_riesz_rep_bilinear' :: \('a::complex_normed_vector \ 'b::chilbert_space \ complex) \ ('a \ 'b)\ where - \the_riesz_rep_bilinear' p x = the_riesz_rep (CBlinfun (p x))\ - -lemma the_riesz_rep_bilinear'_correct: - assumes \bounded_clinear (p x)\ - shows \(the_riesz_rep_bilinear' p x) \\<^sub>C y = p x y\ - by (auto simp add: the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply) - - -lemma the_riesz_rep_bilinear'_plus1: - assumes \\x. bounded_clinear (p x)\ and \\x. bounded_clinear (q x)\ - shows \the_riesz_rep_bilinear' (p + q) = the_riesz_rep_bilinear' p + the_riesz_rep_bilinear' q\ - by (auto intro!: ext simp add: the_riesz_rep_add CBlinfun_plus the_riesz_rep_bilinear'_def - assms bounded_clinear_CBlinfun_apply) - -lemma the_riesz_rep_bilinear'_plus2: +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_bilinear' p (x + y) = the_riesz_rep_bilinear' p x + the_riesz_rep_bilinear' p y\ -proof - - have [simp]: \p (x + y) = p x + p y\ - using assms bounded_sesquilinear.add_left by fastforce - have [simp]: \bounded_clinear (p x)\ for x - by (simp add: assms bounded_sesquilinear.bounded_clinear_right) - show ?thesis - by (auto intro!: ext simp add: the_riesz_rep_add CBlinfun_plus the_riesz_rep_bilinear'_def - assms bounded_clinear_CBlinfun_apply) -qed - -lemma the_riesz_rep_bilinear'_scaleC1: - assumes \\x. bounded_clinear (p x)\ - shows \the_riesz_rep_bilinear' (\x y. c *\<^sub>C p x y) = (\x. cnj c *\<^sub>C the_riesz_rep_bilinear' p x)\ - by (auto intro!: ext simp add: the_riesz_rep_scaleC CBlinfun_scaleC the_riesz_rep_bilinear'_def - assms bounded_clinear_CBlinfun_apply - simp del: complex_scaleC_def scaleC_conv_of_complex) - -lemma the_riesz_rep_bilinear'_scaleC2: - assumes \bounded_sesquilinear p\ - shows \the_riesz_rep_bilinear' p (c *\<^sub>C x) = c *\<^sub>C the_riesz_rep_bilinear' p x\ + shows \(the_riesz_rep_sesqui p 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 - - have [simp]: \p (c *\<^sub>C x) = (\y. cnj c *\<^sub>C p x y)\ - using assms bounded_sesquilinear.scaleC_left by blast - have [simp]: \bounded_clinear (p x)\ for x - by (simp add: assms bounded_sesquilinear.bounded_clinear_right) + 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 - by (auto intro!: ext simp add: the_riesz_rep_scaleC CBlinfun_scaleC the_riesz_rep_bilinear'_def - assms bounded_clinear_CBlinfun_apply - simp del: complex_scaleC_def scaleC_conv_of_complex) + 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 bounded_clinear_the_riesz_rep_bilinear'2: - assumes \bounded_sesquilinear p\ - shows \bounded_clinear (the_riesz_rep_bilinear' p)\ +lemma norm_bidual_embedding[simp]: \norm (bidual_embedding :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ proof - - obtain K0 where K0: \cmod (p x y) \ norm x * norm y * K0\ for x y - using assms bounded_sesquilinear.bounded by blast - define K where \K = max K0 0\ - with K0 have K: \cmod (p x y) \ norm x * norm y * K\ for x y - by (smt (verit, del_insts) mult_nonneg_nonneg mult_nonneg_nonpos norm_ge_zero) - have [simp]: \K \ 0\ - by (simp add: K_def) - have [simp]: \bounded_clinear (p x)\ for x - by (simp add: assms bounded_sesquilinear.bounded_clinear_right) - have "norm (the_riesz_rep_bilinear' p x) \ norm x * K" for x - unfolding the_riesz_rep_bilinear'_def - using K by (auto intro!: norm_cblinfun_bound simp: bounded_clinear_CBlinfun_apply mult.commute mult.left_commute) + 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 + 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 + 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 - apply (rule_tac bounded_clinearI) - by (auto simp: assms the_riesz_rep_bilinear'_plus2 the_riesz_rep_bilinear'_scaleC2) + by (metis cblinfun_eqI surj_def) qed -lift_definition the_riesz_rep_bilinear:: \('a::complex_normed_vector \ 'b::chilbert_space \ complex) \ ('a \\<^sub>C\<^sub>L 'b)\ is - \\p. if bounded_sesquilinear p then the_riesz_rep_bilinear' p else 0\ - by (auto simp: bounded_clinear_the_riesz_rep_bilinear'2 zero_fun_def) - -lemma the_riesz_rep_bilinear_correct: - assumes \bounded_sesquilinear p\ - shows \(the_riesz_rep_bilinear p x) \\<^sub>C y = p x y\ - apply (transfer fixing: p) - by (simp add: assms bounded_sesquilinear.bounded_clinear_right the_riesz_rep_bilinear'_correct) - - 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_exists_norm: \B \ 0 \ norm (cblinfun_extension S f) \ B\ + 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 apply atomize_elim by (meson CauchyD \0 < B'\ \0 < e\ linordered_field_class.divide_pos_pos) have \norm (f (seq x m) - f (seq x n)) < e\ if \n \ M\ and \m \ M\ for n m proof - have \norm (f (seq x m) - f (seq x n)) = norm (f (seq x m - seq x n))\ using f_add f_scale seq_S by (metis add_diff_cancel assms(1) complex_vector.subspace_diff diff_add_cancel) also have \\ \ B' * norm (seq x m - seq x n)\ apply (rule bounded') by (simp add: assms(1) complex_vector.subspace_diff seq_S) also from less_eB have \\ < B' * (e/B')\ by (meson \0 < B'\ linordered_semiring_strict_class.mult_strict_left_mono that) also have \\ \ e\ using \0 < B'\ by auto finally show ?thesis by - qed then show \\M. \m\M. \n\M. norm (f (seq x m) - f (seq x n)) < e\ by auto qed then have f_seq_lim: \(\i. f (seq x i)) \ g x\ for x by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff g_def) have f_xi_lim: \(\i. f (xi i)) \ g x\ if \xi \ x\ and \\i. xi i \ S\ for xi x proof - from seq_lim that have \(\i. B' * norm (xi i - seq x i)) \ 0\ by (metis (no_types) \0 < B'\ cancel_comm_monoid_add_class.diff_cancel norm_not_less_zero norm_zero tendsto_diff tendsto_norm_zero_iff tendsto_zero_mult_left_iff) then have \(\i. f (xi i + (-1) *\<^sub>C seq x i)) \ 0\ apply (rule Lim_null_comparison[rotated]) using bounded' by (simp add: assms(1) complex_vector.subspace_diff seq_S that(2)) then have \(\i. f (xi i) - f (seq x i)) \ 0\ apply (subst (asm) f_add) apply (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S)[2] apply (subst (asm) f_scale) by (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S) then show \(\i. f (xi i)) \ g x\ using Lim_transform f_seq_lim by fastforce qed have g_add: \g (x + y) = g x + g y\ for x y proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto obtain yi :: \nat \ 'a\ where \yi \ y\ and \yi i \ S\ for i using seq_S seq_lim by auto have \(\i. xi i + yi i) \ x + y\ using \xi \ x\ \yi \ y\ tendsto_add by blast then have lim1: \(\i. f (xi i + yi i)) \ g (x + y)\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ assms(1) complex_vector.subspace_add f_xi_lim) have \(\i. f (xi i + yi i)) = (\i. f (xi i) + f (yi i))\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ f_add) also have \\ \ g x + g y\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ \xi \ x\ \yi \ y\ f_xi_lim tendsto_add) finally show ?thesis using lim1 LIMSEQ_unique by blast qed have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ for c x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto have \(\i. c *\<^sub>C xi i) \ c *\<^sub>C x\ using \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at isCont_tendsto_compose by blast then have lim1: \(\i. f (c *\<^sub>C xi i)) \ g (c *\<^sub>C x)\ by (simp add: \\i. xi i \ S\ assms(1) complex_vector.subspace_scale f_xi_lim) have \(\i. f (c *\<^sub>C xi i)) = (\i. c *\<^sub>C f (xi i))\ by (simp add: \\i. xi i \ S\ f_scale) also have \\ \ c *\<^sub>C g x\ using \\i. xi i \ S\ \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at f_xi_lim isCont_tendsto_compose by blast finally show ?thesis using lim1 LIMSEQ_unique by blast qed have [simp]: \f x = g x\ if \x \ S\ for x proof - have \(\_. x) \ x\ by auto then have \(\_. f x) \ g x\ using that by (rule f_xi_lim) then show \f x = g x\ by (simp add: LIMSEQ_const_iff) qed have g_bounded: \norm (g x) \ B' * norm x\ for x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto then have \(\i. f (xi i)) \ g x\ using f_xi_lim by presburger then have \(\i. norm (f (xi i))) \ norm (g x)\ by (metis tendsto_norm) moreover have \(\i. B' * norm (xi i)) \ B' * norm x\ by (simp add: \xi \ x\ tendsto_mult_left tendsto_norm) ultimately show \norm (g x) \ B' * norm x\ apply (rule lim_mono[rotated]) using bounded' using \xi _ \ S\ by blast qed have \bounded_clinear g\ using g_add g_scale apply (rule bounded_clinearI[where K=B']) using g_bounded by (simp add: ordered_field_class.sign_simps(5)) then have [simp]: \CBlinfun g *\<^sub>V x = g x\ for x by (subst CBlinfun_inverse, auto) show \cblinfun_extension_exists S f\ apply (rule cblinfun_extension_existsI[where B=\CBlinfun g\]) by auto 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\ apply auto by (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_exists_norm) + 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 \\P. is_projection_on P (closure S) \ bounded_clinear P\ (* Maybe can be replaced by is_Proj if the latter's type class is widened *) + 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 where P_proj: \is_projection_on P (closure S)\ and P_blin[simp]: \bounded_clinear P\ - using assms(2) by blast - have P_lin[simp]: \clinear P\ - by (simp add: bounded_clinear.clinear) - define f' S' where \f' \ = f (P \)\ and \S' = S + (P -` {0})\ for \ + 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 + (P -` {0})\ - apply (rule set_plus_intro) - using P_proj is_projection_on_in_image - by (auto simp: complex_vector.linear_diff is_projection_on_fixes_image is_projection_on_in_image) - also have \\ \ closure (closure S + (P -` {0}))\ + 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 + (P -` {0}))\ + 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 - by (smt (z3) P_blin P_proj S'_def f'_def add.right_neutral bounded_clinear_CBlinfun_apply cblinfun.add_right closure_subset f_add is_projection_on_fixes_image set_plus_elim singletonD subset_eq that(1) that(2) vimageE) + 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 - by (smt (verit, ccfv_SIG) P_blin P_proj S'_def f'_def add.right_neutral bounded_clinear_CBlinfun_apply cblinfun.add_right cblinfun.scaleC_right closure_subset f_scale is_projection_on_fixes_image set_plus_elim singletonD subset_eq that vimageE) - + using that by (auto simp: f'_def cblinfun.scaleC_right f_scale PS') moreover - from P_blin obtain B' where B': \norm (P x) \ B' * norm x\ for x - by (metis bounded_clinear.bounded mult.commute) - have \norm (f' x) \ (B * B') * norm x\ if \x \ S'\ for x + have \norm (f' x) \ (B * norm P) * norm x\ if \x \ S'\ for x proof - have \norm (f' x) \ B* norm (P x)\ - apply (auto simp: f'_def) - by (smt (verit) P_blin P_proj S'_def add.right_neutral bounded bounded_clinear_CBlinfun_apply cblinfun.add_right closure_subset is_projection_on_fixes_image set_plus_elim singletonD subset_eq that vimageE) - also have \\ \ B * B' * norm x\ - by (simp add: B' mult.assoc mult_mono) + 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'\ - from F_ex have *: \F \ = f' \\ if \\ \ S'\ for \ - by (simp add: F_def cblinfun_extension_apply that) - then have \F \ = f \\ if \\ \ S\ for \ - apply (auto simp: S'_def f'_def) - by (metis (no_types, lifting) P_lin P_proj add.right_neutral closure_subset complex_vector.linear_subspace_vimage complex_vector.subspace_0 complex_vector.subspace_single_0 is_projection_on_fixes_image set_plus_intro subset_eq that) + 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 apply auto by (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) + apply simp + by (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 + 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 = top\ \ccspan F = top\ 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)\ apply transfer by 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 apply auto by (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 apply simp by (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 fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows bij_between_bases_bij: \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) 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) end bundle blinfun_notation begin notation blinfun_apply (infixr "*\<^sub>V" 70) end bundle no_blinfun_notation begin no_notation blinfun_apply (infixr "*\<^sub>V" 70) end unbundle no_cblinfun_notation unbundle no_lattice_syntax end diff --git a/thys/Complex_Bounded_Operators/Complex_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,2446 +1,2436 @@ (* 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 -(* TODO: Remove this eventually. Do not use this syntax. *) -bundle cinner_bracket_notation begin -notation cinner ("\_, _\") -end - -(* TODO: Remove this eventually. Do not use this syntax. *) -bundle no_cinner_bracket_notation begin -no_notation cinner ("\_, _\") -end - 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[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 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 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: \ \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: \ \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 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) 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) 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,1557 +1,1535 @@ section \\Complex_L2\ -- Hilbert space of square-summable functions\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_L2 imports Complex_Bounded_Linear_Function "HOL-Analysis.L2_Norm" "HOL-Library.Rewrite" "HOL-Analysis.Infinite_Sum" begin unbundle lattice_syntax unbundle cblinfun_notation unbundle no_notation_blinfun_apply subsection \l2 norm of functions\ definition "has_ell2_norm (x::_\complex) \ (\i. (x i)\<^sup>2) abs_summable_on UNIV" lemma has_ell2_norm_bdd_above: \has_ell2_norm x \ bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (simp add: has_ell2_norm_def abs_summable_iff_bdd_above) lemma has_ell2_norm_L2_set: "has_ell2_norm x = bdd_above (L2_set (norm o x) ` Collect finite)" proof (rule iffI) have \mono sqrt\ using monoI real_sqrt_le_mono by blast assume \has_ell2_norm x\ then have *: \bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (subst (asm) has_ell2_norm_bdd_above) have \bdd_above ((\F. sqrt (sum (\xa. norm ((x xa)\<^sup>2)) F)) ` Collect finite)\ using bdd_above_image_mono[OF \mono sqrt\ *] by (auto simp: image_image) then show \bdd_above (L2_set (norm o x) ` Collect finite)\ by (auto simp: L2_set_def norm_power) next define p2 where \p2 x = (if x < 0 then 0 else x^2)\ for x :: real have \mono p2\ by (simp add: monoI p2_def) have [simp]: \p2 (L2_set f F) = (\i\F. (f i)\<^sup>2)\ for f and F :: \'a set\ by (smt (verit) L2_set_def L2_set_nonneg p2_def power2_less_0 real_sqrt_pow2 sum.cong sum_nonneg) assume *: \bdd_above (L2_set (norm o x) ` Collect finite)\ have \bdd_above (p2 ` L2_set (norm o x) ` Collect finite)\ using bdd_above_image_mono[OF \mono p2\ *] by auto then show \has_ell2_norm x\ apply (simp add: image_image has_ell2_norm_def abs_summable_iff_bdd_above) by (simp add: norm_power) qed definition ell2_norm :: \('a \ complex) \ real\ where \ell2_norm x = sqrt (\\<^sub>\i. norm (x i)^2)\ lemma ell2_norm_SUP: assumes \has_ell2_norm x\ shows "ell2_norm x = sqrt (SUP F\{F. finite F}. sum (\i. norm (x i)^2) F)" using assms apply (auto simp add: ell2_norm_def has_ell2_norm_def) apply (subst infsum_nonneg_is_SUPREMUM_real) by (auto simp: norm_power) lemma ell2_norm_L2_set: assumes "has_ell2_norm x" shows "ell2_norm x = (SUP F\{F. finite F}. L2_set (norm o x) F)" proof- have "sqrt (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)) = (SUP F\{F. finite F}. sqrt (\i\F. (cmod (x i))\<^sup>2))" proof (subst continuous_at_Sup_mono) show "mono sqrt" by (simp add: mono_def) show "continuous (at_left (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite))) sqrt" using continuous_at_split isCont_real_sqrt by blast show "sum (\i. (cmod (x i))\<^sup>2) ` Collect finite \ {}" by auto show "bdd_above (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)" using has_ell2_norm_bdd_above[THEN iffD1, OF assms] by (auto simp: norm_power) show "\ (sqrt ` sum (\i. (cmod (x i))\<^sup>2) ` Collect finite) = (SUP F\Collect finite. sqrt (\i\F. (cmod (x i))\<^sup>2))" by (metis image_image) qed thus ?thesis using assms by (auto simp: ell2_norm_SUP L2_set_def) qed lemma has_ell2_norm_finite[simp]: "has_ell2_norm (x::'a::finite\_)" unfolding has_ell2_norm_def by simp lemma ell2_norm_finite: "ell2_norm (x::'a::finite\complex) = sqrt (sum (\i. (norm(x i))^2) UNIV)" by (simp add: ell2_norm_def) lemma ell2_norm_finite_L2_set: "ell2_norm (x::'a::finite\complex) = L2_set (norm o x) UNIV" by (simp add: ell2_norm_finite L2_set_def) lemma ell2_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. if a = i then 1 else 0)\ shows has_ell2_norm_ket: \has_ell2_norm f\ and ell2_norm_ket: \ell2_norm f = 1\ proof - have \(\x. (f x)\<^sup>2) abs_summable_on {a}\ apply (rule summable_on_finite) by simp then show \has_ell2_norm f\ unfolding has_ell2_norm_def apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1]) unfolding f_def by auto have \(\\<^sub>\x\{a}. (f x)\<^sup>2) = 1\ apply (subst infsum_finite) by (auto simp: f_def) then show \ell2_norm f = 1\ unfolding ell2_norm_def apply (subst infsum_cong_neutral[where T=\{a}\ and g=\\x. (cmod (f x))\<^sup>2\]) by (auto simp: f_def) qed lemma ell2_norm_geq0: \ell2_norm x \ 0\ by (auto simp: ell2_norm_def intro!: infsum_nonneg) lemma ell2_norm_point_bound: assumes \has_ell2_norm x\ shows \ell2_norm x \ cmod (x i)\ proof - have \(cmod (x i))\<^sup>2 = norm ((x i)\<^sup>2)\ by (simp add: norm_power) also have \norm ((x i)\<^sup>2) = sum (\i. (norm ((x i)\<^sup>2))) {i}\ by auto also have \\ = infsum (\i. (norm ((x i)\<^sup>2))) {i}\ by (rule infsum_finite[symmetric], simp) also have \\ \ infsum (\i. (norm ((x i)\<^sup>2))) UNIV\ apply (rule infsum_mono_neutral) using assms by (auto simp: has_ell2_norm_def) also have \\ = (ell2_norm x)\<^sup>2\ by (metis (no_types, lifting) ell2_norm_def ell2_norm_geq0 infsum_cong norm_power real_sqrt_eq_iff real_sqrt_unique) finally show ?thesis using ell2_norm_geq0 power2_le_imp_le by blast qed lemma ell2_norm_0: assumes "has_ell2_norm x" - shows "(ell2_norm x = 0) = (x = (\_. 0))" + shows "ell2_norm x = 0 \ x = (\_. 0)" proof assume u1: "x = (\_. 0)" have u2: "(SUP x::'a set\Collect finite. (0::real)) = 0" if "x = (\_. 0)" by (metis cSUP_const empty_Collect_eq finite.emptyI) show "ell2_norm x = 0" unfolding ell2_norm_def using u1 u2 by auto next assume norm0: "ell2_norm x = 0" show "x = (\_. 0)" proof fix i have \cmod (x i) \ ell2_norm x\ using assms by (rule ell2_norm_point_bound) also have \\ = 0\ by (fact norm0) finally show "x i = 0" by auto qed qed lemma ell2_norm_smult: assumes "has_ell2_norm x" shows "has_ell2_norm (\i. c * x i)" and "ell2_norm (\i. c * x i) = cmod c * ell2_norm x" proof - have L2_set_mul: "L2_set (cmod \ (\i. c * x i)) F = cmod c * L2_set (cmod \ x) F" for F proof- have "L2_set (cmod \ (\i. c * x i)) F = L2_set (\i. (cmod c * (cmod o x) i)) F" by (metis comp_def norm_mult) also have "\ = cmod c * L2_set (cmod o x) F" by (metis norm_ge_zero L2_set_right_distrib) finally show ?thesis . qed from assms obtain M where M: "M \ L2_set (cmod o x) F" if "finite F" for F unfolding has_ell2_norm_L2_set bdd_above_def by auto hence "cmod c * M \ L2_set (cmod o (\i. c * x i)) F" if "finite F" for F unfolding L2_set_mul by (simp add: ordered_comm_semiring_class.comm_mult_left_mono that) thus has: "has_ell2_norm (\i. c * x i)" unfolding has_ell2_norm_L2_set bdd_above_def using L2_set_mul[symmetric] by auto have "ell2_norm (\i. c * x i) = (SUP F \ Collect finite. (L2_set (cmod \ (\i. c * x i)) F))" by (simp add: ell2_norm_L2_set has) also have "\ = (SUP F \ Collect finite. (cmod c * L2_set (cmod \ x) F))" using L2_set_mul by auto also have "\ = cmod c * ell2_norm x" proof (subst ell2_norm_L2_set) show "has_ell2_norm x" by (simp add: assms) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = cmod c * \ (L2_set (cmod \ x) ` Collect finite)" proof (subst continuous_at_Sup_mono [where f = "\x. cmod c * x"]) show "mono ((*) (cmod c))" by (simp add: mono_def ordered_comm_semiring_class.comm_mult_left_mono) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) ((*) (cmod c))" proof (rule continuous_mult) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. cmod c)" by simp show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. x)" by simp qed show "L2_set (cmod \ x) ` Collect finite \ {}" by auto show "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms has_ell2_norm_L2_set) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = \ ((*) (cmod c) ` L2_set (cmod \ x) ` Collect finite)" by (metis image_image) qed qed finally show "ell2_norm (\i. c * x i) = cmod c * ell2_norm x". qed lemma ell2_norm_triangle: assumes "has_ell2_norm x" and "has_ell2_norm y" shows "has_ell2_norm (\i. x i + y i)" and "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" proof - have triangle: "L2_set (cmod \ (\i. x i + y i)) F \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" (is "?lhs\?rhs") if "finite F" for F proof - have "?lhs \ L2_set (\i. (cmod o x) i + (cmod o y) i) F" proof (rule L2_set_mono) show "(cmod \ (\i. x i + y i)) i \ (cmod \ x) i + (cmod \ y) i" if "i \ F" for i :: 'a using that norm_triangle_ineq by auto show "0 \ (cmod \ (\i. x i + y i)) i" if "i \ F" for i :: 'a using that by simp qed also have "\ \ ?rhs" by (rule L2_set_triangle_ineq) finally show ?thesis . qed obtain Mx My where Mx: "Mx \ L2_set (cmod o x) F" and My: "My \ L2_set (cmod o y) F" if "finite F" for F using assms unfolding has_ell2_norm_L2_set bdd_above_def by auto hence MxMy: "Mx + My \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" if "finite F" for F using that by fastforce hence bdd_plus: "bdd_above ((\xa. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa) ` Collect finite)" unfolding bdd_above_def by auto from MxMy have MxMy': "Mx + My \ L2_set (cmod \ (\i. x i + y i)) F" if "finite F" for F using triangle that by fastforce thus has: "has_ell2_norm (\i. x i + y i)" unfolding has_ell2_norm_L2_set bdd_above_def by auto have SUP_plus: "(SUP x\A. f x + g x) \ (SUP x\A. f x) + (SUP x\A. g x)" if notempty: "A\{}" and bddf: "bdd_above (f`A)"and bddg: "bdd_above (g`A)" for f g :: "'a set \ real" and A proof- have xleq: "x \ (SUP x\A. f x) + (SUP x\A. g x)" if x: "x \ (\x. f x + g x) ` A" for x proof - obtain a where aA: "a:A" and ax: "x = f a + g a" using x by blast have fa: "f a \ (SUP x\A. f x)" by (simp add: bddf aA cSUP_upper) moreover have "g a \ (SUP x\A. g x)" by (simp add: bddg aA cSUP_upper) ultimately have "f a + g a \ (SUP x\A. f x) + (SUP x\A. g x)" by simp with ax show ?thesis by simp qed have "(\x. f x + g x) ` A \ {}" using notempty by auto moreover have "x \ \ (f ` A) + \ (g ` A)" if "x \ (\x. f x + g x) ` A" for x :: real using that by (simp add: xleq) ultimately show ?thesis by (meson bdd_above_def cSup_le_iff) qed have a2: "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms(1) has_ell2_norm_L2_set) have a3: "bdd_above (L2_set (cmod \ y) ` Collect finite)" by (meson assms(2) has_ell2_norm_L2_set) have a1: "Collect finite \ {}" by auto have a4: "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ (SUP xa\Collect finite. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa)" by (metis (mono_tags, lifting) a1 bdd_plus cSUP_mono mem_Collect_eq triangle) have "\r. \ (L2_set (cmod \ (\a. x a + y a)) ` Collect finite) \ r \ \ (SUP A\Collect finite. L2_set (cmod \ x) A + L2_set (cmod \ y) A) \ r" using a4 by linarith hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ \ (L2_set (cmod \ x) ` Collect finite) + \ (L2_set (cmod \ y) ` Collect finite)" by (metis (no_types) SUP_plus a1 a2 a3) hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ ell2_norm x + ell2_norm y" by (simp add: assms(1) assms(2) ell2_norm_L2_set) thus "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" by (simp add: ell2_norm_L2_set has) qed lemma ell2_norm_uminus: assumes "has_ell2_norm x" shows \has_ell2_norm (\i. - x i)\ and \ell2_norm (\i. - x i) = ell2_norm x\ using assms by (auto simp: has_ell2_norm_def ell2_norm_def) subsection \The type \ell2\ of square-summable functions\ typedef 'a ell2 = "{x::'a\complex. has_ell2_norm x}" unfolding has_ell2_norm_def by (rule exI[of _ "\_.0"], auto) setup_lifting type_definition_ell2 instantiation ell2 :: (type)complex_vector begin lift_definition zero_ell2 :: "'a ell2" is "\_. 0" by (auto simp: has_ell2_norm_def) lift_definition uminus_ell2 :: "'a ell2 \ 'a ell2" is uminus by (simp add: has_ell2_norm_def) lift_definition plus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x + g x" by (rule ell2_norm_triangle) lift_definition minus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x - g x" apply (subst add_uminus_conv_diff[symmetric]) apply (rule ell2_norm_triangle) by (auto simp add: ell2_norm_uminus) lift_definition scaleR_ell2 :: "real \ 'a ell2 \ 'a ell2" is "\r f x. complex_of_real r * f x" by (rule ell2_norm_smult) lift_definition scaleC_ell2 :: "complex \ 'a ell2 \ 'a ell2" is "\c f x. c * f x" by (rule ell2_norm_smult) instance proof fix a b c :: "'a ell2" show "((*\<^sub>R) r::'a ell2 \ _) = (*\<^sub>C) (complex_of_real r)" for r apply (rule ext) apply transfer by auto show "a + b + c = a + (b + c)" by (transfer; rule ext; simp) show "a + b = b + a" by (transfer; rule ext; simp) show "0 + a = a" by (transfer; rule ext; simp) show "- a + a = 0" by (transfer; rule ext; simp) show "a - b = a + - b" by (transfer; rule ext; simp) show "r *\<^sub>C (a + b) = r *\<^sub>C a + r *\<^sub>C b" for r apply (transfer; rule ext) by (simp add: vector_space_over_itself.scale_right_distrib) show "(r + r') *\<^sub>C a = r *\<^sub>C a + r' *\<^sub>C a" for r r' apply (transfer; rule ext) by (simp add: ring_class.ring_distribs(2)) show "r *\<^sub>C r' *\<^sub>C a = (r * r') *\<^sub>C a" for r r' by (transfer; rule ext; simp) show "1 *\<^sub>C a = a" by (transfer; rule ext; simp) qed end instantiation ell2 :: (type)complex_normed_vector begin lift_definition norm_ell2 :: "'a ell2 \ real" is ell2_norm . declare norm_ell2_def[code del] definition "dist x y = norm (x - y)" for x y::"'a ell2" definition "sgn x = x /\<^sub>R norm x" for x::"'a ell2" definition [code del]: "uniformity = (INF e\{0<..}. principal {(x::'a ell2, y). norm (x - y) < e})" definition [code del]: "open U = (\x\U. \\<^sub>F (x', y) in INF e\{0<..}. principal {(x, y). norm (x - y) < e}. x' = x \ y \ U)" for U :: "'a ell2 set" instance proof fix a b :: "'a ell2" show "dist a b = norm (a - b)" by (simp add: dist_ell2_def) show "sgn a = a /\<^sub>R norm a" by (simp add: sgn_ell2_def) show "uniformity = (INF e\{0<..}. principal {(x, y). dist (x::'a ell2) y < e})" unfolding dist_ell2_def uniformity_ell2_def by simp show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. (x'::'a ell2) = x \ y \ U)" for U :: "'a ell2 set" unfolding uniformity_ell2_def open_ell2_def by simp_all show "(norm a = 0) = (a = 0)" apply transfer by (fact ell2_norm_0) show "norm (a + b) \ norm a + norm b" apply transfer by (fact ell2_norm_triangle) show "norm (r *\<^sub>R (a::'a ell2)) = \r\ * norm a" for r and a :: "'a ell2" apply transfer by (simp add: ell2_norm_smult(2)) show "norm (r *\<^sub>C a) = cmod r * norm a" for r apply transfer by (simp add: ell2_norm_smult(2)) qed end lemma norm_point_bound_ell2: "norm (Rep_ell2 x i) \ norm x" apply transfer by (simp add: ell2_norm_point_bound) lemma ell2_norm_finite_support: assumes \finite S\ \\ i. i \ S \ Rep_ell2 x i = 0\ shows \norm x = sqrt ((sum (\i. (cmod (Rep_ell2 x i))\<^sup>2)) S)\ proof (insert assms(2), transfer fixing: S) fix x :: \'a \ complex\ assume zero: \\i. i \ S \ x i = 0\ have \ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)\ by (auto simp: ell2_norm_def) also have \\ = sqrt (\\<^sub>\i\S. (cmod (x i))\<^sup>2)\ apply (subst infsum_cong_neutral[where g=\\i. (cmod (x i))\<^sup>2\ and S=UNIV and T=S]) using zero by auto also have \\ = sqrt (\i\S. (cmod (x i))\<^sup>2)\ using \finite S\ by simp finally show \ell2_norm x = sqrt (\i\S. (cmod (x i))\<^sup>2)\ by - qed instantiation ell2 :: (type) complex_inner begin lift_definition cinner_ell2 :: "'a ell2 \ 'a ell2 \ complex" is "\x y. infsum (\i. (cnj (x i) * y i)) UNIV" . declare cinner_ell2_def[code del] instance proof standard fix x y z :: "'a ell2" fix c :: complex show "cinner x y = cnj (cinner y x)" proof transfer fix x y :: "'a\complex" assume "has_ell2_norm x" and "has_ell2_norm y" have "(\\<^sub>\i. cnj (x i) * y i) = (\\<^sub>\i. cnj (cnj (y i) * x i))" by (metis complex_cnj_cnj complex_cnj_mult mult.commute) also have "\ = cnj (\\<^sub>\i. cnj (y i) * x i)" by (metis infsum_cnj) finally show "(\\<^sub>\i. cnj (x i) * y i) = cnj (\\<^sub>\i. cnj (y i) * x i)" . qed show "cinner (x + y) z = cinner x z + cinner y z" proof transfer fix x y z :: "'a \ complex" assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence cnj_y: "(\i. cnj (y i) * cnj (y i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm z" hence z: "(\i. z i * z i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_z:"(\i. cnj (x i) * z i) abs_summable_on UNIV" using cnj_x z by (rule abs_summable_product) have cnj_y_z:"(\i. cnj (y i) * z i) abs_summable_on UNIV" using cnj_y z by (rule abs_summable_product) show "(\\<^sub>\i. cnj (x i + y i) * z i) = (\\<^sub>\i. cnj (x i) * z i) + (\\<^sub>\i. cnj (y i) * z i)" apply (subst infsum_add [symmetric]) using cnj_x_z cnj_y_z by (auto simp add: summable_on_iff_abs_summable_on_complex distrib_left mult.commute) qed show "cinner (c *\<^sub>C x) y = cnj c * cinner x y" proof transfer fix x y :: "'a \ complex" and c :: complex assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence y: "(\i. y i * y i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_y:"(\i. cnj (x i) * y i) abs_summable_on UNIV" using cnj_x y by (rule abs_summable_product) thus "(\\<^sub>\i. cnj (c * x i) * y i) = cnj c * (\\<^sub>\i. cnj (x i) * y i)" by (auto simp flip: infsum_cmult_right simp add: abs_summable_summable mult.commute vector_space_over_itself.scale_left_commute) qed show "0 \ cinner x x" proof transfer fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence "(\i. cnj (x i) * x i) abs_summable_on UNIV" by auto hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square. have "0 = (\\<^sub>\i::'a. 0)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_complex) by (auto simp add: abs_summable_summable sum) finally show "0 \ (\\<^sub>\i. cnj (x i) * x i)" by assumption qed show "(cinner x x = 0) = (x = 0)" proof (transfer, auto) fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (smt (verit, del_insts) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_ge_zero norm_power real_norm_def summable_on_cong) hence cmod_x2: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square by simp assume eq0: "(\\<^sub>\i. cnj (x i) * x i) = 0" show "x = (\_. 0)" proof (rule ccontr) assume "x \ (\_. 0)" then obtain i where "x i \ 0" by auto hence "0 < cnj (x i) * x i" by (metis le_less cnj_x_x_geq0 complex_cnj_zero_iff vector_space_over_itself.scale_eq_0_iff) also have "\ = (\\<^sub>\i\{i}. cnj (x i) * x i)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_neutral_complex) by (auto simp add: abs_summable_summable cmod_x2) also from eq0 have "\ = 0" by assumption finally show False by simp qed qed show "norm x = sqrt (cmod (cinner x x))" proof transfer fix x :: "'a \ complex" assume x: "has_ell2_norm x" have "(\i::'a. cmod (x i) * cmod (x i)) abs_summable_on UNIV \ (\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" by (metis (no_types, lifting) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_power summable_on_cong x) from x have "ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)" unfolding ell2_norm_def by simp also have "\ = sqrt (\\<^sub>\i. cmod (cnj (x i) * x i))" unfolding norm_complex_def power2_eq_square by auto also have "\ = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by (auto simp: infsum_cmod abs_summable_summable sum) finally show "ell2_norm x = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by assumption qed qed end instance ell2 :: (type) chilbert_space -proof +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 -instantiation ell2 :: (CARD_1) complex_algebra_1 -begin -lift_definition one_ell2 :: "'a ell2" is "\_. 1" by simp -lift_definition times_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x * b x" - by simp -instance -proof - fix a b c :: "'a ell2" and r :: complex - show "a * b * c = a * (b * c)" - by (transfer, auto) - show "(a + b) * c = a * c + b * c" - apply (transfer, rule ext) - by (simp add: distrib_left mult.commute) - show "a * (b + c) = a * b + a * c" - apply transfer - by (simp add: ring_class.ring_distribs(1)) - show "r *\<^sub>C a * b = r *\<^sub>C (a * b)" - by (transfer, auto) - show "(a::'a ell2) * r *\<^sub>C b = r *\<^sub>C (a * b)" - by (transfer, auto) - show "1 * a = a" - by (transfer, rule ext, auto) - show "a * 1 = a" - by (transfer, rule ext, auto) - show "(0::'a ell2) \ 1" - apply transfer - by (meson zero_neq_one) -qed -end - -instantiation ell2 :: (CARD_1) field begin -lift_definition divide_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x / b x" - by simp -lift_definition inverse_ell2 :: "'a ell2 \ 'a ell2" is "\a x. inverse (a x)" - by simp -instance -proof (intro_classes; transfer) - fix a :: "'a \ complex" - assume "a \ (\_. 0)" - then obtain y where ay: "a y \ 0" - by auto - show "(\x. inverse (a x) * a x) = (\_. 1)" - proof (rule ext) - fix x - have "x = y" - by auto - with ay have "a x \ 0" - by metis - then show "inverse (a x) * a x = 1" - by auto - qed -qed (auto simp add: divide_complex_def mult.commute ring_class.ring_distribs) -end - 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 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 subsection \Kets and bras\ lift_definition ket :: "'a \ 'a ell2" is "\x y. if x=y then 1 else 0" by (rule has_ell2_norm_ket) abbreviation bra :: "'a \ (_,complex) cblinfun" where "bra i \ vector_to_cblinfun (ket i)*" for i instance ell2 :: (type) not_singleton proof standard have "ket undefined \ (0::'a ell2)" proof transfer show "(\y. if (undefined::'a) = y then 1::complex else 0) \ (\_. 0)" by (meson one_neq_zero) qed thus \\x y::'a ell2. x \ y\ by blast qed lemma cinner_ket_left: \(ket i \\<^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 lemma cinner_ket_eqI: assumes \\i. cinner (ket i) \ = cinner (ket i) \\ shows \\ = \\ by (metis Rep_ell2_inject assms cinner_ket_left ext) lemma norm_ket[simp]: "norm (ket i) = 1" apply transfer by (rule ell2_norm_ket) lemma cinner_ket_same[simp]: \(ket i \\<^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) lemma cinner_ket: \(ket i \\<^sub>C ket j) = (if i=j then 1 else 0)\ by (simp add: cinner_ket_left ket.rep_eq) lemma ket_injective[simp]: \ket i = ket j \ i = j\ by (metis cinner_ket one_neq_zero) lemma inj_ket[simp]: \inj ket\ by (simp add: inj_on_def) - lemma trunc_ell2_ket_cspan: \trunc_ell2 S x \ (cspan (range ket))\ if \finite S\ proof (use that in induction) case empty then show ?case by (auto intro: complex_vector.span_zero) next case (insert a F) from insert.hyps have \trunc_ell2 (insert a F) x = trunc_ell2 F x + Rep_ell2 x a *\<^sub>C ket a\ apply (transfer fixing: F a) by auto with insert.IH show ?case by (simp add: complex_vector.span_add_eq complex_vector.span_base complex_vector.span_scale) qed lemma closed_cspan_range_ket[simp]: \closure (cspan (range ket)) = UNIV\ proof (intro set_eqI iffI UNIV_I closure_approachable[THEN iffD2] allI impI) fix \ :: \'a ell2\ fix e :: real assume \e > 0\ have \((\S. trunc_ell2 S \) \ \) (finite_subsets_at_top UNIV)\ by (rule trunc_ell2_lim_at_UNIV) then obtain F where \finite F\ and \dist (trunc_ell2 F \) \ < e\ apply (drule_tac tendstoD[OF _ \e > 0\]) by (auto dest: simp: eventually_finite_subsets_at_top) moreover have \trunc_ell2 F \ \ cspan (range ket)\ using \finite F\ trunc_ell2_ket_cspan by blast ultimately show \\\\cspan (range ket). dist \ \ < e\ by auto qed lemma ccspan_range_ket[simp]: "ccspan (range ket) = (top::('a ell2 ccsubspace))" proof- have \closure (complex_vector.span (range ket)) = (UNIV::'a ell2 set)\ using Complex_L2.closed_cspan_range_ket by blast thus ?thesis by (simp add: ccspan.abs_eq top_ccsubspace.abs_eq) qed lemma cspan_range_ket_finite[simp]: "cspan (range ket :: 'a::finite ell2 set) = UNIV" by (metis closed_cspan_range_ket closure_finite_cspan finite_class.finite_UNIV finite_imageI) instance ell2 :: (finite) cfinite_dim proof define basis :: \'a ell2 set\ where \basis = range ket\ have \finite basis\ unfolding basis_def by simp moreover have \cspan basis = UNIV\ by (simp add: basis_def) ultimately show \\basis::'a ell2 set. finite basis \ cspan basis = UNIV\ by auto qed instantiation ell2 :: (enum) onb_enum begin definition "canonical_basis_ell2 = map ket Enum.enum" instance proof show "distinct (canonical_basis::'a ell2 list)" proof- have \finite (UNIV::'a set)\ by simp have \distinct (enum_class.enum::'a list)\ using enum_distinct by blast moreover have \inj_on ket (set enum_class.enum)\ by (meson inj_onI ket_injective) ultimately show ?thesis unfolding canonical_basis_ell2_def using distinct_map by blast qed show "is_ortho_set (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (z3) norm_ket f_inv_into_f is_ortho_set_def orthogonal_ket norm_zero) show "cindependent (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (verit, best) norm_ket f_inv_into_f is_ortho_set_def is_ortho_set_cindependent orthogonal_ket norm_zero) show "cspan (set (canonical_basis::'a ell2 list)) = UNIV" by (auto simp: canonical_basis_ell2_def enum_UNIV) show "norm (x::'a ell2) = 1" if "(x::'a ell2) \ set canonical_basis" for x :: "'a ell2" using that unfolding canonical_basis_ell2_def by auto qed - end lemma canonical_basis_length_ell2[code_unfold, simp]: "length (canonical_basis ::'a::enum ell2 list) = CARD('a)" unfolding canonical_basis_ell2_def apply simp using card_UNIV_length_enum by metis lemma ket_canonical_basis: "ket x = canonical_basis ! enum_idx x" proof- have "x = (enum_class.enum::'a list) ! enum_idx x" using enum_idx_correct[where i = x] by simp hence p1: "ket x = ket ((enum_class.enum::'a list) ! enum_idx x)" by simp have "enum_idx x < length (enum_class.enum::'a list)" using enum_idx_bound[where x = x]. hence "(map ket (enum_class.enum::'a list)) ! enum_idx x = ket ((enum_class.enum::'a list) ! enum_idx x)" by auto thus ?thesis unfolding canonical_basis_ell2_def using p1 by auto qed lemma clinear_equal_ket: fixes f g :: \'a::finite ell2 \ _\ assumes \clinear f\ assumes \clinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule complex_vector.linear_eq_on_span[where f=f and g=g and B=\range ket\]) using assms by auto lemma equal_ket: fixes A B :: \('a ell2, 'b::complex_normed_vector) cblinfun\ - assumes \\ x. cblinfun_apply A (ket x) = cblinfun_apply B (ket x)\ + 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[where G=\range ket\ and t=y, rotated 2]) using that by (auto intro!: bounded_linear_intros) then have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ for x y apply (rule bounded_antilinear_eq_on[where G=\range ket\ and t=x, rotated 2]) by (auto intro!: bounded_linear_intros) then show ?thesis by (rule adjoint_eqI) qed lemma ket_nonzero[simp]: "ket i \ 0" using norm_ket[of i] by force - lemma cindependent_ket: "cindependent (range (ket::'a\_))" proof- define S where "S = range (ket::'a\_)" have "is_ortho_set S" unfolding S_def is_ortho_set_def by auto moreover have "0 \ S" unfolding S_def using ket_nonzero by (simp add: image_iff) ultimately show ?thesis using is_ortho_set_cindependent[where A = S] unfolding S_def by blast qed lemma cdim_UNIV_ell2[simp]: \cdim (UNIV::'a::finite ell2 set) = CARD('a)\ apply (subst cspan_range_ket_finite[symmetric]) by (metis card_image cindependent_ket complex_vector.dim_span_eq_card_independent inj_ket) lemma is_ortho_set_ket[simp]: \is_ortho_set (range ket)\ using is_ortho_set_def by fastforce 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[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 ket_CARD_1_is_1: \ket x = 1\ for x :: \'a::CARD_1\ - apply transfer by simp - 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 subsection \Butterflies\ lemma cspan_butterfly_ket: \cspan {butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = UNIV\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cspan_butterfly_UNIV) by auto qed lemma cindependent_butterfly_ket: \cindependent {butterfly (ket i) (ket j)| (i::'b) (j::'a). True}\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b) (j::'a). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cindependent_butterfly) by auto qed lemma clinear_eq_butterfly_ketI: fixes F G :: \('a::finite ell2 \\<^sub>C\<^sub>L 'b::finite ell2) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes "\i j. F (butterfly (ket i) (ket j)) = G (butterfly (ket i) (ket j))" shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_ket) using assms by auto lemma sum_butterfly_ket[simp]: \(\(i::'a::finite)\UNIV. butterfly (ket i) (ket i)) = id_cblinfun\ apply (rule equal_ket) apply (subst complex_vector.linear_sum[where f=\\y. y *\<^sub>V ket _\]) apply (auto simp add: scaleC_cblinfun.rep_eq cblinfun.add_left clinearI butterfly_def cblinfun_compose_image cinner_ket) apply (subst sum.mono_neutral_cong_right[where S=\{_}\]) by auto subsection \One-dimensional spaces\ +instantiation ell2 :: (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 +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 \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/Complex_Vector_Spaces.thy b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy --- a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy +++ b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy @@ -1,3121 +1,3179 @@ section \\Complex_Vector_Spaces\ -- Complex Vector Spaces\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Vector_Spaces imports "HOL-Analysis.Elementary_Topology" "HOL-Analysis.Operator_Norm" "HOL-Analysis.Elementary_Normed_Spaces" "HOL-Library.Set_Algebras" "HOL-Analysis.Starlike" "HOL-Types_To_Sets.Types_To_Sets" "HOL-Library.Complemented_Lattices" + "HOL-Library.Function_Algebras" Extra_Vector_Spaces Extra_Ordered_Fields Extra_Operator_Norm Extra_General Complex_Vector_Spaces0 begin bundle notation_norm begin notation norm ("\_\") end unbundle lattice_syntax subsection \Misc\ (* Should rather be in Extra_Vector_Spaces but then complex_vector.span_image_scale' does not exist for some reason. Ideally this would replace Vector_Spaces.vector_space.span_image_scale. *) lemma (in vector_space) span_image_scale': \ \Strengthening of @{thm [source] vector_space.span_image_scale} without the condition \finite S\\ assumes nz: "\x. x \ S \ c x \ 0" shows "span ((\x. c x *s x) ` S) = span S" proof have \((\x. c x *s x) ` S) \ span S\ by (metis (mono_tags, lifting) image_subsetI in_mono local.span_superset local.subspace_scale local.subspace_span) then show \span ((\x. c x *s x) ` S) \ span S\ by (simp add: local.span_minimal) next have \x \ span ((\x. c x *s x) ` S)\ if \x \ S\ for x proof - have \x = inverse (c x) *s c x *s x\ by (simp add: nz that) moreover have \c x *s x \ (\x. c x *s x) ` S\ using that by blast ultimately show ?thesis by (metis local.span_base local.span_scale) qed then show \span S \ span ((\x. c x *s x) ` S)\ by (simp add: local.span_minimal subsetI) qed lemma (in scaleC) scaleC_real: assumes "r\\" shows "r *\<^sub>C x = Re r *\<^sub>R x" unfolding scaleR_scaleC using assms by simp lemma of_complex_of_real_eq [simp]: "of_complex (of_real n) = of_real n" unfolding of_complex_def of_real_def unfolding scaleR_scaleC by simp lemma Complexs_of_real [simp]: "of_real r \ \" unfolding Complexs_def of_real_def of_complex_def apply (subst scaleR_scaleC) by simp lemma Reals_in_Complexs: "\ \ \" unfolding Reals_def by auto lemma (in bounded_clinear) bounded_linear: "bounded_linear f" by standard lemma clinear_times: "clinear (\x. c * x)" for c :: "'a::complex_algebra" by (auto simp: clinearI distrib_left) lemma (in clinear) linear: \linear f\ by standard lemma bounded_clinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_clinear f" using assms by (auto intro!: exI bounded_clinear.intro clinearI simp: bounded_clinear_axioms_def) lemma bounded_clinear_id[simp]: \bounded_clinear id\ by (simp add: id_def) +lemma bounded_clinear_0[simp]: \bounded_clinear 0\ + by (auto intro!: bounded_clinearI[where K=0]) + definition cbilinear :: \('a::complex_vector \ 'b::complex_vector \ 'c::complex_vector) \ bool\ where \cbilinear = (\ f. (\ y. clinear (\ x. f x y)) \ (\ x. clinear (\ y. f x y)) )\ lemma cbilinear_add_left: assumes \cbilinear f\ shows \f (a + b) c = f a c + f b c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_add_right: assumes \cbilinear f\ shows \f a (b + c) = f a b + f a c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_times: fixes g' :: \'a::complex_vector \ complex\ and g :: \'b::complex_vector \ complex\ assumes \\ x y. h x y = (g' x)*(g y)\ and \clinear g\ and \clinear g'\ shows \cbilinear h\ proof - have w1: "h (b1 + b2) y = h b1 y + h b2 y" for b1 :: 'a and b2 :: 'a and y proof- have \h (b1 + b2) y = g' (b1 + b2) * g y\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = (g' b1 + g' b2) * g y\ using \clinear g'\ unfolding clinear_def by (simp add: assms(3) complex_vector.linear_add) also have \\ = g' b1 * g y + g' b2 * g y\ by (simp add: ring_class.ring_distribs(2)) also have \\ = h b1 y + h b2 y\ using assms(1) by auto finally show ?thesis by blast qed have w2: "h (r *\<^sub>C b) y = r *\<^sub>C h b y" for r :: complex and b :: 'a and y proof- have \h (r *\<^sub>C b) y = g' (r *\<^sub>C b) * g y\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' b * g y)\ by (simp add: assms(3) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h b y)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "clinear (\x. h x y)" for y :: 'b unfolding clinear_def by (meson clinearI clinear_def w1 w2) hence t2: "\y. clinear (\x. h x y)" by simp have v1: "h x (b1 + b2) = h x b1 + h x b2" for b1 :: 'b and b2 :: 'b and x proof- have \h x (b1 + b2) = g' x * g (b1 + b2)\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = g' x * (g b1 + g b2)\ using \clinear g'\ unfolding clinear_def by (simp add: assms(2) complex_vector.linear_add) also have \\ = g' x * g b1 + g' x * g b2\ by (simp add: ring_class.ring_distribs(1)) also have \\ = h x b1 + h x b2\ using assms(1) by auto finally show ?thesis by blast qed have v2: "h x (r *\<^sub>C b) = r *\<^sub>C h x b" for r :: complex and b :: 'b and x proof- have \h x (r *\<^sub>C b) = g' x * g (r *\<^sub>C b)\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' x * g b)\ by (simp add: assms(2) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h x b)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) (h x)" for x :: 'a using v1 v2 by (meson clinearI clinear_def) hence t1: "\x. clinear (h x)" unfolding clinear_def by simp show ?thesis unfolding cbilinear_def by (simp add: t1 t2) qed lemma csubspace_is_subspace: "csubspace A \ subspace A" apply (rule subspaceI) by (auto simp: complex_vector.subspace_def scaleR_scaleC) lemma span_subset_cspan: "span A \ cspan A" unfolding span_def complex_vector.span_def by (simp add: csubspace_is_subspace hull_antimono) lemma cindependent_implies_independent: assumes "cindependent (S::'a::complex_vector set)" shows "independent S" using assms unfolding dependent_def complex_vector.dependent_def using span_subset_cspan by blast lemma cspan_singleton: "cspan {x} = {\ *\<^sub>C x| \. True}" proof - have \cspan {x} = {y. y\cspan {x}}\ by auto also have \\ = {\ *\<^sub>C x| \. True}\ apply (subst complex_vector.span_breakdown_eq) by auto finally show ?thesis by - qed lemma cspan_as_span: "cspan (B::'a::complex_vector set) = span (B \ scaleC \ ` B)" proof (intro set_eqI iffI) let ?cspan = complex_vector.span let ?rspan = real_vector.span fix \ assume cspan: "\ \ ?cspan B" have "\B' r. finite B' \ B' \ B \ \ = (\b\B'. r b *\<^sub>C b)" using complex_vector.span_explicit[of B] cspan by auto then obtain B' r where "finite B'" and "B' \ B" and \_explicit: "\ = (\b\B'. r b *\<^sub>C b)" by atomize_elim define R where "R = B \ scaleC \ ` B" have x2: "(case x of (b, i) \ if i then Im (r b) *\<^sub>R \ *\<^sub>C b else Re (r b) *\<^sub>R b) \ span (B \ (*\<^sub>C) \ ` B)" if "x \ B' \ (UNIV::bool set)" for x :: "'a \ bool" using that \B' \ B\ by (auto simp add: real_vector.span_base real_vector.span_scale subset_iff) have x1: "\ = (\x\B'. \i\UNIV. if i then Im (r x) *\<^sub>R \ *\<^sub>C x else Re (r x) *\<^sub>R x)" if "\b. r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" using that by (simp add: UNIV_bool \_explicit) moreover have "r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" for b using complex_eq scaleC_add_left scaleC_scaleC scaleR_scaleC by (metis (no_types, lifting) complex_of_real_i i_complex_of_real) ultimately have "\ = (\(b,i)\(B'\UNIV). if i then Im (r b) *\<^sub>R (\ *\<^sub>C b) else Re (r b) *\<^sub>R b)" by (simp add: sum.cartesian_product) also have "\ \ ?rspan R" unfolding R_def using x2 by (rule real_vector.span_sum) finally show "\ \ ?rspan R" by - next let ?cspan = complex_vector.span let ?rspan = real_vector.span define R where "R = B \ scaleC \ ` B" fix \ assume rspan: "\ \ ?rspan R" have "subspace {a. a \ cspan B}" by (rule real_vector.subspaceI, auto simp add: complex_vector.span_zero complex_vector.span_add_eq2 complex_vector.span_scale scaleR_scaleC) moreover have "x \ cspan B" if "x \ R" for x :: 'a using that R_def complex_vector.span_base complex_vector.span_scale by fastforce ultimately show "\ \ ?cspan B" using real_vector.span_induct rspan by blast qed lemma isomorphic_equal_cdim: assumes lin_f: \clinear f\ assumes inj_f: \inj_on f (cspan S)\ assumes im_S: \f ` S = T\ shows \cdim S = cdim T\ proof - obtain SB where SB_span: "cspan SB = cspan S" and indep_SB: \cindependent SB\ by (metis complex_vector.basis_exists complex_vector.span_mono complex_vector.span_span subset_antisym) with lin_f inj_f have indep_fSB: \cindependent (f ` SB)\ apply (rule_tac complex_vector.linear_independent_injective_image) by auto from lin_f have \cspan (f ` SB) = f ` cspan SB\ by (meson complex_vector.linear_span_image) also from SB_span lin_f have \\ = cspan T\ by (metis complex_vector.linear_span_image im_S) finally have \cdim T = card (f ` SB)\ using indep_fSB complex_vector.dim_eq_card by blast also have \\ = card SB\ apply (rule card_image) using inj_f by (metis SB_span complex_vector.linear_inj_on_span_iff_independent_image indep_fSB lin_f) also have \\ = cdim S\ using indep_SB SB_span by (metis complex_vector.dim_eq_card) finally show ?thesis by simp qed lemma cindependent_inter_scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 1" shows "B \ (*\<^sub>C) c ` B = {}" proof (rule classical, cases \c = 0\) case True then show ?thesis using a1 by (auto simp add: complex_vector.dependent_zero) next case False assume "\(B \ (*\<^sub>C) c ` B = {})" hence "B \ (*\<^sub>C) c ` B \ {}" by blast then obtain x where u1: "x \ B \ (*\<^sub>C) c ` B" by blast then obtain b where u2: "x = b" and u3: "b\B" by blast then obtain b' where u2': "x = c *\<^sub>C b'" and u3': "b'\B" using u1 by blast have g1: "b = c *\<^sub>C b'" using u2 and u2' by simp hence "b \ complex_vector.span {b'}" using False by (simp add: complex_vector.span_base complex_vector.span_scale) hence "b = b'" by (metis u3' a1 complex_vector.dependent_def complex_vector.span_base complex_vector.span_scale insertE insert_Diff u2 u2' u3) hence "b' = c *\<^sub>C b'" using g1 by blast thus ?thesis by (metis a1 a3 complex_vector.dependent_zero complex_vector.scale_right_imp_eq mult_cancel_right2 scaleC_scaleC u3') qed lemma real_independent_from_complex_independent: assumes "cindependent (B::'a::complex_vector set)" defines "B' == ((*\<^sub>C) \ ` B)" shows "independent (B \ B')" proof (rule notI) assume \dependent (B \ B')\ then obtain T f0 x where [simp]: \finite T\ and \T \ B \ B'\ and f0_sum: \(\v\T. f0 v *\<^sub>R v) = 0\ and x: \x \ T\ and f0_x: \f0 x \ 0\ by (auto simp: real_vector.dependent_explicit) define f T1 T2 T' f' x' where \f v = (if v \ T then f0 v else 0)\ and \T1 = T \ B\ and \T2 = scaleC (-\) ` (T \ B')\ and \T' = T1 \ T2\ and \f' v = f v + \ * f (\ *\<^sub>C v)\ and \x' = (if x \ T1 then x else -\ *\<^sub>C x)\ for v have \B \ B' = {}\ by (simp add: assms cindependent_inter_scaleC_cindependent) have \T' \ B\ by (auto simp: T'_def T1_def T2_def B'_def) have [simp]: \finite T'\ \finite T1\ \finite T2\ by (auto simp add: T'_def T1_def T2_def) have f_sum: \(\v\T. f v *\<^sub>R v) = 0\ unfolding f_def using f0_sum by auto have f_x: \f x \ 0\ using f0_x x by (auto simp: f_def) have f'_sum: \(\v\T'. f' v *\<^sub>C v) = 0\ proof - have \(\v\T'. f' v *\<^sub>C v) = (\v\T'. complex_of_real (f v) *\<^sub>C v) + (\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ by (auto simp: f'_def sum.distrib scaleC_add_left) also have \(\v\T'. complex_of_real (f v) *\<^sub>C v) = (\v\T1. f v *\<^sub>R v)\ (is \_ = ?left\) apply (auto simp: T'_def scaleR_scaleC intro!: sum.mono_neutral_cong_right) using T'_def T1_def \T' \ B\ f_def by auto also have \(\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v) = (\v\T2. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ (is \_ = ?right\) apply (auto simp: T'_def intro!: sum.mono_neutral_cong_right) by (smt (z3) B'_def IntE IntI T1_def T2_def \f \ \v. if v \ T then f0 v else 0\ add.inverse_inverse complex_vector.vector_space_axioms i_squared imageI mult_minus_left vector_space.vector_space_assms(3) vector_space.vector_space_assms(4)) also have \?right = (\v\T\B'. f v *\<^sub>R v)\ (is \_ = ?right\) apply (rule sum.reindex_cong[symmetric, where l=\scaleC \\]) apply (auto simp: T2_def image_image scaleR_scaleC) using inj_on_def by fastforce also have \?left + ?right = (\v\T. f v *\<^sub>R v)\ apply (subst sum.union_disjoint[symmetric]) using \B \ B' = {}\ \T \ B \ B'\ apply (auto simp: T1_def) by (metis Int_Un_distrib Un_Int_eq(4) sup.absorb_iff1) also have \\ = 0\ by (rule f_sum) finally show ?thesis by - qed have x': \x' \ T'\ using \T \ B \ B'\ x by (auto simp: x'_def T'_def T1_def T2_def) have f'_x': \f' x' \ 0\ using Complex_eq Complex_eq_0 f'_def f_x x'_def by auto from \finite T'\ \T' \ B\ f'_sum x' f'_x' have \cdependent B\ using complex_vector.independent_explicit_module by blast with assms show False by auto qed lemma crepresentation_from_representation: assumes a1: "cindependent B" and a2: "b \ B" and a3: "finite B" shows "crepresentation B \ b = (representation (B \ (*\<^sub>C) \ ` B) \ b) + \ *\<^sub>C (representation (B \ (*\<^sub>C) \ ` B) \ (\ *\<^sub>C b))" proof (cases "\ \ cspan B") define B' where "B' = B \ (*\<^sub>C) \ ` B" case True define r where "r v = real_vector.representation B' \ v" for v define r' where "r' v = real_vector.representation B' \ (\ *\<^sub>C v)" for v define f where "f v = r v + \ *\<^sub>C r' v" for v define g where "g v = crepresentation B \ v" for v have "(\v | g v \ 0. g v *\<^sub>C v) = \" unfolding g_def using Collect_cong Collect_mono_iff DiffD1 DiffD2 True a1 complex_vector.finite_representation complex_vector.sum_nonzero_representation_eq sum.mono_neutral_cong_left by fastforce moreover have "finite {v. g v \ 0}" unfolding g_def by (simp add: complex_vector.finite_representation) moreover have "v \ B" if "g v \ 0" for v using that unfolding g_def by (simp add: complex_vector.representation_ne_zero) ultimately have rep1: "(\v\B. g v *\<^sub>C v) = \" unfolding g_def using a3 True a1 complex_vector.sum_representation_eq by blast have l0': "inj ((*\<^sub>C) \::'a \'a)" unfolding inj_def by simp have l0: "inj ((*\<^sub>C) (- \)::'a \'a)" unfolding inj_def by simp have l1: "(*\<^sub>C) (- \) ` B \ B = {}" using cindependent_inter_scaleC_cindependent[where B=B and c = "- \"] by (metis Int_commute a1 add.inverse_inverse complex_i_not_one i_squared mult_cancel_left1 neg_equal_0_iff_equal) have l2: "B \ (*\<^sub>C) \ ` B = {}" by (simp add: a1 cindependent_inter_scaleC_cindependent) have rr1: "r (\ *\<^sub>C v) = r' v" for v unfolding r_def r'_def by simp have k1: "independent B'" unfolding B'_def using a1 real_independent_from_complex_independent by simp have "\ \ span B'" using B'_def True cspan_as_span by blast have "v \ B'" if "r v \ 0" for v unfolding r_def using r_def real_vector.representation_ne_zero that by auto have "finite B'" unfolding B'_def using a3 by simp have "(\v\B'. r v *\<^sub>R v) = \" unfolding r_def using True Real_Vector_Spaces.real_vector.sum_representation_eq[where B = B' and basis = B' and v = \] by (smt Real_Vector_Spaces.dependent_raw_def \\ \ Real_Vector_Spaces.span B'\ \finite B'\ equalityD2 k1) have d1: "(\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v)) = (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using l0' by (metis (mono_tags, lifting) inj_eq inj_on_def sum.reindex_cong) have "(\v\B. (r v + \ * (r' v)) *\<^sub>C v) = (\v\B. r v *\<^sub>C v + (\ * r' v) *\<^sub>C v)" by (meson scaleC_left.add) also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. (\ * r' v) *\<^sub>C v)" using sum.distrib by fastforce also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. \ *\<^sub>C (r' v *\<^sub>C v))" by auto also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. \ *\<^sub>C (r (\ *\<^sub>C v) *\<^sub>R v))" unfolding r'_def r_def by (metis (mono_tags, lifting) scaleR_scaleC sum.cong) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v))" by (metis (no_types, lifting) complex_vector.scale_left_commute scaleR_scaleC) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using d1 by simp also have "\ = \" using l2 \(\v\B'. r v *\<^sub>R v) = \\ unfolding B'_def by (simp add: a3 sum.union_disjoint) finally have "(\v\B. f v *\<^sub>C v) = \" unfolding r'_def r_def f_def by simp hence "0 = (\v\B. f v *\<^sub>C v) - (\v\B. crepresentation B \ v *\<^sub>C v)" using rep1 unfolding g_def by simp also have "\ = (\v\B. f v *\<^sub>C v - crepresentation B \ v *\<^sub>C v)" by (simp add: sum_subtractf) also have "\ = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)" by (metis scaleC_left.diff) finally have "0 = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)". hence "(\v\B. (f v - crepresentation B \ v) *\<^sub>C v) = 0" by simp hence "f b - crepresentation B \ b = 0" using a1 a2 a3 complex_vector.independentD[where s = B and t = B and u = "\v. f v - crepresentation B \ v" and v = b] order_refl by smt hence "crepresentation B \ b = f b" by simp thus ?thesis unfolding f_def r_def r'_def B'_def by auto next define B' where "B' = B \ (*\<^sub>C) \ ` B" case False have b2: "\ \ real_vector.span B'" unfolding B'_def using False cspan_as_span by auto have "\ \ complex_vector.span B" using False by blast have "crepresentation B \ b = 0" unfolding complex_vector.representation_def by (simp add: False) moreover have "real_vector.representation B' \ b = 0" unfolding real_vector.representation_def by (simp add: b2) moreover have "real_vector.representation B' \ ((*\<^sub>C) \ b) = 0" unfolding real_vector.representation_def by (simp add: b2) ultimately show ?thesis unfolding B'_def by simp qed lemma CARD_1_vec_0[simp]: \(\ :: _ ::{complex_vector,CARD_1}) = 0\ by auto lemma scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 0" shows "cindependent ((*\<^sub>C) c ` B)" proof- have "u y = 0" if g1: "y\S" and g2: "(\x\S. u x *\<^sub>C x) = 0" and g3: "finite S" and g4: "S\(*\<^sub>C) c ` B" for u y S proof- define v where "v x = u (c *\<^sub>C x)" for x obtain S' where "S'\B" and S_S': "S = (*\<^sub>C) c ` S'" by (meson g4 subset_imageE) have "inj ((*\<^sub>C) c::'a\_)" unfolding inj_def using a3 by auto hence "finite S'" using S_S' finite_imageD g3 subset_inj_on by blast have "t \ (*\<^sub>C) (inverse c) ` S" if "t \ S'" for t proof- have "c *\<^sub>C t \ S" using \S = (*\<^sub>C) c ` S'\ that by blast hence "(inverse c) *\<^sub>C (c *\<^sub>C t) \ (*\<^sub>C) (inverse c) ` S" by blast moreover have "(inverse c) *\<^sub>C (c *\<^sub>C t) = t" by (simp add: a3) ultimately show ?thesis by simp qed moreover have "t \ S'" if "t \ (*\<^sub>C) (inverse c) ` S" for t proof- obtain t' where "t = (inverse c) *\<^sub>C t'" and "t' \ S" using \t \ (*\<^sub>C) (inverse c) ` S\ by auto have "c *\<^sub>C t = c *\<^sub>C ((inverse c) *\<^sub>C t')" using \t = (inverse c) *\<^sub>C t'\ by simp also have "\ = (c * (inverse c)) *\<^sub>C t'" by simp also have "\ = t'" by (simp add: a3) finally have "c *\<^sub>C t = t'". thus ?thesis using \t' \ S\ using \S = (*\<^sub>C) c ` S'\ a3 complex_vector.scale_left_imp_eq by blast qed ultimately have "S' = (*\<^sub>C) (inverse c) ` S" by blast hence "inverse c *\<^sub>C y \ S'" using that(1) by blast have t: "inj (((*\<^sub>C) c)::'a \ _)" using a3 complex_vector.injective_scale[where c = c] by blast have "0 = (\x\(*\<^sub>C) c ` S'. u x *\<^sub>C x)" using \S = (*\<^sub>C) c ` S'\ that(2) by auto also have "\ = (\x\S'. v x *\<^sub>C (c *\<^sub>C x))" unfolding v_def using t Groups_Big.comm_monoid_add_class.sum.reindex[where h = "((*\<^sub>C) c)" and A = S' and g = "\x. u x *\<^sub>C x"] subset_inj_on by auto also have "\ = c *\<^sub>C (\x\S'. v x *\<^sub>C x)" by (metis (mono_tags, lifting) complex_vector.scale_left_commute scaleC_right.sum sum.cong) finally have "0 = c *\<^sub>C (\x\S'. v x *\<^sub>C x)". hence "(\x\S'. v x *\<^sub>C x) = 0" using a3 by auto hence "v (inverse c *\<^sub>C y) = 0" using \inverse c *\<^sub>C y \ S'\ \finite S'\ \S' \ B\ a1 complex_vector.independentD by blast thus "u y = 0" unfolding v_def by (simp add: a3) qed thus ?thesis using complex_vector.dependent_explicit by (simp add: complex_vector.dependent_explicit ) qed lemma cspan_eqI: assumes \\a. a\A \ a\cspan B\ assumes \\b. b\B \ b\cspan A\ shows \cspan A = cspan B\ apply (rule complex_vector.span_subspace[rotated]) apply (rule complex_vector.span_minimal) using assms by auto lemma (in bounded_cbilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by standard lemma norm_scaleC_sgn[simp]: \complex_of_real (norm \) *\<^sub>C sgn \ = \\ for \ :: "'a::complex_normed_vector" apply (cases \\ = 0\) by (auto simp: sgn_div_norm scaleR_scaleC) lemma scaleC_of_complex[simp]: \scaleC x (of_complex y) = of_complex (x * y)\ unfolding of_complex_def using scaleC_scaleC by blast lemma bounded_clinear_inv: assumes [simp]: \bounded_clinear f\ assumes b: \b > 0\ assumes bound: \\x. norm (f x) \ b * norm x\ assumes \surj f\ shows \bounded_clinear (inv f)\ proof (rule bounded_clinear_intro) fix x y :: 'b and r :: complex define x' y' where \x' = inv f x\ and \y' = inv f y\ have [simp]: \clinear f\ by (simp add: bounded_clinear.clinear) have [simp]: \inj f\ proof (rule injI) fix x y assume \f x = f y\ then have \norm (f (x - y)) = 0\ by (simp add: complex_vector.linear_diff) with bound b have \norm (x - y) = 0\ by (metis linorder_not_le mult_le_0_iff nle_le norm_ge_zero) then show \x = y\ by simp qed from \surj f\ have [simp]: \x = f x'\ \y = f y'\ by (simp_all add: surj_f_inv_f x'_def y'_def) show "inv f (x + y) = inv f x + inv f y" by (simp flip: complex_vector.linear_add) show "inv f (r *\<^sub>C x) = r *\<^sub>C inv f x" by (simp flip: clinear.scaleC) from bound have "b * norm (inv f x) \ norm x" by (simp flip: clinear.scaleC) with b show "norm (inv f x) \ norm x * inverse b" by (smt (verit, ccfv_threshold) left_inverse mult.commute mult_cancel_right1 mult_le_cancel_left_pos vector_space_over_itself.scale_scale) qed lemma range_is_csubspace[simp]: assumes a1: "clinear f" shows "csubspace (range f)" using assms complex_vector.linear_subspace_image complex_vector.subspace_UNIV by blast lemma csubspace_is_convex[simp]: assumes a1: "csubspace M" shows "convex M" proof- have \\x\M. \y\ M. \u. \v. u *\<^sub>C x + v *\<^sub>C y \ M\ using a1 by (simp add: complex_vector.subspace_def) hence \\x\M. \y\M. \u::real. \v::real. u *\<^sub>R x + v *\<^sub>R y \ M\ by (simp add: scaleR_scaleC) hence \\x\M. \y\M. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \M\ by blast thus ?thesis using convex_def by blast qed lemma kernel_is_csubspace[simp]: assumes a1: "clinear f" shows "csubspace (f -` {0})" by (simp add: assms complex_vector.linear_subspace_vimage) +lemma bounded_cbilinear_0[simp]: \bounded_cbilinear (\_ _. 0)\ + by (auto intro!: bounded_cbilinear.intro exI[where x=0]) +lemma bounded_cbilinear_0'[simp]: \bounded_cbilinear 0\ + by (auto intro!: bounded_cbilinear.intro exI[where x=0]) + +lemma bounded_cbilinear_apply_bounded_clinear: \bounded_clinear (f x)\ if \bounded_cbilinear f\ +proof - + interpret f: bounded_cbilinear f + by (fact that) + from f.bounded obtain K where \norm (f a b) \ norm a * norm b * K\ for a b + by auto + then show ?thesis + by (auto intro!: bounded_clinearI[where K=\norm x * K\] + simp add: f.add_right f.scaleC_right mult.commute mult.left_commute) +qed + subsection \Antilinear maps and friends\ locale antilinear = additive f for f :: "'a::complex_vector \ 'b::complex_vector" + assumes scaleC: "f (scaleC r x) = cnj r *\<^sub>C f x" sublocale antilinear \ linear proof (rule linearI) show "f (b1 + b2) = f b1 + f b2" for b1 :: 'a and b2 :: 'a by (simp add: add) show "f (r *\<^sub>R b) = r *\<^sub>R f b" for r :: real and b :: 'a unfolding scaleR_scaleC by (subst scaleC, simp) qed lemma antilinear_imp_scaleC: fixes D :: "complex \ 'a::complex_vector" assumes "antilinear D" obtains d where "D = (\x. cnj x *\<^sub>C d)" proof - interpret clinear "D o cnj" apply standard apply auto apply (simp add: additive.add assms antilinear.axioms(1)) using assms antilinear.scaleC by fastforce obtain d where "D o cnj = (\x. x *\<^sub>C d)" using clinear_axioms complex_vector.linear_imp_scale by blast then have \D = (\x. cnj x *\<^sub>C d)\ by (metis comp_apply complex_cnj_cnj) then show ?thesis by (rule that) qed corollary complex_antilinearD: fixes f :: "complex \ complex" assumes "antilinear f" obtains c where "f = (\x. c * cnj x)" by (rule antilinear_imp_scaleC [OF assms]) (force simp: scaleC_conv_of_complex) lemma antilinearI: assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>C x) = cnj c *\<^sub>C f x" shows "antilinear f" by standard (rule assms)+ lemma antilinear_o_antilinear: "antilinear f \ antilinear g \ clinear (g o f)" apply (rule clinearI) apply (simp add: additive.add antilinear_def) by (simp add: antilinear.scaleC) lemma clinear_o_antilinear: "antilinear f \ clinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) lemma antilinear_o_clinear: "clinear f \ antilinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) locale bounded_antilinear = antilinear f for f :: "'a::complex_normed_vector \ 'b::complex_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" lemma bounded_antilinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = cnj r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_antilinear f" using assms by (auto intro!: exI bounded_antilinear.intro antilinearI simp: bounded_antilinear_axioms_def) sublocale bounded_antilinear \ real: bounded_linear \ \Gives access to all lemmas from \<^locale>\linear\ using prefix \real.\\ apply standard by (fact bounded) lemma (in bounded_antilinear) bounded_linear: "bounded_linear f" by (fact real.bounded_linear) lemma (in bounded_antilinear) antilinear: "antilinear f" by (fact antilinear_axioms) lemma bounded_antilinear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleC r x) = scaleC (cnj r) (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_antilinear f" by standard (blast intro: assms)+ lemma bounded_antilinear_0[simp]: \bounded_antilinear (\_. 0)\ - by (rule bounded_antilinear_intro[where K=0], auto) + by (auto intro!: bounded_antilinearI[where K=0]) + +lemma bounded_antilinear_0'[simp]: \bounded_antilinear 0\ + by (auto intro!: bounded_antilinearI[where K=0]) lemma cnj_bounded_antilinear[simp]: "bounded_antilinear cnj" apply (rule bounded_antilinear_intro [where K = 1]) by auto lemma bounded_antilinear_o_bounded_antilinear: assumes "bounded_antilinear f" and "bounded_antilinear g" shows "bounded_clinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_antilinear g by fact fix b1 b2 b r show "f (g (b1 + b2)) = f (g b1) + f (g b2)" by (simp add: f.add g.add) show "f (g (r *\<^sub>C b)) = r *\<^sub>C f (g b)" by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed +lemma bounded_antilinear_o_bounded_antilinear': + assumes "bounded_antilinear f" + and "bounded_antilinear g" + shows "bounded_clinear (g o f)" + using assms by (simp add: o_def bounded_antilinear_o_bounded_antilinear) + lemma bounded_antilinear_o_bounded_clinear: assumes "bounded_antilinear f" and "bounded_clinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_clinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed +lemma bounded_antilinear_o_bounded_clinear': + assumes "bounded_clinear f" + and "bounded_antilinear g" + shows "bounded_antilinear (g o f)" + using assms by (simp add: o_def bounded_antilinear_o_bounded_clinear) + lemma bounded_clinear_o_bounded_antilinear: assumes "bounded_clinear f" and "bounded_antilinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_clinear f by fact interpret g: bounded_antilinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x using f.scaleC g.scaleC by fastforce have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed +lemma bounded_clinear_o_bounded_antilinear': + assumes "bounded_antilinear f" + and "bounded_clinear g" + shows "bounded_antilinear (g o f)" + using assms by (simp add: o_def bounded_clinear_o_bounded_antilinear) + lemma bij_clinear_imp_inv_clinear: "clinear (inv f)" if a1: "clinear f" and a2: "bij f" proof fix b1 b2 r b show "inv f (b1 + b2) = inv f b1 + inv f b2" by (simp add: a1 a2 bij_is_inj bij_is_surj complex_vector.linear_add inv_f_eq surj_f_inv_f) show "inv f (r *\<^sub>C b) = r *\<^sub>C inv f b" using that by (smt bij_inv_eq_iff clinear_def complex_vector.linear_scale) qed locale bounded_sesquilinear = fixes prod :: "'a::complex_normed_vector \ 'b::complex_normed_vector \ 'c::complex_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleC_left: "prod (r *\<^sub>C a) b = (cnj r) *\<^sub>C (prod a b)" and scaleC_right: "prod a (r *\<^sub>C b) = r *\<^sub>C (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" sublocale bounded_sesquilinear \ real: bounded_bilinear \ \Gives access to all lemmas from \<^locale>\linear\ using prefix \real.\\ apply standard by (auto simp: add_left add_right scaleC_left scaleC_right bounded scaleR_scaleC) lemma (in bounded_sesquilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by intro_locales lemma (in bounded_sesquilinear) bounded_antilinear_left: "bounded_antilinear (\a. prod a b)" apply standard apply (auto simp add: scaleC_left add_left) by (metis ab_semigroup_mult_class.mult_ac(1) bounded) lemma (in bounded_sesquilinear) bounded_clinear_right: "bounded_clinear (\b. prod a b)" apply standard apply (auto simp add: scaleC_right add_right) by (metis ab_semigroup_mult_class.mult_ac(1) ordered_field_class.sign_simps(34) real.pos_bounded) lemma (in bounded_sesquilinear) comp1: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x. prod (g x))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (g (a + a')) b = prod (g a) b + prod (g a') b" by (simp add: add add_left) show "prod (g a) (b + b') = prod (g a) b + prod (g a) b'" by (simp add: add add_right) show "prod (g (r *\<^sub>C a)) b = cnj r *\<^sub>C prod (g a) b" by (simp add: scaleC scaleC_left) show "prod (g a) (r *\<^sub>C b) = r *\<^sub>C prod (g a) b" by (simp add: scaleC_right) interpret bi: bounded_bilinear \(\x. prod (g x))\ by (simp add: bounded_linear real.comp1) show "\K. \a b. norm (prod (g a) b) \ norm a * norm b * K" using bi.bounded by blast qed lemma (in bounded_sesquilinear) comp2: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x y. prod x (g y))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (a + a') (g b) = prod a (g b) + prod a' (g b)" by (simp add: add add_left) show "prod a (g (b + b')) = prod a (g b) + prod a (g b')" by (simp add: add add_right) show "prod (r *\<^sub>C a) (g b) = cnj r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_left) show "prod a (g (r *\<^sub>C b)) = r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_right) interpret bi: bounded_bilinear \(\x y. prod x (g y))\ apply (rule bounded_bilinear.flip) using _ bounded_linear apply (rule bounded_bilinear.comp1) using bounded_bilinear by (rule bounded_bilinear.flip) show "\K. \a b. norm (prod a (g b)) \ norm a * norm b * K" using bi.bounded by blast qed lemma (in bounded_sesquilinear) comp: "bounded_clinear f \ bounded_clinear g \ bounded_sesquilinear (\x y. prod (f x) (g y))" using comp1 bounded_sesquilinear.comp2 by auto lemma bounded_clinear_const_scaleR: fixes c :: real assumes \bounded_clinear f\ shows \bounded_clinear (\ x. c *\<^sub>R f x )\ proof- have \bounded_clinear (\ x. (complex_of_real c) *\<^sub>C f x )\ by (simp add: assms bounded_clinear_const_scaleC) thus ?thesis by (simp add: scaleR_scaleC) qed lemma bounded_linear_bounded_clinear: \bounded_linear A \ \c x. A (c *\<^sub>C x) = c *\<^sub>C A x \ bounded_clinear A\ apply standard by (simp_all add: linear_simps bounded_linear.bounded) lemma comp_bounded_clinear: fixes A :: \'b::complex_normed_vector \ 'c::complex_normed_vector\ and B :: \'a::complex_normed_vector \ 'b\ assumes \bounded_clinear A\ and \bounded_clinear B\ shows \bounded_clinear (A \ B)\ by (metis clinear_compose assms(1) assms(2) bounded_clinear_axioms_def bounded_clinear_compose bounded_clinear_def o_def) lemma bounded_sesquilinear_add: \bounded_sesquilinear (\ x y. A x y + B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "A (a + a') b + B (a + a') b = (A a b + B a b) + (A a' b + B a' b)" by (simp add: bounded_sesquilinear.add_left that(1) that(2)) show \A a (b + b') + B a (b + b') = (A a b + B a b) + (A a b' + B a b')\ by (simp add: bounded_sesquilinear.add_right that(1) that(2)) show \A (r *\<^sub>C a) b + B (r *\<^sub>C a) b = cnj r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_left scaleC_add_right that(1) that(2)) show \A a (r *\<^sub>C b) + B a (r *\<^sub>C b) = r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_right scaleC_add_right that(1) that(2)) show \\K. \a b. norm (A a b + B a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \\ KB. \ a b. norm (B a b) \ norm a * norm b * KB\ by (simp add: bounded_sesquilinear.bounded that(2)) then obtain KB where \\ a b. norm (B a b) \ norm a * norm b * KB\ by blast have \norm (A a b + B a b) \ norm a * norm b * (KA + KB)\ for a b proof- have \norm (A a b + B a b) \ norm (A a b) + norm (B a b)\ using norm_triangle_ineq by blast also have \\ \ norm a * norm b * KA + norm a * norm b * KB\ using \\ a b. norm (A a b) \ norm a * norm b * KA\ \\ a b. norm (B a b) \ norm a * norm b * KB\ using add_mono by blast also have \\= norm a * norm b * (KA + KB)\ by (simp add: mult.commute ring_class.ring_distribs(2)) finally show ?thesis by blast qed thus ?thesis by blast qed qed lemma bounded_sesquilinear_uminus: \bounded_sesquilinear (\ x y. - A x y)\ if \bounded_sesquilinear A\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "- A (a + a') b = (- A a b) + (- A a' b)" by (simp add: bounded_sesquilinear.add_left that) show \- A a (b + b') = (- A a b) + (- A a b')\ by (simp add: bounded_sesquilinear.add_right that) show \- A (r *\<^sub>C a) b = cnj r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_left that) show \- A a (r *\<^sub>C b) = r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_right that) show \\K. \a b. norm (- A a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \norm (- A a b) \ norm a * norm b * KA\ for a b by (simp add: \\a b. norm (A a b) \ norm a * norm b * KA\) thus ?thesis by blast qed qed lemma bounded_sesquilinear_diff: \bounded_sesquilinear (\ x y. A x y - B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof - have \bounded_sesquilinear (\ x y. - B x y)\ using that(2) by (rule bounded_sesquilinear_uminus) then have \bounded_sesquilinear (\ x y. A x y + (- B x y))\ using that(1) by (rule bounded_sesquilinear_add[rotated]) then show ?thesis by auto qed lemmas isCont_scaleC [simp] = bounded_bilinear.isCont [OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] +lemma bounded_sesquilinear_0[simp]: \bounded_sesquilinear (\_ _.0)\ + by (auto intro!: bounded_sesquilinear.intro exI[where x=0]) + +lemma bounded_sesquilinear_0'[simp]: \bounded_sesquilinear 0\ + by (auto intro!: bounded_sesquilinear.intro exI[where x=0]) + +lemma bounded_sesquilinear_apply_bounded_clinear: \bounded_clinear (f x)\ if \bounded_sesquilinear f\ +proof - + interpret f: bounded_sesquilinear f + by (fact that) + from f.bounded obtain K where \norm (f a b) \ norm a * norm b * K\ for a b + by auto + then show ?thesis + by (auto intro!: bounded_clinearI[where K=\norm x * K\] + simp add: f.add_right f.scaleC_right mult.commute mult.left_commute) +qed + subsection \Misc 2\ lemma summable_on_scaleC_left [intro]: fixes c :: \'a :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. f x *\<^sub>C c) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>C c) = (\y. y *\<^sub>C c) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleC_left.additive_axioms) lemma summable_on_scaleC_right [intro]: fixes f :: \'a \ 'b :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. c *\<^sub>C f x) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. c *\<^sub>C f x) = (\y. c *\<^sub>C y) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleC_right.additive_axioms) lemma infsum_scaleC_left: fixes c :: \'a :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "infsum (\x. f x *\<^sub>C c) A = infsum f A *\<^sub>C c" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>C c) = (\y. y *\<^sub>C c) o f\], simp add: o_def) apply (rule infsum_comm_additive) using assms by (auto simp add: scaleC_left.additive_axioms) lemma infsum_scaleC_right: fixes f :: \'a \ 'b :: complex_normed_vector\ shows "infsum (\x. c *\<^sub>C f x) A = c *\<^sub>C infsum f A" proof - consider (summable) \f summable_on A\ | (c0) \c = 0\ | (not_summable) \\ f summable_on A\ \c \ 0\ by auto then show ?thesis proof cases case summable then show ?thesis apply (subst asm_rl[of \(\x. c *\<^sub>C f x) = (\y. c *\<^sub>C y) o f\], simp add: o_def) apply (rule infsum_comm_additive) using summable by (auto simp add: scaleC_right.additive_axioms) next case c0 then show ?thesis by auto next case not_summable have \\ (\x. c *\<^sub>C f x) summable_on A\ proof (rule notI) assume \(\x. c *\<^sub>C f x) summable_on A\ then have \(\x. inverse c *\<^sub>C c *\<^sub>C f x) summable_on A\ using summable_on_scaleC_right by blast then have \f summable_on A\ using not_summable by auto with not_summable show False by simp qed then show ?thesis by (simp add: infsum_not_exists not_summable(1)) qed qed lemmas sums_of_complex = bounded_linear.sums [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas summable_of_complex = bounded_linear.summable [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas suminf_of_complex = bounded_linear.suminf [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_left = bounded_linear.sums[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_left = bounded_linear.summable[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_left = bounded_linear.suminf[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_right = bounded_linear.sums[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_right = bounded_linear.summable[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_right = bounded_linear.suminf[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemma closed_scaleC: fixes S::\'a::complex_normed_vector set\ and a :: complex assumes \closed S\ shows \closed ((*\<^sub>C) a ` S)\ proof (cases \a = 0\) case True then show ?thesis apply (cases \S = {}\) by (auto simp: image_constant) next case False then have \(*\<^sub>C) a ` S = (*\<^sub>C) (inverse a) -` S\ by (auto simp add: rev_image_eqI) moreover have \closed ((*\<^sub>C) (inverse a) -` S)\ by (simp add: assms continuous_closed_vimage) ultimately show ?thesis by simp qed lemma closure_scaleC: fixes S::\'a::complex_normed_vector set\ shows \closure ((*\<^sub>C) a ` S) = (*\<^sub>C) a ` closure S\ proof have \closed (closure S)\ by simp show "closure ((*\<^sub>C) a ` S) \ (*\<^sub>C) a ` closure S" by (simp add: closed_scaleC closure_minimal closure_subset image_mono) have "x \ closure ((*\<^sub>C) a ` S)" if "x \ (*\<^sub>C) a ` closure S" for x :: 'a proof- obtain t where \x = ((*\<^sub>C) a) t\ and \t \ closure S\ using \x \ (*\<^sub>C) a ` closure S\ by auto have \\s. (\n. s n \ S) \ s \ t\ using \t \ closure S\ Elementary_Topology.closure_sequential by blast then obtain s where \\n. s n \ S\ and \s \ t\ by blast have \(\ n. scaleC a (s n) \ ((*\<^sub>C) a ` S))\ using \\n. s n \ S\ by blast moreover have \(\ n. scaleC a (s n)) \ x\ proof- have \isCont (scaleC a) t\ by simp thus ?thesis using \s \ t\ \x = ((*\<^sub>C) a) t\ by (simp add: isCont_tendsto_compose) qed ultimately show ?thesis using Elementary_Topology.closure_sequential by metis qed thus "(*\<^sub>C) a ` closure S \ closure ((*\<^sub>C) a ` S)" by blast qed lemma onorm_scalarC: fixes f :: \'a::complex_normed_vector \ 'b::complex_normed_vector\ assumes a1: \bounded_clinear f\ shows \onorm (\ x. r *\<^sub>C (f x)) = (cmod r) * onorm f\ proof- have \(norm (f x)) / norm x \ onorm f\ for x using a1 by (simp add: bounded_clinear.bounded_linear le_onorm) hence t2: \bdd_above {(norm (f x)) / norm x | x. True}\ by fastforce have \continuous_on UNIV ( (*) w ) \ for w::real by simp hence \isCont ( ((*) (cmod r)) ) x\ for x by simp hence t3: \continuous (at_left (Sup {(norm (f x)) / norm x | x. True})) ((*) (cmod r))\ using Elementary_Topology.continuous_at_imp_continuous_within by blast have \{(norm (f x)) / norm x | x. True} \ {}\ by blast moreover have \mono ((*) (cmod r))\ by (simp add: monoI ordered_comm_semiring_class.comm_mult_left_mono) ultimately have \Sup {((*) (cmod r)) ((norm (f x)) / norm x) | x. True} = ((*) (cmod r)) (Sup {(norm (f x)) / norm x | x. True})\ using t2 t3 by (simp add: continuous_at_Sup_mono full_SetCompr_eq image_image) hence \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (cmod r) * (Sup {(norm (f x)) / norm x | x. True})\ by blast moreover have \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (SUP x. cmod r * norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) moreover have \(Sup {(norm (f x)) / norm x | x. True}) = (SUP x. norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) ultimately have t1: "(SUP x. cmod r * norm (f x) / norm x) = cmod r * (SUP x. norm (f x) / norm x)" by simp have \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. norm ( (\ t. r *\<^sub>C (f t)) x) / norm x)\ by (simp add: onorm_def) hence \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. (cmod r) * (norm (f x)) / norm x)\ by simp also have \... = (cmod r) * (SUP x. (norm (f x)) / norm x)\ using t1. finally show ?thesis by (simp add: onorm_def) qed lemma onorm_scaleC_left_lemma: fixes f :: "'a::complex_normed_vector" assumes r: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" proof (rule onorm_bound) fix x have "norm (r x *\<^sub>C f) = norm (r x) * norm f" by simp also have "\ \ onorm r * norm x * norm f" by (simp add: bounded_clinear.bounded_linear mult.commute mult_left_mono onorm r) finally show "norm (r x *\<^sub>C f) \ onorm r * norm f * norm x" by (simp add: ac_simps) show "0 \ onorm r * norm f" by (simp add: bounded_clinear.bounded_linear onorm_pos_le r) qed lemma onorm_scaleC_left: fixes f :: "'a::complex_normed_vector" assumes f: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) = onorm r * norm f" proof (cases "f = 0") assume "f \ 0" show ?thesis proof (rule order_antisym) show "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" using f by (rule onorm_scaleC_left_lemma) next have bl1: "bounded_clinear (\x. r x *\<^sub>C f)" by (metis bounded_clinear_scaleC_const f) have x1:"bounded_clinear (\x. r x * norm f)" by (metis bounded_clinear_mult_const f) have "onorm r \ onorm (\x. r x * complex_of_real (norm f)) / norm f" if "onorm r \ onorm (\x. r x * complex_of_real (norm f)) * cmod (1 / complex_of_real (norm f))" and "f \ 0" using that by (metis complex_of_real_cmod complex_of_real_nn_iff field_class.field_divide_inverse inverse_eq_divide nice_ordered_field_class.zero_le_divide_1_iff norm_ge_zero of_real_1 of_real_divide of_real_eq_iff) hence "onorm r \ onorm (\x. r x * norm f) * inverse (norm f)" using \f \ 0\ onorm_scaleC_left_lemma[OF x1, of "inverse (norm f)"] by (simp add: inverse_eq_divide) also have "onorm (\x. r x * norm f) \ onorm (\x. r x *\<^sub>C f)" proof (rule onorm_bound) have "bounded_linear (\x. r x *\<^sub>C f)" using bl1 bounded_clinear.bounded_linear by auto thus "0 \ onorm (\x. r x *\<^sub>C f)" by (rule Operator_Norm.onorm_pos_le) show "cmod (r x * complex_of_real (norm f)) \ onorm (\x. r x *\<^sub>C f) * norm x" for x :: 'b by (smt \bounded_linear (\x. r x *\<^sub>C f)\ complex_of_real_cmod complex_of_real_nn_iff complex_scaleC_def norm_ge_zero norm_scaleC of_real_eq_iff onorm) qed finally show "onorm r * norm f \ onorm (\x. r x *\<^sub>C f)" using \f \ 0\ by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) qed qed (simp add: onorm_zero) subsection \Finite dimension and canonical basis\ lemma vector_finitely_spanned: assumes \z \ cspan T\ shows \\ S. finite S \ S \ T \ z \ cspan S\ proof- have \\ S r. finite S \ S \ T \ z = (\a\S. r a *\<^sub>C a)\ using complex_vector.span_explicit[where b = "T"] assms by auto then obtain S r where \finite S\ and \S \ T\ and \z = (\a\S. r a *\<^sub>C a)\ by blast thus ?thesis by (meson complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset subset_iff) qed setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a set \ 'a set\)\ class cfinite_dim = complex_vector + assumes cfinitely_spanned: "\S::'a set. finite S \ cspan S = UNIV" class basis_enum = complex_vector + fixes canonical_basis :: "'a list" assumes distinct_canonical_basis[simp]: "distinct canonical_basis" and is_cindependent_set[simp]: "cindependent (set canonical_basis)" and is_generator_set[simp]: "cspan (set canonical_basis) = UNIV" setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a::complex_vector set \ 'a set\)\ instantiation complex :: basis_enum begin definition "canonical_basis = [1::complex]" instance proof show "distinct (canonical_basis::complex list)" by (simp add: canonical_basis_complex_def) show "cindependent (set (canonical_basis::complex list))" unfolding canonical_basis_complex_def by auto show "cspan (set (canonical_basis::complex list)) = UNIV" unfolding canonical_basis_complex_def apply (auto simp add: cspan_raw_def vector_space_over_itself.span_Basis) by (metis complex_scaleC_def complex_vector.span_base complex_vector.span_scale cspan_raw_def insertI1 mult.right_neutral) qed end lemma cdim_UNIV_basis_enum[simp]: \cdim (UNIV::'a::basis_enum set) = length (canonical_basis::'a list)\ apply (subst is_generator_set[symmetric]) apply (subst complex_vector.dim_span_eq_card_independent) apply (rule is_cindependent_set) using distinct_canonical_basis distinct_card by blast lemma finite_basis: "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = UNIV" proof - from cfinitely_spanned obtain S :: \'a set\ where \finite S\ and \cspan S = UNIV\ by auto from complex_vector.maximal_independent_subset obtain B :: \'a set\ where \B \ S\ and \cindependent B\ and \S \ cspan B\ by metis moreover have \finite B\ using \B \ S\ \finite S\ by (meson finite_subset) moreover have \cspan B = UNIV\ using \cspan S = UNIV\ \S \ cspan B\ by (metis complex_vector.span_eq top_greatest) ultimately show ?thesis by auto qed instance basis_enum \ cfinite_dim apply intro_classes apply (rule exI[of _ \set canonical_basis\]) using is_cindependent_set is_generator_set by auto lemma cindependent_cfinite_dim_finite: assumes \cindependent (S::'a::cfinite_dim set)\ shows \finite S\ by (metis assms cfinitely_spanned complex_vector.independent_span_bound top_greatest) lemma cfinite_dim_finite_subspace_basis: assumes \csubspace X\ shows "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = X" by (meson assms cindependent_cfinite_dim_finite complex_vector.basis_exists complex_vector.span_subspace) text \The following auxiliary lemma (\finite_span_complete_aux\) shows more or less the same as \finite_span_representation_bounded\, \finite_span_complete\ below (see there for an intuition about the mathematical content of the lemmas). However, there is one difference: Here we additionally assume here that there is a bijection rep/abs between a finite type \<^typ>\'basis\ and the set $B$. This is needed to be able to use results about euclidean spaces that are formulated w.r.t. the type class \<^class>\finite\ Since we anyway assume that $B$ is finite, this added assumption does not make the lemma weaker. However, we cannot derive the existence of \<^typ>\'basis\ inside the proof (HOL does not support such reasoning). Therefore we have the type \<^typ>\'basis\ as an explicit assumption and remove it using @{attribute internalize_sort} after the proof.\ (* TODO: Maybe this should be in Extra_Vector_Spaces *) lemma finite_span_complete_aux: fixes b :: "'b::real_normed_vector" and B :: "'b set" and rep :: "'basis::finite \ 'b" and abs :: "'b \ 'basis" assumes t: "type_definition rep abs B" and t1: "finite B" and t2: "b\B" and t3: "independent B" shows "\D>0. \\. norm (representation B \ b) \ norm \ * D" and "complete (span B)" proof - define repr where "repr = real_vector.representation B" define repr' where "repr' \ = Abs_euclidean_space (repr \ o rep)" for \ define comb where "comb l = (\b\B. l b *\<^sub>R b)" for l define comb' where "comb' l = comb (Rep_euclidean_space l o abs)" for l have comb_cong: "comb x = comb y" if "\z. z\B \ x z = y z" for x y unfolding comb_def using that by auto have comb_repr[simp]: "comb (repr \) = \" if "\ \ real_vector.span B" for \ using \comb \ \l. \b\B. l b *\<^sub>R b\ local.repr_def real_vector.sum_representation_eq t1 t3 that by fastforce have w5:"(\b | (b \ B \ x b \ 0) \ b \ B. x b *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" for x using \finite B\ by (smt DiffD1 DiffD2 mem_Collect_eq real_vector.scale_eq_0_iff subset_eq sum.mono_neutral_left) have "representation B (\b\B. x b *\<^sub>R b) = (\b. if b \ B then x b else 0)" for x proof (rule real_vector.representation_eqI) show "independent B" by (simp add: t3) show "(\b\B. x b *\<^sub>R b) \ span B" by (meson real_vector.span_scale real_vector.span_sum real_vector.span_superset subset_iff) show "b \ B" if "(if b \ B then x b else 0) \ 0" for b :: 'b using that by meson show "finite {b. (if b \ B then x b else 0) \ 0}" using t1 by auto show "(\b | (if b \ B then x b else 0) \ 0. (if b \ B then x b else 0) *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" using w5 by simp qed hence repr_comb[simp]: "repr (comb x) = (\b. if b\B then x b else 0)" for x unfolding repr_def comb_def. have repr_bad[simp]: "repr \ = (\_. 0)" if "\ \ real_vector.span B" for \ unfolding repr_def using that by (simp add: real_vector.representation_def) have [simp]: "repr' \ = 0" if "\ \ real_vector.span B" for \ unfolding repr'_def repr_bad[OF that] apply transfer by auto have comb'_repr'[simp]: "comb' (repr' \) = \" if "\ \ real_vector.span B" for \ proof - have x1: "(repr \ \ rep \ abs) z = repr \ z" if "z \ B" for z unfolding o_def using t that type_definition.Abs_inverse by fastforce have "comb' (repr' \) = comb ((repr \ \ rep) \ abs)" unfolding comb'_def repr'_def by (subst Abs_euclidean_space_inverse; simp) also have "\ = comb (repr \)" using x1 comb_cong by blast also have "\ = \" using that by simp finally show ?thesis by - qed have t1: "Abs_euclidean_space (Rep_euclidean_space t) = t" if "\x. rep x \ B" for t::"'a euclidean_space" apply (subst Rep_euclidean_space_inverse) by simp have "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x y else 0) = x" for x using type_definition.Rep[OF t] apply simp using t1 by blast hence "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x (abs (rep y)) else 0) = x" for x apply (subst type_definition.Rep_inverse[OF t]) by simp hence repr'_comb'[simp]: "repr' (comb' x) = x" for x unfolding comb'_def repr'_def o_def by simp have sphere: "compact (sphere 0 d :: 'basis euclidean_space set)" for d using compact_sphere by blast have "complete (UNIV :: 'basis euclidean_space set)" by (simp add: complete_UNIV) have "(\b\B. (Rep_euclidean_space (x + y) \ abs) b *\<^sub>R b) = (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b) + (\b\B. (Rep_euclidean_space y \ abs) b *\<^sub>R b)" for x :: "'basis euclidean_space" and y :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: scaleR_add_left sum.distrib) moreover have "(\b\B. (Rep_euclidean_space (c *\<^sub>R x) \ abs) b *\<^sub>R b) = c *\<^sub>R (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b)" for c :: real and x :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: real_vector.scale_sum_right) ultimately have blin_comb': "bounded_linear comb'" unfolding comb_def comb'_def by (rule bounded_linearI') hence "continuous_on X comb'" for X by (simp add: linear_continuous_on) hence "compact (comb' ` sphere 0 d)" for d using sphere by (rule compact_continuous_image) hence compact_norm_comb': "compact (norm ` comb' ` sphere 0 1)" using compact_continuous_image continuous_on_norm_id by blast have not0: "0 \ norm ` comb' ` sphere 0 1" proof (rule ccontr, simp) assume "0 \ norm ` comb' ` sphere 0 1" then obtain x where nc0: "norm (comb' x) = 0" and x: "x \ sphere 0 1" by auto hence "comb' x = 0" by simp hence "repr' (comb' x) = 0" unfolding repr'_def o_def repr_def apply simp by (smt repr'_comb' blin_comb' dist_0_norm linear_simps(3) mem_sphere norm_zero x) hence "x = 0" by auto with x show False by simp qed have "closed (norm ` comb' ` sphere 0 1)" using compact_imp_closed compact_norm_comb' by blast moreover have "0 \ norm ` comb' ` sphere 0 1" by (simp add: not0) ultimately have "\d>0. \x\norm ` comb' ` sphere 0 1. d \ dist 0 x" by (meson separate_point_closed) then obtain d where d: "x\norm ` comb' ` sphere 0 1 \ d \ dist 0 x" and "d > 0" for x by metis define D where "D = 1/d" hence "D > 0" using \d>0\ unfolding D_def by auto have "x \ d" if "x\norm ` comb' ` sphere 0 1" for x using d that apply auto by fastforce hence *: "norm (comb' x) \ d" if "norm x = 1" for x using that by auto have norm_comb': "norm (comb' x) \ d * norm x" for x proof (cases "x=0") show "d * norm x \ norm (comb' x)" if "x = 0" using that by simp show "d * norm x \ norm (comb' x)" if "x \ 0" using that using *[of "(1/norm x) *\<^sub>R x"] unfolding linear_simps(5)[OF blin_comb'] apply auto by (simp add: le_divide_eq) qed have *: "norm (repr' \) \ norm \ * D" for \ proof (cases "\ \ real_vector.span B") show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that unfolding D_def using norm_comb'[of "repr' \"] \d>0\ by (simp_all add: linordered_field_class.mult_imp_le_div_pos mult.commute) show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that \0 < D\ by auto qed hence "norm (Rep_euclidean_space (repr' \) (abs b)) \ norm \ * D" for \ proof - have "(Rep_euclidean_space (repr' \) (abs b)) = repr' \ \ euclidean_space_basis_vector (abs b)" apply (transfer fixing: abs b) by auto also have "\\\ \ norm (repr' \)" apply (rule Basis_le_norm) unfolding Basis_euclidean_space_def by simp also have "\ \ norm \ * D" using * by auto finally show ?thesis by simp qed hence "norm (repr \ b) \ norm \ * D" for \ unfolding repr'_def by (smt \comb' \ \l. comb (Rep_euclidean_space l \ abs)\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' comp_apply norm_le_zero_iff repr_bad repr_comb) thus "\D>0. \\. norm (repr \ b) \ norm \ * D" using \D>0\ by auto from \d>0\ have complete_comb': "complete (comb' ` UNIV)" proof (rule complete_isometric_image) show "subspace (UNIV::'basis euclidean_space set)" by simp show "bounded_linear comb'" by (simp add: blin_comb') show "\x\UNIV. d * norm x \ norm (comb' x)" by (simp add: norm_comb') show "complete (UNIV::'basis euclidean_space set)" by (simp add: \complete UNIV\) qed have range_comb': "comb' ` UNIV = real_vector.span B" proof (auto simp: image_def) show "comb' x \ real_vector.span B" for x by (metis comb'_def comb_cong comb_repr local.repr_def repr_bad repr_comb real_vector.representation_zero real_vector.span_zero) next fix \ assume "\ \ real_vector.span B" then obtain f where f: "comb f = \" apply atomize_elim unfolding span_finite[OF \finite B\] comb_def by auto define f' where "f' b = (if b\B then f b else 0)" for b :: 'b have f': "comb f' = \" unfolding f[symmetric] apply (rule comb_cong) unfolding f'_def by simp define x :: "'basis euclidean_space" where "x = Abs_euclidean_space (f' o rep)" have "\ = comb' x" by (metis (no_types, lifting) \\ \ span B\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' f' fun.map_cong repr_comb t type_definition.Rep_range x_def) thus "\x. \ = comb' x" by auto qed from range_comb' complete_comb' show "complete (real_vector.span B)" by simp qed (* TODO: Maybe this should be in Extra_Vector_Spaces *) lemma finite_span_complete[simp]: fixes A :: "'a::real_normed_vector set" assumes "finite A" shows "complete (span A)" text \The span of a finite set is complete.\ proof (cases "A \ {} \ A \ {0}") case True obtain B where BT: "real_vector.span B = real_vector.span A" and "independent B" and "finite B" by (meson True assms finite_subset real_vector.maximal_independent_subset real_vector.span_eq real_vector.span_superset subset_trans) have "B\{}" apply (rule ccontr, simp) using BT True by (metis real_vector.span_superset real_vector.span_empty subset_singletonD) (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) { (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux, otherwise "internalize_sort" below fails *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) note finite_span_complete_aux(2)[internalize_sort "'basis::finite"] note this[OF basisT_finite t] } note this[cancel_type_definition, OF \B\{}\ \finite B\ _ \independent B\] hence "complete (real_vector.span B)" using \B\{}\ by auto thus "complete (real_vector.span A)" unfolding BT by simp next case False thus ?thesis using complete_singleton by auto qed (* TODO: Maybe this should be in Extra_Vector_Spaces *) lemma finite_span_representation_bounded: fixes B :: "'a::real_normed_vector set" assumes "finite B" and "independent B" shows "\D>0. \\ b. abs (representation B \ b) \ norm \ * D" text \ Assume $B$ is a finite linear independent set of vectors (in a real normed vector space). Let $\alpha^\psi_b$ be the coefficients of $\psi$ expressed as a linear combination over $B$. Then $\alpha$ is is uniformly cblinfun (i.e., $\lvert\alpha^\psi_b \leq D \lVert\psi\rVert\psi$ for some $D$ independent of $\psi,b$). (This also holds when $b$ is not in the span of $B$ because of the way \real_vector.representation\ is defined in this corner case.)\ proof (cases "B\{}") case True (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) define repr where "repr = real_vector.representation B" { (* Step 1: Create a fake type definition by introducing a new type variable 'basis and then assuming the existence of the morphisms Rep/Abs to B This is then roughly equivalent to "typedef 'basis = B" *) (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux (I.e., we cannot call it 'basis) *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto (* Step 2: We show that our fake typedef 'basisT could be instantiated as type class finite *) have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) (* Step 3: We take the finite_span_complete_aux and remove the requirement that 'basis::finite (instead, a precondition "class.finite TYPE('basisT)" is introduced) *) note finite_span_complete_aux(1)[internalize_sort "'basis::finite"] (* Step 4: We instantiate the premises *) note this[OF basisT_finite t] } (* Now we have the desired fact, except that it still assumes that B is isomorphic to some type 'basis together with the assumption that there are morphisms between 'basis and B. 'basis and that premise are removed using cancel_type_definition *) note this[cancel_type_definition, OF True \finite B\ _ \independent B\] hence d2:"\D. \\. D>0 \ norm (repr \ b) \ norm \ * D" if \b\B\ for b by (simp add: repr_def that True) have d1: " (\b. b \ B \ \D. \\. 0 < D \ norm (repr \ b) \ norm \ * D) \ \D. \b \. b \ B \ 0 < D b \ norm (repr \ b) \ norm \ * D b" apply (rule choice) by auto then obtain D where D: "D b > 0 \ norm (repr \ b) \ norm \ * D b" if "b\B" for b \ apply atomize_elim using d2 by blast hence Dpos: "D b > 0" and Dbound: "norm (repr \ b) \ norm \ * D b" if "b\B" for b \ using that by auto define Dall where "Dall = Max (D`B)" have "Dall > 0" unfolding Dall_def using \finite B\ \B\{}\ Dpos by (metis (mono_tags, lifting) Max_in finite_imageI image_iff image_is_empty) have "Dall \ D b" if "b\B" for b unfolding Dall_def using \finite B\ that by auto with Dbound have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ using that by (smt mult_left_mono norm_not_less_zero) moreover have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ unfolding repr_def using real_vector.representation_ne_zero True by (metis calculation empty_subsetI less_le_trans local.repr_def norm_ge_zero norm_zero not_less subsetI subset_antisym) ultimately show "\D>0. \\ b. abs (repr \ b) \ norm \ * D" using \Dall > 0\ real_norm_def by metis next case False thus ?thesis unfolding repr_def using real_vector.representation_ne_zero[of B] using nice_ordered_field_class.linordered_field_no_ub by fastforce qed hide_fact finite_span_complete_aux lemma finite_cspan_complete[simp]: fixes B :: "'a::complex_normed_vector set" assumes "finite B" shows "complete (cspan B)" by (simp add: assms cspan_as_span) (* TODO: Maybe this should be in Extra_Vector_Spaces *) lemma finite_span_closed[simp]: fixes B :: "'a::real_normed_vector set" assumes "finite B" shows "closed (real_vector.span B)" by (simp add: assms complete_imp_closed) lemma finite_cspan_closed[simp]: fixes S::\'a::complex_normed_vector set\ assumes a1: \finite S\ shows \closed (cspan S)\ by (simp add: assms complete_imp_closed) lemma closure_finite_cspan: fixes T::\'a::complex_normed_vector set\ assumes \finite T\ shows \closure (cspan T) = cspan T\ by (simp add: assms) lemma finite_cspan_crepresentation_bounded: fixes B :: "'a::complex_normed_vector set" assumes a1: "finite B" and a2: "cindependent B" shows "\D>0. \\ b. cmod (crepresentation B \ b) \ norm \ * D" proof - define B' where "B' = (B \ scaleC \ ` B)" have independent_B': "independent B'" using B'_def \cindependent B\ by (simp add: real_independent_from_complex_independent a1) have "finite B'" unfolding B'_def using \finite B\ by simp obtain D' where "D' > 0" and D': "norm (real_vector.representation B' \ b) \ norm \ * D'" for \ b apply atomize_elim using independent_B' \finite B'\ by (simp add: finite_span_representation_bounded) define D where "D = 2*D'" from \D' > 0\ have \D > 0\ unfolding D_def by simp have "norm (crepresentation B \ b) \ norm \ * D" for \ b proof (cases "b\B") case True have d3: "norm \ = 1" by simp have "norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm \ * norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_scaleC by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d3 by simp finally have d2:"norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))". have "norm (crepresentation B \ b) = norm (complex_of_real (real_vector.representation B' \ b) + \ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" by (simp add: B'_def True a1 a2 crepresentation_from_representation) also have "\ \ norm (complex_of_real (real_vector.representation B' \ b)) + norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_triangle_ineq by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ b)) + norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d2 by simp also have "\ = norm (real_vector.representation B' \ b) + norm (real_vector.representation B' \ (\ *\<^sub>C b))" by simp also have "\ \ norm \ * D' + norm \ * D'" by (rule add_mono; rule D') also have "\ \ norm \ * D" unfolding D_def by linarith finally show ?thesis by auto next case False hence "crepresentation B \ b = 0" using complex_vector.representation_ne_zero by blast thus ?thesis by (smt \0 < D\ norm_ge_zero norm_zero split_mult_pos_le) qed with \D > 0\ show ?thesis by auto qed lemma bounded_clinear_finite_dim[simp]: fixes f :: \'a::{cfinite_dim,complex_normed_vector} \ 'b::complex_normed_vector\ assumes \clinear f\ shows \bounded_clinear f\ proof - include notation_norm obtain basis :: \'a set\ where b1: "complex_vector.span basis = UNIV" and b2: "cindependent basis" and b3:"finite basis" using finite_basis by auto have "\C>0. \\ b. cmod (crepresentation basis \ b) \ \\\ * C" using finite_cspan_crepresentation_bounded[where B = basis] b2 b3 by blast then obtain C where s1: "cmod (crepresentation basis \ b) \ \\\ * C" and s2: "C > 0" for \ b by blast define M where "M = C * (\a\basis. \f a\)" have "\f x\ \ \x\ * M" for x proof- define r where "r b = crepresentation basis x b" for b have x_span: "x \ complex_vector.span basis" by (simp add: b1) have f0: "v \ basis" if "r v \ 0" for v using complex_vector.representation_ne_zero r_def that by auto have w:"{a|a. r a \ 0} \ basis" using f0 by blast hence f1: "finite {a|a. r a \ 0}" using b3 rev_finite_subset by auto have f2: "(\a| r a \ 0. r a *\<^sub>C a) = x" unfolding r_def using b2 complex_vector.sum_nonzero_representation_eq x_span Collect_cong by fastforce have g1: "(\a\basis. crepresentation basis x a *\<^sub>C a) = x" by (simp add: b2 b3 complex_vector.sum_representation_eq x_span) have f3: "(\a\basis. r a *\<^sub>C a) = x" unfolding r_def by (simp add: g1) hence "f x = f (\a\basis. r a *\<^sub>C a)" by simp also have "\ = (\a\basis. r a *\<^sub>C f a)" by (smt (verit, ccfv_SIG) assms complex_vector.linear_scale complex_vector.linear_sum sum.cong) finally have "f x = (\a\basis. r a *\<^sub>C f a)". hence "\f x\ = \(\a\basis. r a *\<^sub>C f a)\" by simp also have "\ \ (\a\basis. \r a *\<^sub>C f a\)" by (simp add: sum_norm_le) also have "\ \ (\a\basis. \r a\ * \f a\)" by simp also have "\ \ (\a\basis. \x\ * C * \f a\)" using sum_mono s1 unfolding r_def by (simp add: sum_mono mult_right_mono) also have "\ \ \x\ * C * (\a\basis. \f a\)" using sum_distrib_left by (smt sum.cong) also have "\ = \x\ * M" unfolding M_def by linarith finally show ?thesis . qed thus ?thesis using assms bounded_clinear_def bounded_clinear_axioms_def by blast qed (* TODO: Maybe this should be in Extra_Vector_Spaces if we move finite_span_closed, too. *) lemma summable_on_scaleR_left_converse: \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here.\ fixes f :: \'b \ real\ and c :: \'a :: real_normed_vector\ assumes \c \ 0\ assumes \(\x. f x *\<^sub>R c) summable_on A\ shows \f summable_on A\ proof - define fromR toR T where \fromR x = x *\<^sub>R c\ and \toR = inv fromR\ and \T = range fromR\ for x :: real have \additive fromR\ by (simp add: fromR_def additive.intro scaleR_left_distrib) have \inj fromR\ by (simp add: fromR_def \c \ 0\ inj_def) have toR_fromR: \toR (fromR x) = x\ for x by (simp add: \inj fromR\ toR_def) have fromR_toR: \fromR (toR x) = x\ if \x \ T\ for x by (metis T_def f_inv_into_f that toR_def) have 1: \sum (toR \ (fromR \ f)) F = toR (sum (fromR \ f) F)\ if \finite F\ for F by (simp add: o_def additive.sum[OF \additive fromR\, symmetric] toR_fromR) have 2: \sum (fromR \ f) F \ T\ if \finite F\ for F by (simp add: o_def additive.sum[OF \additive fromR\, symmetric] T_def) have 3: \(toR \ toR x) (at x within T)\ for x proof (cases \x \ T\) case True have \dist (toR y) (toR x) < e\ if \y\T\ \e>0\ \dist y x < e * norm c\ for e y proof - obtain x' y' where x: \x = fromR x'\ and y: \y = fromR y'\ using T_def True \y \ T\ by blast have \dist (toR y) (toR x) = dist (fromR (toR y)) (fromR (toR x)) / norm c\ by (auto simp: dist_real_def fromR_def \c \ 0\) also have \\ = dist y x / norm c\ using \x\T\ \y\T\ by (simp add: fromR_toR) also have \\ < e\ using \dist y x < e * norm c\ by (simp add: divide_less_eq that(2)) finally show ?thesis by (simp add: x y toR_fromR) qed then show ?thesis apply (auto simp: tendsto_iff at_within_def eventually_inf_principal eventually_nhds_metric) by (metis assms(1) div_0 divide_less_eq zero_less_norm_iff) next case False have \T = span {c}\ by (simp add: T_def fromR_def span_singleton) then have \closed T\ by simp with False have \x \ closure T\ by simp then have \(at x within T) = bot\ by (rule not_in_closure_trivial_limitI) then show ?thesis by simp qed have 4: \(fromR \ f) summable_on A\ by (simp add: assms(2) fromR_def summable_on_cong) have \(toR o (fromR o f)) summable_on A\ using 1 2 3 4 by (rule summable_on_comm_additive_general[where T=T]) with toR_fromR show ?thesis by (auto simp: o_def) qed (* TODO: Maybe this should be in Extra_Vector_Spaces if we move finite_span_closed, too. *) lemma infsum_scaleR_left: \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here. It is a strengthening of @{thm [source] infsum_scaleR_left}.\ fixes c :: \'a :: real_normed_vector\ shows "infsum (\x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c" proof (cases \f summable_on A\) case True then show ?thesis apply (subst asm_rl[of \(\x. f x *\<^sub>R c) = (\y. y *\<^sub>R c) o f\], simp add: o_def) apply (rule infsum_comm_additive) using True by (auto simp add: scaleR_left.additive_axioms) next case False then have \\ (\x. f x *\<^sub>R c) summable_on A\ if \c \ 0\ using summable_on_scaleR_left_converse[where A=A and f=f and c=c] using that by auto with False show ?thesis apply (cases \c = 0\) by (auto simp add: infsum_not_exists) qed (* TODO: Maybe this should be in Extra_Vector_Spaces if we move finite_span_closed, too. *) lemma infsum_of_real: shows \(\\<^sub>\x\A. of_real (f x) :: 'b::{real_normed_vector, real_algebra_1}) = of_real (\\<^sub>\x\A. f x)\ \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here.\ unfolding of_real_def by (rule infsum_scaleR_left) subsection \Closed subspaces\ lemma csubspace_INF[simp]: "(\x. x \ A \ csubspace x) \ csubspace (\A)" by (simp add: complex_vector.subspace_Inter) locale closed_csubspace = fixes A::"('a::{complex_vector,topological_space}) set" assumes subspace: "csubspace A" assumes closed: "closed A" declare closed_csubspace.subspace[simp] lemma closure_is_csubspace[simp]: fixes A::"('a::complex_normed_vector) set" assumes \csubspace A\ shows \csubspace (closure A)\ proof- have "x \ closure A \ y \ closure A \ x+y \ closure A" for x y proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast assume \y\(closure A)\ then obtain yy where \\ n::nat. yy n \ A\ and \yy \ y\ using closure_sequential by blast have \\ n::nat. (xx n) + (yy n) \ A\ using \\n. xx n \ A\ \\n. yy n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) hence \(\ n. (xx n) + (yy n)) \ x + y\ using \xx \ x\ \yy \ y\ by (simp add: tendsto_add) thus ?thesis using \\ n::nat. (xx n) + (yy n) \ A\ by (meson closure_sequential) qed moreover have "x\(closure A) \ c *\<^sub>C x \ (closure A)" for x c proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast have \\ n::nat. c *\<^sub>C (xx n) \ A\ using \\n. xx n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) have \isCont (\ t. c *\<^sub>C t) x\ using bounded_clinear.bounded_linear bounded_clinear_scaleC_right linear_continuous_at by auto hence \(\ n. c *\<^sub>C (xx n)) \ c *\<^sub>C x\ using \xx \ x\ by (simp add: isCont_tendsto_compose) thus ?thesis using \\ n::nat. c *\<^sub>C (xx n) \ A\ by (meson closure_sequential) qed moreover have "0 \ (closure A)" using assms closure_subset complex_vector.subspace_def by (metis in_mono) ultimately show ?thesis by (simp add: complex_vector.subspaceI) qed lemma csubspace_set_plus: assumes \csubspace A\ and \csubspace B\ shows \csubspace (A + B)\ proof - define C where \C = {\+\| \ \. \\A \ \\B}\ have "x\C \ y\C \ x+y\C" for x y using C_def assms(1) assms(2) complex_vector.subspace_add complex_vector.subspace_sums by blast moreover have "c *\<^sub>C x \ C" if \x\C\ for x c proof - have "csubspace C" by (simp add: C_def assms(1) assms(2) complex_vector.subspace_sums) then show ?thesis using that by (simp add: complex_vector.subspace_def) qed moreover have "0 \ C" using \C = {\ + \ |\ \. \ \ A \ \ \ B}\ add.inverse_neutral add_uminus_conv_diff assms(1) assms(2) diff_0 mem_Collect_eq add.right_inverse by (metis (mono_tags, lifting) complex_vector.subspace_0) ultimately show ?thesis unfolding C_def complex_vector.subspace_def by (smt mem_Collect_eq set_plus_elim set_plus_intro) qed lemma closed_csubspace_0[simp]: "closed_csubspace ({0} :: ('a::{complex_vector,t1_space}) set)" proof- have \csubspace {0}\ using add.right_neutral complex_vector.subspace_def scaleC_right.zero by blast moreover have "closed ({0} :: 'a set)" by simp ultimately show ?thesis by (simp add: closed_csubspace_def) qed lemma closed_csubspace_UNIV[simp]: "closed_csubspace (UNIV::('a::{complex_vector,topological_space}) set)" proof- have \csubspace UNIV\ by simp moreover have \closed UNIV\ by simp ultimately show ?thesis unfolding closed_csubspace_def by auto qed lemma closed_csubspace_inter[simp]: assumes "closed_csubspace A" and "closed_csubspace B" shows "closed_csubspace (A\B)" proof- obtain C where \C = A \ B\ by blast have \csubspace C\ proof- have "x\C \ y\C \ x+y\C" for x y by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "x\C \ c *\<^sub>C x \ C" for x c by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "0 \ C" using \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def by fastforce ultimately show ?thesis by (simp add: complex_vector.subspace_def) qed moreover have \closed C\ using \C = A \ B\ by (simp add: assms(1) assms(2) closed_Int closed_csubspace.closed) ultimately show ?thesis using \C = A \ B\ by (simp add: closed_csubspace_def) qed lemma closed_csubspace_INF[simp]: assumes a1: "\A\\. closed_csubspace A" shows "closed_csubspace (\\)" proof- have \csubspace (\\)\ by (simp add: assms closed_csubspace.subspace complex_vector.subspace_Inter) moreover have \closed (\\)\ by (simp add: assms closed_Inter closed_csubspace.closed) ultimately show ?thesis by (simp add: closed_csubspace.intro) qed typedef (overloaded) ('a::"{complex_vector,topological_space}") ccsubspace = \{S::'a set. closed_csubspace S}\ morphisms space_as_set Abs_clinear_space using Complex_Vector_Spaces.closed_csubspace_UNIV by blast setup_lifting type_definition_ccsubspace lemma csubspace_space_as_set[simp]: \csubspace (space_as_set S)\ by (metis closed_csubspace_def mem_Collect_eq space_as_set) lemma closed_space_as_set[simp]: \closed (space_as_set S)\ apply transfer by (simp add: closed_csubspace.closed) lemma zero_space_as_set[simp]: \0 \ space_as_set A\ by (simp add: complex_vector.subspace_0) instantiation ccsubspace :: (complex_normed_vector) scaleC begin lift_definition scaleC_ccsubspace :: "complex \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>C) c ` S" proof show "csubspace ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: complex_vector.linear_subspace_image) show "closed ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: closed_scaleC closed_csubspace.closed) qed lift_definition scaleR_ccsubspace :: "real \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>R) c ` S" proof show "csubspace ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that using bounded_clinear_def bounded_clinear_scaleC_right scaleR_scaleC by (simp add: scaleR_scaleC complex_vector.linear_subspace_image) show "closed ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that by (simp add: closed_scaling closed_csubspace.closed) qed instance proof show "((*\<^sub>R) r::'a ccsubspace \ _) = (*\<^sub>C) (complex_of_real r)" for r :: real by (simp add: scaleR_scaleC scaleC_ccsubspace_def scaleR_ccsubspace_def) qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") bot begin lift_definition bot_ccsubspace :: \'a ccsubspace\ is \{0}\ by simp instance.. end lemma zero_cblinfun_image[simp]: "0 *\<^sub>C S = bot" for S :: "_ ccsubspace" proof transfer have "(0::'b) \ (\x. 0) ` S" if "closed_csubspace S" for S::"'b set" using that unfolding closed_csubspace_def by (simp add: complex_vector.linear_subspace_image complex_vector.module_hom_zero complex_vector.subspace_0) thus "(*\<^sub>C) 0 ` S = {0::'b}" if "closed_csubspace (S::'b set)" for S :: "'b set" using that by (auto intro !: exI [of _ 0]) qed lemma csubspace_scaleC_invariant: fixes a S assumes \a \ 0\ and \csubspace S\ shows \(*\<^sub>C) a ` S = S\ proof- have \x \ (*\<^sub>C) a ` S \ x \ S\ for x using assms(2) complex_vector.subspace_scale by blast moreover have \x \ S \ x \ (*\<^sub>C) a ` S\ for x proof - assume "x \ S" hence "\c aa. (c / a) *\<^sub>C aa \ S \ c *\<^sub>C aa = x" using assms(2) complex_vector.subspace_def scaleC_one by metis hence "\aa. aa \ S \ a *\<^sub>C aa = x" using assms(1) by auto thus ?thesis by (meson image_iff) qed ultimately show ?thesis by blast qed lemma ccsubspace_scaleC_invariant[simp]: "a \ 0 \ a *\<^sub>C S = S" for S :: "_ ccsubspace" apply transfer by (simp add: closed_csubspace.subspace csubspace_scaleC_invariant) instantiation ccsubspace :: ("{complex_vector,topological_space}") "top" begin lift_definition top_ccsubspace :: \'a ccsubspace\ is \UNIV\ by simp instance .. end lemma space_as_set_bot[simp]: \space_as_set bot = {0}\ by (rule bot_ccsubspace.rep_eq) lemma ccsubspace_top_not_bot[simp]: "(top::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ bot" (* The type class t1_space is needed because the definition of bot in ccsubspace needs it *) by (metis UNIV_not_singleton bot_ccsubspace.rep_eq top_ccsubspace.rep_eq) lemma ccsubspace_bot_not_top[simp]: "(bot::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ top" using ccsubspace_top_not_bot by metis instantiation ccsubspace :: ("{complex_vector,topological_space}") "Inf" begin lift_definition Inf_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\ S. \ S\ proof fix S :: "'a set set" assume closed: "closed_csubspace x" if \x \ S\ for x show "csubspace (\ S::'a set)" by (simp add: closed closed_csubspace.subspace) show "closed (\ S::'a set)" by (simp add: closed closed_csubspace.closed) qed instance .. end lift_definition ccspan :: "'a::complex_normed_vector set \ 'a ccsubspace" is "\G. closure (cspan G)" proof (rule closed_csubspace.intro) fix S :: "'a set" show "csubspace (closure (cspan S))" by (simp add: closure_is_csubspace) show "closed (closure (cspan S))" by simp qed lemma ccspan_superset: \A \ space_as_set (ccspan A)\ for A :: \'a::complex_normed_vector set\ apply transfer by (meson closure_subset complex_vector.span_superset subset_trans) lemma ccspan_canonical_basis[simp]: "ccspan (set canonical_basis) = top" using ccspan.rep_eq space_as_set_inject top_ccsubspace.rep_eq closure_UNIV is_generator_set by metis lemma ccspan_Inf_def: \ccspan A = Inf {S. A \ space_as_set S}\ for A::\('a::cbanach) set\ proof - have \x \ space_as_set (ccspan A) \ x \ space_as_set (Inf {S. A \ space_as_set S})\ for x::'a proof- assume \x \ space_as_set (ccspan A)\ hence "x \ closure (cspan A)" by (simp add: ccspan.rep_eq) hence \x \ closure (complex_vector.span A)\ unfolding ccspan_def by simp hence \\ y::nat \ 'a. (\ n. y n \ (complex_vector.span A)) \ y \ x\ by (simp add: closure_sequential) then obtain y where \\ n. y n \ (complex_vector.span A)\ and \y \ x\ by blast have \y n \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ for n using \\ n. y n \ (complex_vector.span A)\ by auto have \closed_csubspace S \ closed S\ for S::\'a set\ by (simp add: closed_csubspace.closed) hence \closed ( \ {S. (complex_vector.span A) \ S \ closed_csubspace S})\ by simp hence \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using \y \ x\ using \\n. y n \ \ {S. complex_vector.span A \ S \ closed_csubspace S}\ closed_sequentially by blast moreover have \{S. A \ S \ closed_csubspace S} \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using Collect_mono_iff by (simp add: Collect_mono_iff closed_csubspace.subspace complex_vector.span_minimal) ultimately have \x \ \ {S. A \ S \ closed_csubspace S}\ by blast moreover have "(x::'a) \ \ {x. A \ x \ closed_csubspace x}" if "(x::'a) \ \ {S. A \ S \ closed_csubspace S}" for x :: 'a and A :: "'a set" using that by simp ultimately show \x \ space_as_set (Inf {S. A \ space_as_set S})\ apply transfer. qed moreover have \x \ space_as_set (Inf {S. A \ space_as_set S}) \ x \ space_as_set (ccspan A)\ for x::'a proof- assume \x \ space_as_set (Inf {S. A \ space_as_set S})\ hence \x \ \ {S. A \ S \ closed_csubspace S}\ apply transfer by blast moreover have \{S. (complex_vector.span A) \ S \ closed_csubspace S} \ {S. A \ S \ closed_csubspace S}\ using Collect_mono_iff complex_vector.span_superset by fastforce ultimately have \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ by blast thus \x \ space_as_set (ccspan A)\ by (metis (no_types, lifting) Inter_iff space_as_set closure_subset mem_Collect_eq ccspan.rep_eq) qed ultimately have \space_as_set (ccspan A) = space_as_set (Inf {S. A \ space_as_set S})\ by blast thus ?thesis using space_as_set_inject by auto qed lemma cspan_singleton_scaleC[simp]: "(a::complex)\0 \ cspan { a *\<^sub>C \ } = cspan {\}" for \::"'a::complex_vector" by (smt complex_vector.dependent_single complex_vector.independent_insert complex_vector.scale_eq_0_iff complex_vector.span_base complex_vector.span_redundant complex_vector.span_scale doubleton_eq_iff insert_absorb insert_absorb2 insert_commute singletonI) lemma closure_is_closed_csubspace[simp]: fixes S::\'a::complex_normed_vector set\ assumes \csubspace S\ shows \closed_csubspace (closure S)\ using assms closed_csubspace.intro closure_is_csubspace by blast lemma ccspan_singleton_scaleC[simp]: "(a::complex)\0 \ ccspan {a *\<^sub>C \} = ccspan {\}" apply transfer by simp lemma clinear_continuous_at: assumes \bounded_clinear f\ shows \isCont f x\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_at) lemma clinear_continuous_within: assumes \bounded_clinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_within) lemma antilinear_continuous_at: assumes \bounded_antilinear f\ shows \isCont f x\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_at) lemma antilinear_continuous_within: assumes \bounded_antilinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_within) lemma bounded_clinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_clinear A\ and \bounded_clinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - have eq': \A t = B t\ if \t \ cspan G\ for t using _ _ that eq apply (rule complex_vector.linear_eq_on) by (auto simp: assms bounded_clinear.clinear) have \A t - B t = 0\ using _ _ t apply (rule continuous_constant_on_closure) by (auto simp add: eq' assms(1) assms(2) clinear_continuous_at continuous_at_imp_continuous_on) then show ?thesis by auto qed instantiation ccsubspace :: ("{complex_vector,topological_space}") "order" begin lift_definition less_eq_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_eq_ccsubspace_def[code del] lift_definition less_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_ccsubspace_def[code del] instance proof fix x y z :: "'a ccsubspace" show "(x < y) = (x \ y \ \ y \ x)" by (simp add: less_eq_ccsubspace.rep_eq less_le_not_le less_ccsubspace.rep_eq) show "x \ x" by (simp add: less_eq_ccsubspace.rep_eq) show "x \ z" if "x \ y" and "y \ z" using that less_eq_ccsubspace.rep_eq by auto show "x = y" if "x \ y" and "y \ x" using that by (simp add: space_as_set_inject less_eq_ccsubspace.rep_eq) qed end lemma ccspan_leqI: assumes \M \ space_as_set S\ shows \ccspan M \ S\ using assms apply transfer by (simp add: closed_csubspace.closed closure_minimal complex_vector.span_minimal) lemma ccspan_mono: assumes \A \ B\ shows \ccspan A \ ccspan B\ apply (transfer fixing: A B) by (simp add: assms closure_mono complex_vector.span_mono) lemma ccsubspace_leI: assumes t1: "space_as_set A \ space_as_set B" shows "A \ B" using t1 apply transfer by - lemma ccspan_of_empty[simp]: "ccspan {} = bot" proof transfer show "closure (cspan {}) = {0::'a}" by simp qed instantiation ccsubspace :: ("{complex_vector,topological_space}") inf begin lift_definition inf_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" is "(\)" by simp instance .. end lemma space_as_set_inf[simp]: "space_as_set (A \ B) = space_as_set A \ space_as_set B" by (rule inf_ccsubspace.rep_eq) instantiation ccsubspace :: ("{complex_vector,topological_space}") order_top begin instance proof show "a \ \" for a :: "'a ccsubspace" apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") order_bot begin instance proof show "(\::'a ccsubspace) \ a" for a :: "'a ccsubspace" apply transfer apply auto using closed_csubspace.subspace complex_vector.subspace_0 by blast qed end instantiation ccsubspace :: ("{complex_vector,topological_space}") semilattice_inf begin instance proof fix x y z :: \'a ccsubspace\ show "x \ y \ x" apply transfer by simp show "x \ y \ y" apply transfer by simp show "x \ y \ z" if "x \ y" and "x \ z" using that apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") zero begin definition zero_ccsubspace :: "'a ccsubspace" where [simp]: "zero_ccsubspace = bot" lemma zero_ccsubspace_transfer[transfer_rule]: \pcr_ccsubspace (=) {0} 0\ unfolding zero_ccsubspace_def by transfer_prover instance .. end lemma ccspan_0[simp]: \ccspan {0} = 0\ apply transfer by simp definition \rel_ccsubspace R x y = rel_set R (space_as_set x) (space_as_set y)\ lemma left_unique_rel_ccsubspace[transfer_rule]: \left_unique (rel_ccsubspace R)\ if \left_unique R\ proof (rule left_uniqueI) fix S T :: \'a ccsubspace\ and U assume assms: \rel_ccsubspace R S U\ \rel_ccsubspace R T U\ have \space_as_set S = space_as_set T\ apply (rule left_uniqueD) using that apply (rule left_unique_rel_set) using assms unfolding rel_ccsubspace_def by auto then show \S = T\ by (simp add: space_as_set_inject) qed lemma right_unique_rel_ccsubspace[transfer_rule]: \right_unique (rel_ccsubspace R)\ if \right_unique R\ by (metis rel_ccsubspace_def right_unique_def right_unique_rel_set space_as_set_inject that) lemma bi_unique_rel_ccsubspace[transfer_rule]: \bi_unique (rel_ccsubspace R)\ if \bi_unique R\ by (metis (no_types, lifting) bi_unique_def bi_unique_rel_set rel_ccsubspace_def space_as_set_inject that) lemma converse_rel_ccsubspace: \conversep (rel_ccsubspace R) = rel_ccsubspace (conversep R)\ by (auto simp: rel_ccsubspace_def[abs_def]) lemma space_as_set_top[simp]: \space_as_set top = UNIV\ by (rule top_ccsubspace.rep_eq) lemma ccsubspace_eqI: assumes \\x. x \ space_as_set S \ x \ space_as_set T\ shows \S = T\ by (metis Abs_clinear_space_cases Abs_clinear_space_inverse antisym assms subsetI) lemma ccspan_remove_0: \ccspan (A - {0}) = ccspan A\ apply transfer by auto lemma sgn_in_spaceD: \\ \ space_as_set A\ if \sgn \ \ space_as_set A\ and \\ \ 0\ for \ :: \_ :: complex_normed_vector\ using that apply (transfer fixing: \) by (metis closed_csubspace.subspace complex_vector.subspace_scale divideC_field_simps(1) scaleR_eq_0_iff scaleR_scaleC sgn_div_norm sgn_zero_iff) lemma sgn_in_spaceI: \sgn \ \ space_as_set A\ if \\ \ space_as_set A\ for \ :: \_ :: complex_normed_vector\ using that by (auto simp: sgn_div_norm scaleR_scaleC complex_vector.subspace_scale) lemma ccsubspace_leI_unit: fixes A B :: \_ :: complex_normed_vector ccsubspace\ assumes "\\. norm \ = 1 \ \ \ space_as_set A \ \ \ space_as_set B" shows "A \ B" proof (rule ccsubspace_leI, rule subsetI) fix \ assume \A: \\ \ space_as_set A\ show \\ \ space_as_set B\ apply (cases \\ = 0\) apply simp using assms[of \sgn \\] \A sgn_in_spaceD sgn_in_spaceI by (auto simp: norm_sgn) qed lemma kernel_is_closed_csubspace[simp]: assumes a1: "bounded_clinear f" shows "closed_csubspace (f -` {0})" proof- have \csubspace (f -` {0})\ using assms bounded_clinear.clinear complex_vector.linear_subspace_vimage complex_vector.subspace_single_0 by blast have "L \ {x. f x = 0}" if "r \ L" and "\ n. r n \ {x. f x = 0}" for r and L proof- have d1: \\ n. f (r n) = 0\ using that(2) by auto have \(\ n. f (r n)) \ f L\ using assms clinear_continuous_at continuous_within_tendsto_compose' that(1) by fastforce hence \(\ n. 0) \ f L\ using d1 by simp hence \f L = 0\ using limI by fastforce thus ?thesis by blast qed then have s3: \closed (f -` {0})\ using closed_sequential_limits by force with \csubspace (f -` {0})\ show ?thesis using closed_csubspace.intro by blast qed subsection \Closed sums\ definition closed_sum:: \'a::{semigroup_add,topological_space} set \ 'a set \ 'a set\ where \closed_sum A B = closure (A + B)\ notation closed_sum (infixl "+\<^sub>M" 65) lemma closed_sum_comm: \A +\<^sub>M B = B +\<^sub>M A\ for A B :: "_::ab_semigroup_add" by (simp add: add.commute closed_sum_def) lemma closed_sum_left_subset: \0 \ B \ A \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.right_neutral closed_sum_def closure_subset in_mono set_plus_intro subsetI) lemma closed_sum_right_subset: \0 \ A \ B \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.left_neutral closed_sum_def closure_subset set_plus_intro subset_iff) lemma finite_cspan_closed_csubspace: assumes "finite (S::'a::complex_normed_vector set)" shows "closed_csubspace (cspan S)" by (simp add: assms closed_csubspace.intro) lemma closed_sum_is_sup: fixes A B C:: \('a::{complex_vector,topological_space}) set\ assumes \closed_csubspace C\ assumes \A \ C\ and \B \ C\ shows \(A +\<^sub>M B) \ C\ proof - have \A + B \ C\ using assms unfolding set_plus_def using closed_csubspace.subspace complex_vector.subspace_add by blast then show \(A +\<^sub>M B) \ C\ unfolding closed_sum_def using \closed_csubspace C\ by (simp add: closed_csubspace.closed closure_minimal) qed lemma closed_subspace_closed_sum: fixes A B::"('a::complex_normed_vector) set" assumes a1: \csubspace A\ and a2: \csubspace B\ shows \closed_csubspace (A +\<^sub>M B)\ using a1 a2 closed_sum_def by (metis closure_is_closed_csubspace csubspace_set_plus) lemma closed_sum_assoc: fixes A B C::"'a::real_normed_vector set" shows \A +\<^sub>M (B +\<^sub>M C) = (A +\<^sub>M B) +\<^sub>M C\ proof - have \A + closure B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \A +\<^sub>M (B +\<^sub>M C) = closure (A + (B + C))\ unfolding closed_sum_def by (meson antisym_conv closed_closure closure_minimal closure_mono closure_subset equalityD1 set_plus_mono2) moreover have \closure A + B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \(A +\<^sub>M B) +\<^sub>M C = closure ((A + B) + C)\ unfolding closed_sum_def by (meson closed_closure closure_minimal closure_mono closure_subset eq_iff set_plus_mono2) ultimately show ?thesis by (simp add: ab_semigroup_add_class.add_ac(1)) qed lemma closed_sum_zero_left[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \{0} +\<^sub>M A = closure A\ unfolding closed_sum_def by (metis add.left_neutral set_zero) lemma closed_sum_zero_right[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \A +\<^sub>M {0} = closure A\ unfolding closed_sum_def by (metis add.right_neutral set_zero) lemma closed_sum_closure_right[simp]: fixes A B :: \'a::real_normed_vector set\ shows \A +\<^sub>M closure B = A +\<^sub>M B\ by (metis closed_sum_assoc closed_sum_def closed_sum_zero_right closure_closure) lemma closed_sum_closure_left[simp]: fixes A B :: \'a::real_normed_vector set\ shows \closure A +\<^sub>M B = A +\<^sub>M B\ by (simp add: closed_sum_comm) lemma closed_sum_mono_left: assumes \A \ B\ shows \A +\<^sub>M C \ B +\<^sub>M C\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) lemma closed_sum_mono_right: assumes \A \ B\ shows \C +\<^sub>M A \ C +\<^sub>M B\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) instantiation ccsubspace :: (complex_normed_vector) sup begin lift_definition sup_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" \ \Note that \<^term>\A+B\ would not be a closed subspace, we need the closure. See, e.g., \<^url>\https://math.stackexchange.com/a/1786792/403528\.\ is "\A B::'a set. A +\<^sub>M B" by (simp add: closed_subspace_closed_sum) instance .. end lemma closed_sum_cspan[simp]: shows \cspan X +\<^sub>M cspan Y = closure (cspan (X \ Y))\ by (smt (verit, best) Collect_cong closed_sum_def complex_vector.span_Un set_plus_def) lemma closure_image_closed_sum: assumes \bounded_linear U\ shows \closure (U ` (A +\<^sub>M B)) = closure (U ` A) +\<^sub>M closure (U ` B)\ proof - have \closure (U ` (A +\<^sub>M B)) = closure (U ` closure (closure A + closure B))\ unfolding closed_sum_def by (smt (verit, best) closed_closure closure_minimal closure_mono closure_subset closure_sum set_plus_mono2 subset_antisym) also have \\ = closure (U ` (closure A + closure B))\ using assms closure_bounded_linear_image_subset_eq by blast also have \\ = closure (U ` closure A + U ` closure B)\ apply (subst image_set_plus) by (simp_all add: assms bounded_linear.linear) also have \\ = closure (closure (U ` A) + closure (U ` B))\ by (smt (verit, ccfv_SIG) assms closed_closure closure_bounded_linear_image_subset closure_bounded_linear_image_subset_eq closure_minimal closure_mono closure_sum dual_order.eq_iff set_plus_mono2) also have \\ = closure (U ` A) +\<^sub>M closure (U ` B)\ using closed_sum_def by blast finally show ?thesis by - qed lemma ccspan_union: "ccspan A \ ccspan B = ccspan (A \ B)" apply transfer by simp instantiation ccsubspace :: (complex_normed_vector) "Sup" begin lift_definition Sup_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\S. closure (complex_vector.span (Union S))\ proof show "csubspace (closure (complex_vector.span (\ S::'a set)))" if "\x::'a set. x \ S \ closed_csubspace x" for S :: "'a set set" using that by (simp add: closure_is_closed_csubspace) show "closed (closure (complex_vector.span (\ S::'a set)))" if "\x. (x::'a set) \ S \ closed_csubspace x" for S :: "'a set set" using that by simp qed instance.. end instance ccsubspace :: ("{complex_normed_vector}") semilattice_sup proof fix x y z :: \'a ccsubspace\ show \x \ sup x y\ apply transfer by (simp add: closed_csubspace_def closed_sum_left_subset complex_vector.subspace_0) show "y \ sup x y" apply transfer by (simp add: closed_csubspace_def closed_sum_right_subset complex_vector.subspace_0) show "sup x y \ z" if "x \ z" and "y \ z" using that apply transfer apply (rule closed_sum_is_sup) by auto qed instance ccsubspace :: (complex_normed_vector) complete_lattice proof show "Inf A \ x" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by auto have b1: "z \ \ A" if "Ball A closed_csubspace" and "closed_csubspace z" and "(\x. closed_csubspace x \ x \ A \ z \ x)" for z::"'a set" and A using that by auto show "z \ Inf A" if "\x::'a ccsubspace. x \ A \ z \ x" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer using b1 by blast show "x \ Sup A" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by (meson Union_upper closure_subset complex_vector.span_superset dual_order.trans) show "Sup A \ z" if "\x::'a ccsubspace. x \ A \ x \ z" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer proof - fix A :: "'a set set" and z :: "'a set" assume A_closed: "Ball A closed_csubspace" assume "closed_csubspace z" assume in_z: "\x. closed_csubspace x \ x \ A \ x \ z" from A_closed in_z have \V \ z\ if \V \ A\ for V by (simp add: that) then have \\ A \ z\ by (simp add: Sup_le_iff) with \closed_csubspace z\ show "closure (cspan (\ A)) \ z" by (simp add: closed_csubspace_def closure_minimal complex_vector.span_def subset_hull) qed show "Inf {} = (top::'a ccsubspace)" using \\z A. (\x. x \ A \ z \ x) \ z \ Inf A\ top.extremum_uniqueI by auto show "Sup {} = (bot::'a ccsubspace)" using \\z A. (\x. x \ A \ x \ z) \ Sup A \ z\ bot.extremum_uniqueI by auto qed instantiation ccsubspace :: (complex_normed_vector) comm_monoid_add begin definition plus_ccsubspace :: "'a ccsubspace \ _ \ _" where [simp]: "plus_ccsubspace = sup" instance proof fix a b c :: \'a ccsubspace\ show "a + b + c = a + (b + c)" using sup.assoc by auto show "a + b = b + a" by (simp add: sup.commute) show "0 + a = a" by (simp add: zero_ccsubspace_def) qed end lemma ccsubspace_plus_sup: "y \ x \ z \ x \ y + z \ x" for x y z :: "'a::complex_normed_vector ccsubspace" unfolding plus_ccsubspace_def by auto lemma ccsubspace_Sup_empty: "Sup {} = (0::_ ccsubspace)" unfolding zero_ccsubspace_def by auto lemma ccsubspace_add_right_incr[simp]: "a \ a + c" for a::"_ ccsubspace" by (simp add: add_increasing2) lemma ccsubspace_add_left_incr[simp]: "a \ c + a" for a::"_ ccsubspace" by (simp add: add_increasing) lemma sum_bot_ccsubspace[simp]: \(\x\X. \) = (\ :: _ ccsubspace)\ by (simp flip: zero_ccsubspace_def) subsection \Conjugate space\ typedef 'a conjugate_space = "UNIV :: 'a set" morphisms from_conjugate_space to_conjugate_space .. setup_lifting type_definition_conjugate_space instantiation conjugate_space :: (complex_vector) complex_vector begin lift_definition scaleC_conjugate_space :: \complex \ 'a conjugate_space \ 'a conjugate_space\ is \\c x. cnj c *\<^sub>C x\. lift_definition scaleR_conjugate_space :: \real \ 'a conjugate_space \ 'a conjugate_space\ is \\r x. r *\<^sub>R x\. lift_definition plus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(+)". lift_definition uminus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is \\x. -x\. lift_definition zero_conjugate_space :: "'a conjugate_space" is 0. lift_definition minus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(-)". instance apply (intro_classes; transfer) by (simp_all add: scaleR_scaleC scaleC_add_right scaleC_left.add) end instantiation conjugate_space :: (complex_normed_vector) complex_normed_vector begin lift_definition sgn_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is "sgn". lift_definition norm_conjugate_space :: "'a conjugate_space \ real" is norm. lift_definition dist_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ real" is dist. lift_definition uniformity_conjugate_space :: "('a conjugate_space \ 'a conjugate_space) filter" is uniformity. lift_definition open_conjugate_space :: "'a conjugate_space set \ bool" is "open". instance apply (intro_classes; transfer) by (simp_all add: dist_norm sgn_div_norm open_uniformity uniformity_dist norm_triangle_ineq) end instantiation conjugate_space :: (cbanach) cbanach begin instance apply intro_classes unfolding Cauchy_def convergent_def LIMSEQ_def apply transfer using Cauchy_convergent unfolding Cauchy_def convergent_def LIMSEQ_def by metis end lemma bounded_antilinear_to_conjugate_space[simp]: \bounded_antilinear to_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma bounded_antilinear_from_conjugate_space[simp]: \bounded_antilinear from_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma antilinear_to_conjugate_space[simp]: \antilinear to_conjugate_space\ by (rule antilinearI; transfer, auto) lemma antilinear_from_conjugate_space[simp]: \antilinear from_conjugate_space\ by (rule antilinearI; transfer, auto) lemma cspan_to_conjugate_space[simp]: "cspan (to_conjugate_space ` X) = to_conjugate_space ` cspan X" unfolding complex_vector.span_def complex_vector.subspace_def hull_def apply transfer apply simp by (metis (no_types, opaque_lifting) complex_cnj_cnj) lemma surj_to_conjugate_space[simp]: "surj to_conjugate_space" by (meson surj_def to_conjugate_space_cases) lemmas has_derivative_scaleC[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] lemma norm_to_conjugate_space[simp]: \norm (to_conjugate_space x) = norm x\ by (fact norm_conjugate_space.abs_eq) lemma norm_from_conjugate_space[simp]: \norm (from_conjugate_space x) = norm x\ by (simp add: norm_conjugate_space.rep_eq) lemma closure_to_conjugate_space: \closure (to_conjugate_space ` X) = to_conjugate_space ` closure X\ proof - have 1: \to_conjugate_space ` closure X \ closure (to_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = to_conjugate_space ` from_conjugate_space ` closure (to_conjugate_space ` X)\ by (simp add: from_conjugate_space_inverse image_image) also have \\ \ to_conjugate_space ` closure (from_conjugate_space ` to_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = to_conjugate_space ` closure X\ by (simp add: to_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma closure_from_conjugate_space: \closure (from_conjugate_space ` X) = from_conjugate_space ` closure X\ proof - have 1: \from_conjugate_space ` closure X \ closure (from_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = from_conjugate_space ` to_conjugate_space ` closure (from_conjugate_space ` X)\ by (simp add: to_conjugate_space_inverse image_image) also have \\ \ from_conjugate_space ` closure (to_conjugate_space ` from_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = from_conjugate_space ` closure X\ by (simp add: from_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma bounded_antilinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_antilinear A\ and \bounded_antilinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - let ?A = \\x. A (from_conjugate_space x)\ and ?B = \\x. B (from_conjugate_space x)\ and ?G = \to_conjugate_space ` G\ and ?t = \to_conjugate_space t\ have \bounded_clinear ?A\ and \bounded_clinear ?B\ by (auto intro!: bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear A\] bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear B\]) moreover from eq have \\x. x \ ?G \ ?A x = ?B x\ by (metis image_iff iso_tuple_UNIV_I to_conjugate_space_inverse) moreover from t have \?t \ closure (cspan ?G)\ by (metis bounded_antilinear.bounded_linear bounded_antilinear_to_conjugate_space closure_bounded_linear_image_subset cspan_to_conjugate_space imageI subsetD) ultimately have \?A ?t = ?B ?t\ by (rule bounded_clinear_eq_on) then show \A t = B t\ by (simp add: to_conjugate_space_inverse) qed subsection \Product is a Complex Vector Space\ (* Follows closely Product_Vector.thy *) instantiation prod :: (complex_vector, complex_vector) complex_vector begin definition scaleC_prod_def: "scaleC r A = (scaleC r (fst A), scaleC r (snd A))" lemma fst_scaleC [simp]: "fst (scaleC r A) = scaleC r (fst A)" unfolding scaleC_prod_def by simp lemma snd_scaleC [simp]: "snd (scaleC r A) = scaleC r (snd A)" unfolding scaleC_prod_def by simp proposition scaleC_Pair [simp]: "scaleC r (a, b) = (scaleC r a, scaleC r b)" unfolding scaleC_prod_def by simp instance proof fix a b :: complex and x y :: "'a \ 'b" show "scaleC a (x + y) = scaleC a x + scaleC a y" by (simp add: scaleC_add_right scaleC_prod_def) show "scaleC (a + b) x = scaleC a x + scaleC b x" by (simp add: Complex_Vector_Spaces.scaleC_prod_def scaleC_left.add) show "scaleC a (scaleC b x) = scaleC (a * b) x" by (simp add: prod_eq_iff) show "scaleC 1 x = x" by (simp add: prod_eq_iff) show \(scaleR :: _ \ _ \ 'a*'b) r = (*\<^sub>C) (complex_of_real r)\ for r by (auto intro!: ext simp: scaleR_scaleC scaleC_prod_def scaleR_prod_def) qed end lemma module_prod_scale_eq_scaleC: "module_prod.scale (*\<^sub>C) (*\<^sub>C) = scaleC" apply (rule ext) apply (rule ext) apply (subst module_prod.scale_def) subgoal by unfold_locales by (simp add: scaleC_prod_def) interpretation complex_vector?: vector_space_prod "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" rewrites "scale = ((*\<^sub>C)::_\_\('a \ 'b))" and "module.dependent (*\<^sub>C) = cdependent" and "module.representation (*\<^sub>C) = crepresentation" and "module.subspace (*\<^sub>C) = csubspace" and "module.span (*\<^sub>C) = cspan" and "vector_space.extend_basis (*\<^sub>C) = cextend_basis" and "vector_space.dim (*\<^sub>C) = cdim" and "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" subgoal by unfold_locales subgoal by (fact module_prod_scale_eq_scaleC) unfolding cdependent_raw_def crepresentation_raw_def csubspace_raw_def cspan_raw_def cextend_basis_raw_def cdim_raw_def clinear_def by (rule refl)+ instance prod :: (complex_normed_vector, complex_normed_vector) complex_normed_vector proof fix c :: complex and x y :: "'a \ 'b" show "norm (c *\<^sub>C x) = cmod c * norm x" unfolding norm_prod_def apply (simp add: power_mult_distrib) apply (simp add: distrib_left [symmetric]) by (simp add: real_sqrt_mult) qed lemma cspan_Times: \cspan (S \ T) = cspan S \ cspan T\ if \0 \ S\ and \0 \ T\ proof have \fst ` cspan (S \ T) \ cspan S\ apply (subst complex_vector.linear_span_image[symmetric]) using that complex_vector.module_hom_fst by auto moreover have \snd ` cspan (S \ T) \ cspan T\ apply (subst complex_vector.linear_span_image[symmetric]) using that complex_vector.module_hom_snd by auto ultimately show \cspan (S \ T) \ cspan S \ cspan T\ by auto show \cspan S \ cspan T \ cspan (S \ T)\ proof fix x assume assm: \x \ cspan S \ cspan T\ then have \fst x \ cspan S\ by auto then obtain t1 r1 where fst_x: \fst x = (\a\t1. r1 a *\<^sub>C a)\ and [simp]: \finite t1\ and \t1 \ S\ by (auto simp add: complex_vector.span_explicit) from assm have \snd x \ cspan T\ by auto then obtain t2 r2 where snd_x: \snd x = (\a\t2. r2 a *\<^sub>C a)\ and [simp]: \finite t2\ and \t2 \ T\ by (auto simp add: complex_vector.span_explicit) define t :: \('a+'b) set\ and r :: \('a+'b) \ complex\ and f :: \('a+'b) \ ('a\'b)\ where \t = t1 <+> t2\ and \r a = (case a of Inl a1 \ r1 a1 | Inr a2 \ r2 a2)\ and \f a = (case a of Inl a1 \ (a1,0) | Inr a2 \ (0,a2))\ for a have \finite t\ by (simp add: t_def) moreover have \f ` t \ S \ T\ using \t1 \ S\ \t2 \ T\ that by (auto simp: f_def t_def) moreover have \(fst x, snd x) = (\a\t. r a *\<^sub>C f a)\ apply (simp only: fst_x snd_x) by (auto simp: t_def sum.Plus r_def f_def sum_prod) ultimately show \x \ cspan (S \ T)\ apply auto by (smt (verit, best) complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset image_subset_iff subset_iff) qed qed lemma onorm_case_prod_plus: \onorm (case_prod plus :: _ \ 'a::{real_normed_vector, not_singleton}) = sqrt 2\ proof - obtain x :: 'a where \x \ 0\ apply atomize_elim by auto show ?thesis apply (rule onormI[where x=\(x,x)\]) using norm_plus_leq_norm_prod apply force using \x \ 0\ by (auto simp add: zero_prod_def norm_prod_def real_sqrt_mult simp flip: scaleR_2) qed subsection \Copying existing theorems into sublocales\ context bounded_clinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = real.continuous lemmas uniform_limit = real.uniform_limit lemmas Cauchy = real.Cauchy end context bounded_antilinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = real.continuous lemmas uniform_limit = real.uniform_limit end context bounded_cbilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = real.tendsto lemmas isCont = real.isCont lemmas scaleR_right = real.scaleR_right lemmas scaleR_left = real.scaleR_left end context bounded_sesquilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = real.tendsto lemmas isCont = real.isCont lemmas scaleR_right = real.scaleR_right lemmas scaleR_left = real.scaleR_left end lemmas tendsto_scaleC [tendsto_intros] = bounded_cbilinear.tendsto [OF bounded_cbilinear_scaleC] unbundle no_lattice_syntax end