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,1572 +1,1576 @@ section \\Cblinfun_Matrix\ -- Matrix representation of bounded operators\ theory Cblinfun_Matrix imports Complex_L2 "Jordan_Normal_Form.Gram_Schmidt" "HOL-Analysis.Starlike" "Complex_Bounded_Operators.Extra_Jordan_Normal_Form" begin hide_const (open) Order.bottom Order.top hide_type (open) Finite_Cartesian_Product.vec hide_const (open) Finite_Cartesian_Product.mat hide_fact (open) Finite_Cartesian_Product.mat_def hide_const (open) Finite_Cartesian_Product.vec hide_fact (open) Finite_Cartesian_Product.vec_def hide_const (open) Finite_Cartesian_Product.row hide_fact (open) Finite_Cartesian_Product.row_def no_notation Finite_Cartesian_Product.vec_nth (infixl "$" 90) unbundle jnf_notation unbundle cblinfun_notation subsection \Isomorphism between vectors\ text \We define the canonical isomorphism between vectors in some complex vector space \<^typ>\'a::basis_enum\ and the complex \<^term>\n\-dimensional vectors (where \<^term>\n\ is the dimension of \<^typ>\'a\). This is possible if \<^typ>\'a\, \<^typ>\'b\ are of class \<^class>\basis_enum\ since that class fixes a finite canonical basis. Vector are represented using the \<^typ>\complex vec\ type from \<^session>\Jordan_Normal_Form\. (The isomorphism will be called \<^term>\vec_of_onb_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 + unfolding vec_of_basis_enum_def + by simp definition basis_enum_of_vec :: \complex vec \ 'a::basis_enum\ where - \basis_enum_of_vec v = + \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 + 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\ + 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)\ + 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)\ + 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, hide_lams) 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) + 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, basis_enum_of_vec y\ = y \c x" proof - have \\basis_enum_of_vec x :: 'a, 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..). + 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..). + 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" + 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) + 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 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) + 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 + 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) + 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 + 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) + 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 + by fastforce ultimately show ?thesis using \a \ \j. list_of_vec (unit_vec n i) ! j *\<^sub>C basis ! j\ - by simp + 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 + 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 + 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\) + 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) + 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 + \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) + 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" + = (canonical_basis::'a list) ! i" by (simp add: assms) - thus ?thesis + thus ?thesis unfolding basis_enum_of_vec_def - by (simp add: assms) + by (simp add: assms) qed lemma vec_of_basis_enum_ket: - "vec_of_basis_enum (ket i) = unit_vec (CARD('a)) (enum_idx i)" + "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)) + 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)) + have "dim_vec (unit_vec (CARD('a)) (enum_idx i)) = CARD('a)" - by simp + 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 + 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) + 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 + 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 + 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)" + 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)) + 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)" + 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, hide_lams) 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) + 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) + 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 + 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)" + 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)" + 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 + 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" + 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: +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: +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: +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\ + 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, \\\ 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 + 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\, + 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_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 +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" + 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))" + 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: + 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) + 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 = + 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) = + 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) +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\ lemma cblinfun_of_mat_plus: - defines "nA \ length (canonical_basis :: 'a::{basis_enum,complex_normed_vector} list)" + 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})) + "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" + "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) + 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 + 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) + 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" + "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)" + 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 + 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)" + 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)" + 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, hide_lams) 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) + 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 + 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 "\ = + 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 + 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) + 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)" + 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, y\ = \x, 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 + 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, y\ = \x, cblinfun_of_mat M *\<^sub>V y\" unfolding u_def v_def - by (simp add: cscalar_prod_vec_of_basis_enum) + 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:) + 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) + 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 + 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 + 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) + by (meson injI ket_injective list.inj_map) hence "length (Enum.enum::'b list) = length (map (ket::'b \_) (Enum.enum::'b list))" - by simp + by simp hence lengthBasisB: "CARD('b) = length BasisB" - unfolding BasisB_def canonical_basis_ell2_def using cardb + 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) + 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 "\BasisB!r, (classical_operator f) *\<^sub>V BasisA!c\ = 0" 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'" + 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] + 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 + 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 + 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) + 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 "\BasisB!r, BasisB!r'\ = 0" using h1 using BasisB_def \r < length BasisB\ \r' < length BasisB\ is_ortho_set_def is_orthonormal nth_mem by metis ultimately have "\BasisB!r, (classical_operator f) *\<^sub>V BasisA!c\ = 0" 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'" + 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 + 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 + by simp hence "map (ket::'a\_) enum_class.enum ! c = ket (enum_class.enum ! c)" - by (simp add: a2) + 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) + 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 + by simp have "(mat_of_cblinfun (classical_operator f))$$(r,c) = \BasisB!r, (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, BasisB!r\" - using w1 by simp + 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) + by (simp add: a1 a2) qed - ultimately show ?thesis + 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" + "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}" + 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) = + 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 + 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 dim_vec_of_basis_enum') + 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, hide_lams) ccspan_singleton_scaleC complex_vector.scale_eq_0_iff nonzero_imp_inverse_nonzero norm_eq_zero scaleR_scaleC scale_left_imp_eq) + 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\ + 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 = +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)) + 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" + and "Bo = gram_schmidt0 d Bv" interpret complex_vec_space d. have Av_carrier: "set Av \ carrier_vec d" unfolding Av_def apply auto by (simp add: carrier_vecI d_def dim_vec_of_basis_enum') have Bv_carrier: "set Bv \ carrier_vec d" unfolding Bv_def apply auto by (simp add: carrier_vecI d_def dim_vec_of_basis_enum') have Bo_carrier: "set Bo \ carrier_vec d" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(1)) have orth_Bo: "corthogonal Bo" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(3)) have distinct_Bo: "distinct Bo" apply (simp add: Bo_def) using Bv_carrier by (rule gram_schmidt0_result(2)) have "ccspan (set A) \ ccspan (set B) \ cspan (set A) \ cspan (set B)" apply (transfer fixing: A B) apply (subst closure_finite_cspan, simp) by (subst closure_finite_cspan, simp_all) also have "\ \ span (set Av) \ span (set Bv)" apply (simp add: Av_def Bv_def) apply (subst vec_of_basis_enum_cspan[symmetric], simp add: d_def) apply (subst vec_of_basis_enum_cspan[symmetric], simp add: d_def) by (metis inj_image_subset_iff inj_on_def vec_of_basis_enum_inverse) also have "\ \ span (set Av) \ span (set Bo)" unfolding Bo_def Av_def Bv_def apply (subst gram_schmidt0_result(4)[symmetric]) by (simp_all add: carrier_dim_vec d_def dim_vec_of_basis_enum' image_subset_iff) also have "\ \ (\v\set Av. adjuster d v Bo = - v)" proof (intro iffI ballI) fix v assume "v \ set Av" and "span (set Av) \ span (set Bo)" then have "v \ span (set Bo)" using Av_carrier span_mem by auto have "adjuster d v Bo + v = 0\<^sub>v d" apply (rule adjuster_already_in_span) using Av_carrier \v \ set Av\ Bo_carrier orth_Bo \v \ span (set Bo)\ by simp_all then show "adjuster d v Bo = - v" using Av_carrier Bo_carrier \v \ set Av\ by (metis (no_types, lifting) add.inv_equality adjuster_carrier' local.vec_neg subsetD) next fix v assume adj_minusv: "\v\set Av. adjuster d v Bo = - v" have *: "adjuster d v Bo \ span (set Bo)" if "v \ set Av" for v apply (rule adjuster_in_span) using Bo_carrier that Av_carrier distinct_Bo by auto have "v \ span (set Bo)" if "v \ set Av" for v using *[OF that] adj_minusv[rule_format, OF that] apply auto by (metis (no_types, lifting) Av_carrier Bo_carrier adjust_nonzero distinct_Bo subsetD that uminus_l_inv_vec) then show "span (set Av) \ span (set Bo)" apply auto by (meson Bo_carrier in_mono span_subsetI subsetI) qed also have "\ = is_subspace_of_vec_list d Av Bv" by (simp add: is_subspace_of_vec_list_def d_def Bo_def) finally show "ccspan (set A) \ ccspan (set B) \ is_subspace_of_vec_list d Av Bv" by simp qed -lemma cblinfun_apply_ccspan_using_vec: +lemma cblinfun_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 +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]) + 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) + 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, s /\<^sub>R norm s\" by simp also have \\s /\<^sub>R norm s, s /\<^sub>R norm s\ = \s /\<^sub>R norm s, s' /\<^sub>R norm s'\\ by (simp add: same) also have \\s /\<^sub>R norm s, s' /\<^sub>R norm s'\ = \s, 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 + 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, y\)" apply (subst (2) xcy) by simp also have "\ = cmod \x, y\" using \cmod c = 1\ by (simp add: norm_mult) finally have "\x, 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) + 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 + 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] + = 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: +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 + (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 + 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, hide_lams) 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) + 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 + 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,2981 +1,2981 @@ 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 + imports Complex_Inner_Product One_Dimensional_Spaces Banach_Steinhaus.Banach_Steinhaus "HOL-Types_To_Sets.Types_To_Sets" Complex_Bounded_Linear_Function0 begin subsection \Misc basic facts and declarations\ notation cblinfun_apply (infixr "*\<^sub>V" 70) lemma id_cblinfun_apply[simp]: "id_cblinfun *\<^sub>V \ = \" apply transfer by simp lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" apply transfer - by (simp add: clinear_continuous_at) + 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" + shows "f = g" apply (rule cblinfun_eq_on_UNIV_span[where basis=basis]) using assms is_generator_set is_cindependent_set by auto lemma cblinfun_eq_0_on_canonical_basis: fixes f ::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = 0" shows "f = 0" by (simp add: assms cblinfun_eq_on_canonical_basis) lemma cinner_canonical_basis_eq_0: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = 0" shows "F = 0" proof- have "F *\<^sub>V u = 0" if "u\basisA" for u proof- have "\v. v\basisB \ \v, F *\<^sub>V u\ = 0" by (simp add: assms(3) that) moreover have "(\v. v\basisB \ \v, x\ = 0) \ x = 0" for x - proof- - assume r1: "\v. v\basisB \ \v, x\ = 0" + proof- + assume r1: "\v. v\basisB \ \v, x\ = 0" have "\v, x\ = 0" for v proof- have "cspan basisB = UNIV" - using basisB_def is_generator_set by auto + using basisB_def is_generator_set by auto hence "v \ cspan basisB" by (smt iso_tuple_UNIV_I) hence "\t s. v = (\a\t. s a *\<^sub>C a) \ finite t \ t \ basisB" using complex_vector.span_explicit by (smt mem_Collect_eq) then obtain t s where b1: "v = (\a\t. s a *\<^sub>C a)" and b2: "finite t" and b3: "t \ basisB" by blast have "\v, x\ = \(\a\t. s a *\<^sub>C a), x\" by (simp add: b1) also have "\ = (\a\t. \s a *\<^sub>C a, x\)" using cinner_sum_left by blast also have "\ = (\a\t. cnj (s a) * \a, x\)" by auto also have "\ = 0" using b3 r1 subsetD by force finally show ?thesis by simp qed thus ?thesis - by (simp add: \\v. \v, x\ = 0\ cinner_extensionality) + by (simp add: \\v. \v, x\ = 0\ cinner_extensionality) qed ultimately show ?thesis by simp qed thus ?thesis - using basisA_def cblinfun_eq_0_on_canonical_basis by auto + using basisA_def cblinfun_eq_0_on_canonical_basis by auto qed lemma cinner_canonical_basis_eq: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = \v, G *\<^sub>V u\" shows "F = G" proof- define H where "H = F - G" have "\u v. u\basisA \ v\basisB \ \v, H *\<^sub>V u\ = 0" unfolding H_def - by (simp add: assms(3) cinner_diff_right minus_cblinfun.rep_eq) + 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) + by (simp add: basisA_def basisB_def cinner_canonical_basis_eq_0) thus ?thesis unfolding H_def by simp qed lemma cinner_canonical_basis_eq': defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \F *\<^sub>V u, v\ = \G *\<^sub>V u, v\" shows "F = G" using cinner_canonical_basis_eq assms by (metis cinner_commute') lemma cblinfun_norm_approx_witness: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ > 0\ shows \\\. norm (A *\<^sub>V \) \ norm A - \ \ norm \ = 1\ proof (transfer fixing: \) fix A :: \'a \ 'b\ assume [simp]: \bounded_clinear A\ have \\y\{norm (A x) |x. norm x = 1}. y > \ {norm (A x) |x. norm x = 1} - \\ apply (rule Sup_real_close) using assms by (auto simp: ex_norm1 bounded_clinear.bounded_linear bdd_above_norm_f) also have \\ {norm (A x) |x. norm x = 1} = onorm A\ by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear onorm_sphere) - finally + 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)\ + 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\ + 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) + 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) + 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 +qed subsection \Relationship to real bounded operators (\<^typ>\_ \\<^sub>L _\)\ instantiation blinfun :: (real_normed_vector, complex_normed_vector) "complex_normed_vector" begin lift_definition scaleC_blinfun :: \complex \ ('a::real_normed_vector, 'b::complex_normed_vector) blinfun \ ('a, 'b) blinfun\ - is \\ c::complex. \ f::'a\'b. (\ x. c *\<^sub>C (f x) )\ + 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) + 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\ + 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) + 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\ + 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) + 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))\)\ + \ (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 + 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) + 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 ) + 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) ) + 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) ) + 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 + 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) + 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 + 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) + by (simp add: ereal_SUP) thus ?thesis - by simp + 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 + 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\ + 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 +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 :: +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 + 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 + using that bounded_linear_bounded_clinear by auto +qed lemma blinfun_of_cblinfun_zero: \blinfun_of_cblinfun 0 = 0\ apply transfer by simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ apply transfer by auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_norm: fixes f::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ shows \norm f = norm (blinfun_of_cblinfun f)\ apply transfer by auto subsection \Composition\ lemma blinfun_of_cblinfun_cblinfun_compose: fixes f::\'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector\ and g::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b\ shows \blinfun_of_cblinfun (f o\<^sub>C\<^sub>L g) = (blinfun_of_cblinfun f) o\<^sub>L (blinfun_of_cblinfun g)\ apply transfer by auto -lemma cblinfun_compose_assoc: +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" + 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" + 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]: +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]: +lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" apply transfer by auto lemma cblinfun_eq_on: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \t \ closure (cspan G)\ shows "A *\<^sub>V t = B *\<^sub>V t" using assms apply transfer using bounded_clinear_eq_on by blast lemma cblinfun_eq_gen_eqI: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \ccspan G = \\ shows "A = B" apply (rule cblinfun_eqI) apply (rule cblinfun_eq_on[where G=G]) using assms apply auto by (metis ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) lemma cblinfun_compose_add_left: \(a + b) o\<^sub>C\<^sub>L c = (a o\<^sub>C\<^sub>L c) + (b o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_add_right: \a o\<^sub>C\<^sub>L (b + c) = (a o\<^sub>C\<^sub>L b) + (a o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) lemma cbilinear_cblinfun_compose[simp]: "cbilinear cblinfun_compose" by (auto intro!: clinearI simp add: cbilinear_def bounded_cbilinear.add_left bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) subsection \Adjoint\ lift_definition adj :: "'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ 'b \\<^sub>C\<^sub>L 'a" ("_*" [99] 100) is cadjoint by (fact cadjoint_bounded_clinear) lemma id_cblinfun_adjoint[simp]: "id_cblinfun* = id_cblinfun" apply transfer using cadjoint_id by (metis eq_id_iff) -lemma double_adj[simp]: "(A*)* = A" +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\ + and A::\'b \\<^sub>C\<^sub>L 'c::complex_inner\ shows "(A o\<^sub>C\<^sub>L B)* = (B*) o\<^sub>C\<^sub>L (A*)" proof transfer fix A :: \'b \ 'c\ and B :: \'a \ 'b\ assume \bounded_clinear A\ and \bounded_clinear B\ hence \bounded_clinear (A \ B)\ by (simp add: comp_bounded_clinear) have \\ (A \ B) u, v \ = \ u, (B\<^sup>\ \ A\<^sup>\) v \\ for u v - by (metis (no_types, lifting) cadjoint_univ_prop \bounded_clinear A\ \bounded_clinear B\ cinner_commute' comp_def) + 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*)" +lemma scaleC_adj[simp]: "(a *\<^sub>C A)* = (cnj a) *\<^sub>C (A*)" apply transfer by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear bounded_clinear_def complex_vector.linear_scale scaleC_cadjoint) -lemma scaleR_adj[simp]: "(a *\<^sub>R A)* = a *\<^sub>R (A*)" +lemma scaleR_adj[simp]: "(a *\<^sub>R A)* = a *\<^sub>R (A*)" by (simp add: scaleR_scaleC) lemma adj_plus: \(A + B)* = (A*) + (B*)\ proof transfer fix A B::\'b \ 'a\ assume a1: \bounded_clinear A\ and a2: \bounded_clinear B\ define F where \F = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ define G where \G = (\x. A x + B x)\ have \bounded_clinear G\ unfolding G_def by (simp add: a1 a2 bounded_clinear_add) moreover have \\F u, v\ = \u, G v\\ for u v unfolding F_def G_def using cadjoint_univ_prop a1 a2 cinner_add_left - by (simp add: cadjoint_univ_prop cinner_add_left cinner_add_right) + by (simp add: cadjoint_univ_prop cinner_add_left cinner_add_right) ultimately have \F = G\<^sup>\ \ using cadjoint_eqI by blast thus \(\x. A x + B x)\<^sup>\ = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ unfolding F_def G_def by auto qed -lemma cinner_sup_norm_cblinfun: +lemma cinner_sup_norm_cblinfun: fixes A :: \'a::{complex_normed_vector,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ shows \norm A = (SUP (\,\). cmod (cinner \ (A *\<^sub>V \)) / (norm \ * norm \))\ apply transfer apply (rule cinner_sup_onorm) by (simp add: bounded_clinear.bounded_linear) lemma cinner_adj_left: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\G* *\<^sub>V x, y\ = \x, G *\<^sub>V y\\ apply transfer using cadjoint_univ_prop by blast lemma cinner_adj_right: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" - shows \\x, G* *\<^sub>V y\ = \G *\<^sub>V x, y\\ + shows \\x, G* *\<^sub>V y\ = \G *\<^sub>V x, y\\ apply transfer using cadjoint_univ_prop' by blast lemma adj_0[simp]: \0* = 0\ by (metis add_cancel_right_left adj_plus) -lemma norm_adj[simp]: \norm (A*) = norm A\ +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) + apply (metis (full_types) bounded_clinear_def complex_vector.linear_0) by auto then show \norm (A*) = norm A\ by simp qed lemma antilinear_adj[simp]: \antilinear adj\ apply (rule antilinearI) by (auto simp add: adj_plus) lemma bounded_antilinear_adj[bounded_antilinear, simp]: \bounded_antilinear adj\ by (auto intro!: antilinearI exI[of _ 1] simp: bounded_antilinear_def bounded_antilinear_axioms_def adj_plus) lemma adjoint_eqI: fixes G:: \'b::chilbert_space \\<^sub>C\<^sub>L 'a::chilbert_space\ and F:: \'a \\<^sub>C\<^sub>L 'b\ assumes \\x y. \(cblinfun_apply F) x, y\ = \x, (cblinfun_apply G) y\\ shows \F = G*\ using assms apply transfer using cadjoint_eqI by auto -lemma cinner_real_hermiteanI: +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) + 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 qed then have 1: \(norm A)\<^sup>2 \ norm (A o\<^sub>C\<^sub>L A*)\ by (metis field_le_mult_one_interval less_eq_real_def ordered_field_class.sign_simps(5)) have 2: \norm (A o\<^sub>C\<^sub>L A*) \ (norm A)\<^sup>2\ proof (rule norm_cblinfun_bound) show \0 \ (norm A)\<^sup>2\ by simp fix \ have \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) = norm (A *\<^sub>V A* *\<^sub>V \)\ by auto also have \\ \ norm A * norm (A* *\<^sub>V \)\ by (simp add: norm_cblinfun) also have \\ \ norm A * norm (A*) * norm \\ by (metis mult.assoc norm_cblinfun norm_imp_pos_and_ge ordered_comm_semiring_class.comm_mult_left_mono) also have \\ = (norm A)\<^sup>2 * norm \\ by (simp add: power2_eq_square) finally show \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) \ (norm A)\<^sup>2 * norm \\ by - qed from 1 2 show ?thesis by simp next case False then have [simp]: \class.CARD_1 TYPE('b)\ by (rule not_singleton_vs_CARD_1) have \A = 0\ apply (rule cblinfun_to_CARD_1_0[internalize_sort' 'b]) by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) then show ?thesis by auto qed subsection \Unitaries / isometries\ definition isometry::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun\ definition unitary::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \unitary U \ (U* o\<^sub>C\<^sub>L U = id_cblinfun) \ (U o\<^sub>C\<^sub>L U* = id_cblinfun)\ lemma unitary_twosided_isometry: "unitary U \ isometry U \ isometry (U*)" unfolding unitary_def isometry_def by simp -lemma isometryD[simp]: "isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" +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" +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" +lemma unitary_adj[simp]: "unitary (U*) = unitary U" unfolding unitary_def by auto -lemma isometry_cblinfun_compose[simp]: - assumes "isometry A" and "isometry B" +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 + 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: + 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) + by (simp add: unitary_def) lemma orthogonal_on_basis_is_isometry: assumes spanB: \ccspan B = \\ assumes orthoU: \\b c. b\B \ c\B \ cinner (U *\<^sub>V b) (U *\<^sub>V c) = cinner b c\ shows \isometry U\ proof - have [simp]: \b \ closure (cspan B)\ for b using spanB apply transfer by simp have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ and \\\B\ for \ \ by (simp add: cinner_adj_left orthoU that(1) that(2)) have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ for \ \ apply (rule bounded_clinear_eq_on[where t=\ and G=B]) using bounded_clinear_cinner_right *[OF that] by auto have \U* *\<^sub>V U *\<^sub>V \ = \\ if \\\B\ for \ apply (rule cinner_extensionality) apply (subst cinner_eq_flip) by (simp add: * that) then have \U* o\<^sub>C\<^sub>L U = id_cblinfun\ by (metis cblinfun_apply_cblinfun_compose cblinfun_eq_gen_eqI cblinfun_id_cblinfun_apply spanB) then show \isometry U\ using isometry_def by blast qed subsection \Images\ (* Closure is necessary. See email 47a3bb3d-3cc3-0934-36eb-3ef0f7b70a85@ut.ee *) lift_definition cblinfun_image :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a ccsubspace \ 'b ccsubspace\ (infixr "*\<^sub>S" 70) is "\A S. closure (A ` S)" using bounded_clinear_def closed_closure closed_csubspace.intro - by (simp add: bounded_clinear_def complex_vector.linear_subspace_image closure_is_closed_csubspace) + 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]: +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]: +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 + 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) + 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) + 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]: +lemma cblinfun_image_id[simp]: "id_cblinfun *\<^sub>S \ = \" apply transfer - by (simp add: closed_csubspace.closed) - -lemma cblinfun_compose_image: + 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] +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" + 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) + 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" + fixes S :: "'a::cbanach ccsubspace" and A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::cbanach" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and "ccspan G \ S" shows "A *\<^sub>S S = B *\<^sub>S S" proof (use assms in transfer) fix G :: "'a set" and A :: "'a \ 'b" and B :: "'a \ 'b" and S :: "'a set" assume a1: "bounded_clinear A" assume a2: "bounded_clinear B" assume a3: "\x. x \ G \ A x = B x" assume a4: "S \ closure (cspan G)" have "A ` closure S = B ` closure S" by (smt (verit, best) UnCI a1 a2 a3 a4 bounded_clinear_eq_on closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis Complex_Vector_Spaces0.bounded_clinear.bounded_linear a1 a2 closure_bounded_linear_image_subset_eq) qed lemma cblinfun_fixes_range: assumes "A o\<^sub>C\<^sub>L B = B" and "\ \ space_as_set (B *\<^sub>S top)" - shows "A *\<^sub>V \ = \" + shows "A *\<^sub>V \ = \" proof- - define rangeB rangeB' where "rangeB = space_as_set (B *\<^sub>S top)" + 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" + 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 + 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" + 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" + 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" + 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)" + 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) + 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 + by blast qed - 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) + 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) + by (simp add: closed_csubspace.closed) qed hence "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>S INFUV = INFUV" - by (metis (mono_tags, hide_lams) x cblinfun_image.rep_eq cblinfun_image_id id_cblinfun_apply image_cong + 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) + 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]: +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) + 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" +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: +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 + then show ?thesis by auto qed subsection \Sandwiches\ lift_definition sandwich :: \('a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner) \ (('a \\<^sub>C\<^sub>L 'a) \\<^sub>C\<^sub>L ('b \\<^sub>C\<^sub>L 'b))\ is \\(A::'a\\<^sub>C\<^sub>L'b) B. A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ -proof +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" +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\ + 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))\ + = (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 + 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 + by auto qed lift_definition is_Proj::\'a::chilbert_space \\<^sub>C\<^sub>L 'a \ bool\ is \\P. \M. closed_csubspace M \ is_projection_on P M\ . lemma Proj_on_own_range': fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \P o\<^sub>C\<^sub>L P = P\ and \P = P*\ shows \Proj (P *\<^sub>S top) = P\ proof- define M where "M = P *\<^sub>S top" have v3: "x \ (\x. x - P *\<^sub>V x) -` {0}" if "x \ range (cblinfun_apply P)" for x :: 'a proof- have v3_1: \cblinfun_apply P \ cblinfun_apply P = cblinfun_apply P\ by (metis \P o\<^sub>C\<^sub>L P = P\ cblinfun_compose.rep_eq) have \\t. P *\<^sub>V t = x\ using that by blast then obtain t where t_def: \P *\<^sub>V t = x\ - by blast + 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) + 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 + 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 + 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) + by (simp add: v1 dual_order.antisym) have p1: \closed {(0::'a)}\ - by simp + 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) + by (simp add: id_cblinfun.rep_eq minus_cblinfun.rep_eq) hence \bounded_clinear (cblinfun_apply (id_cblinfun - P))\ using cblinfun_apply - by blast + by blast hence \continuous (at x) (cblinfun_apply (id_cblinfun - P))\ by (simp add: clinear_continuous_at) thus ?thesis using \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by simp qed have i2: \closed ( (\ x. x - P *\<^sub>V x) -` {0} )\ using p1 p2 by (rule Abstract_Topology.continuous_closed_vimage) have \closed (range (cblinfun_apply P))\ using i1 i2 by simp have u2: \cblinfun_apply P x \ space_as_set M\ for x by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq top_ccsubspace.rep_eq) have xy: \\ x - P *\<^sub>V x, y \ = 0\ if y1: \y \ space_as_set M\ for x y proof- have \\t. y = P *\<^sub>V t\ using y1 - by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq image_iff + by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq image_iff top_ccsubspace.rep_eq) then obtain t where t_def: \y = P *\<^sub>V t\ by blast have \\ x - P *\<^sub>V x, y \ = \ x - P *\<^sub>V x, P *\<^sub>V t \\ by (simp add: t_def) also have \\ = \ P *\<^sub>V (x - P *\<^sub>V x), t \\ by (metis \P = P*\ cinner_adj_left) also have \\ = \ P *\<^sub>V x - P *\<^sub>V (P *\<^sub>V x), t \\ by (simp add: cblinfun.diff_right) also have \\ = \ P *\<^sub>V x - P *\<^sub>V x, t \\ - by (metis assms(1) comp_apply cblinfun_compose.rep_eq) + by (metis assms(1) comp_apply cblinfun_compose.rep_eq) also have \\ = \ 0, t \\ by simp also have \\ = 0\ by simp finally show ?thesis by blast qed - hence u1: \x - P *\<^sub>V x \ orthogonal_complement (space_as_set M)\ + hence u1: \x - P *\<^sub>V x \ orthogonal_complement (space_as_set M)\ for x - by (simp add: orthogonal_complementI) + 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) + 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: id_cblinfun.rep_eq minus_cblinfun.rep_eq) + 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, hide_lams) M_def) + by (metis (no_types, opaque_lifting) M_def) qed lemma Proj_range_closed: assumes "is_Proj P" shows "closed (range (cblinfun_apply P))" - using assms apply transfer + using assms apply transfer using closed_csubspace.closed is_projection_on_image by blast lemma Proj_is_Proj[simp]: fixes M::\'a::chilbert_space ccsubspace\ shows \is_Proj (Proj M)\ proof- have u1: "closed_csubspace (space_as_set M)" using space_as_set by blast have v1: "h - Proj M *\<^sub>V h \ orthogonal_complement (space_as_set M)" for h by (simp add: Proj.rep_eq orthogonal_complementI projection_orthogonal u1) have v2: "Proj M *\<^sub>V h \ space_as_set M" for h by (metis Proj.rep_eq mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) have u2: "is_projection_on ((*\<^sub>V) (Proj M)) (space_as_set M)" unfolding is_projection_on_def by (simp add: smallest_dist_is_ortho u1 v1 v2) show ?thesis - using u1 u2 is_Proj.rep_eq by blast + using u1 u2 is_Proj.rep_eq by blast qed -lemma is_Proj_algebraic: +lemma is_Proj_algebraic: fixes P::\'a::chilbert_space \\<^sub>C\<^sub>L 'a\ shows \is_Proj P \ P o\<^sub>C\<^sub>L P = P \ P = P*\ proof have "P o\<^sub>C\<^sub>L P = P" if "is_Proj P" using that apply transfer using is_projection_on_idem by fastforce moreover have "P = P*" if "is_Proj P" using that apply transfer by (metis is_projection_on_cadjoint) ultimately show "P o\<^sub>C\<^sub>L P = P \ P = P*" if "is_Proj P" using that by blast show "is_Proj P" if "P o\<^sub>C\<^sub>L P = P \ P = P*" using that Proj_on_own_range' Proj_is_Proj by metis qed lemma Proj_on_own_range: fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \is_Proj P\ shows \Proj (P *\<^sub>S top) = P\ using Proj_on_own_range' assms is_Proj_algebraic by blast lemma Proj_image_leq: "(Proj S) *\<^sub>S A \ S" by (metis Proj_range inf_top_left le_inf_iff isometry_cblinfun_image_inf_distrib') lemma Proj_sandwich: fixes A::"'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space" assumes "isometry A" - shows "sandwich A *\<^sub>V Proj S = Proj (A *\<^sub>S S)" + 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 + unfolding P_def by (metis adj_Proj adj_cblinfun_compose cblinfun_assoc_left(1) double_adj) - ultimately have + 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" + have f1: "A o\<^sub>C\<^sub>L Proj S = P o\<^sub>C\<^sub>L A" by (simp add: P_def assms cblinfun_compose_assoc) hence "P o\<^sub>C\<^sub>L A o\<^sub>C\<^sub>L A* = P" using P_def by presburger hence "(P o\<^sub>C\<^sub>L A) *\<^sub>S (c \ A* *\<^sub>S d) = P *\<^sub>S (A *\<^sub>S c \ d)" for c d by (simp add: cblinfun_assoc_left(2)) hence "P *\<^sub>S (A *\<^sub>S \ \ c) = (P o\<^sub>C\<^sub>L A) *\<^sub>S \" for c by (metis sup_top_left) hence \M = A *\<^sub>S S\ using f1 by (metis \P = Proj M\ cblinfun_assoc_left(2) Proj_range sup_top_right) thus ?thesis using \P = Proj M\ unfolding P_def sandwich_apply by blast qed lemma Proj_orthog_ccspan_union: assumes "\x y. x \ X \ y \ Y \ is_orthogonal x y" shows \Proj (ccspan (X \ Y)) = Proj (ccspan X) + Proj (ccspan Y)\ proof - have \x \ cspan X \ y \ cspan Y \ is_orthogonal x y\ for x y apply (rule is_orthogonal_closure_cspan[where X=X and Y=Y]) using closure_subset assms by auto then have \x \ closure (cspan X) \ y \ closure (cspan Y) \ is_orthogonal x y\ for x y by (metis orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') then show ?thesis apply (transfer fixing: X Y) apply (subst projection_plus[symmetric]) by auto qed abbreviation proj :: "'a::chilbert_space \ 'a \\<^sub>C\<^sub>L 'a" where "proj \ \ Proj (ccspan {\})" lemma proj_0[simp]: \proj 0 = 0\ apply transfer by auto lemma surj_isometry_is_unitary: fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes \isometry U\ assumes \U *\<^sub>S \ = \\ shows \unitary U\ by (metis Proj_sandwich sandwich_apply Proj_on_own_range' assms(1) assms(2) cblinfun_compose_id_right isometry_def unitary_def unitary_id unitary_range) lemma ccsubspace_supI_via_Proj: fixes A B C::"'a::chilbert_space ccsubspace" assumes a1: \Proj (- C) *\<^sub>S A \ B\ shows "A \ sup B C" proof- have x2: \x \ space_as_set B\ if "x \ closure ( (projection (orthogonal_complement (space_as_set C))) ` space_as_set A)" for x using that - by (metis Proj.rep_eq cblinfun_image.rep_eq assms less_eq_ccsubspace.rep_eq subsetD + 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 + 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 + 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 + thus ?thesis by (simp add: x1 less_eq_ccsubspace.rep_eq subset_eq sup_ccsubspace.rep_eq) qed lemma is_Proj_idempotent: assumes "is_Proj P" shows "P o\<^sub>C\<^sub>L P = P" using assms unfolding is_Proj_def using assms is_Proj_algebraic by auto lemma is_proj_selfadj: assumes "is_Proj P" shows "P* = P" using assms unfolding is_Proj_def - by (metis is_Proj_algebraic is_Proj_def) - -lemma is_Proj_I: + by (metis is_Proj_algebraic is_Proj_def) + +lemma is_Proj_I: assumes "P o\<^sub>C\<^sub>L P = P" and "P* = P" shows "is_Proj P" using assms is_Proj_algebraic by metis lemma is_Proj_0[simp]: "is_Proj 0" by (metis add_left_cancel adj_plus bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose group_cancel.rule0 is_Proj_I) -lemma is_Proj_complement[simp]: +lemma is_Proj_complement[simp]: assumes a1: "is_Proj P" shows "is_Proj (id_cblinfun-P)" by (smt (z3) add_diff_cancel_left add_diff_cancel_left' adj_cblinfun_compose adj_plus assms bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel id_cblinfun_adjoint is_Proj_algebraic cblinfun_compose_id_left) lemma Proj_bot[simp]: "Proj bot = 0" - by (metis zero_cblinfun_image Proj_on_own_range' is_Proj_0 is_Proj_algebraic + 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: +lemma Proj_inj: assumes "Proj X = Proj Y" shows "X = Y" by (metis assms Proj_range) subsection \Kernel\ lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector - \ 'a ccsubspace" + \ '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)" + "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]: +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 + thus ?thesis + unfolding eigenspace_def by blast qed lemma eigenspace_memberD: assumes "x \ space_as_set (eigenspace e A)" shows "A *\<^sub>V x = e *\<^sub>C x" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms apply transfer by auto lemma eigenspace_memberI: assumes "A *\<^sub>V x = e *\<^sub>C x" shows "x \ space_as_set (eigenspace e A)" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms apply transfer by auto subsection \Isomorphisms and inverses\ definition iso_cblinfun :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ bool\ where \iso_cblinfun A = (\ B. A o\<^sub>C\<^sub>L B = id_cblinfun \ B o\<^sub>C\<^sub>L A = id_cblinfun)\ definition cblinfun_inv :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ ('b,'a) cblinfun\ where \cblinfun_inv A = (SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun)\ -lemma +lemma assumes \iso_cblinfun A\ shows cblinfun_inv_left: \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and cblinfun_inv_right: \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ proof - from assms obtain B where AB: \A o\<^sub>C\<^sub>L B = id_cblinfun\ and BA: \B o\<^sub>C\<^sub>L A = id_cblinfun\ using iso_cblinfun_def by blast from BA have \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ by (metis (mono_tags, lifting) cblinfun_inv_def someI_ex) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB BA show \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ by auto qed lemma cblinfun_inv_uniq: assumes "A o\<^sub>C\<^sub>L B = id_cblinfun" and "B o\<^sub>C\<^sub>L A = id_cblinfun" shows "cblinfun_inv A = B" using assms by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) subsection \One-dimensional spaces\ instantiation cblinfun :: (one_dim, one_dim) complex_inner begin text \Once we have a theory for the trace, we could instead define the Hilbert-Schmidt inner product and relax the \<^class>\one_dim\-sort constraint to (\<^class>\cfinite_dim\,\<^class>\complex_normed_vector\) or similar\ definition "cinner_cblinfun (A::'a \\<^sub>C\<^sub>L 'b) (B::'a \\<^sub>C\<^sub>L 'b) = cnj (one_dim_iso (A *\<^sub>V 1)) * one_dim_iso (B *\<^sub>V 1)" instance proof intro_classes fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "\A, B\ = cnj \B, A\" unfolding cinner_cblinfun_def by auto show "\A + B, C\ = \A, C\ + \B, C\" - by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) + by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) show "\c *\<^sub>C A, B\ = cnj c * \A, B\" by (simp add: cblinfun.scaleC_left cinner_cblinfun_def) show "0 \ \A, A\" unfolding cinner_cblinfun_def by auto have "bounded_clinear A \ A 1 = 0 \ A = (\_. 0)" for A::"'a \ 'b" proof (rule one_dim_clinear_eqI [where x = 1] , auto) show "clinear A" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: bounded_clinear.clinear) show "clinear ((\_. 0)::'a \ 'b)" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that - by (simp add: complex_vector.module_hom_zero) + 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) + by (metis one_dim_iso_of_zero one_dim_iso_inj) thus "(\A, A\ = 0) = (A = 0)" by (auto simp: cinner_cblinfun_def) show "norm A = sqrt (cmod \A, A\)" - unfolding cinner_cblinfun_def - apply transfer + 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 + 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]: +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 + apply transfer by simp lemma scaleC_1_right[simp]: \scaleC x (1::'a::one_dim) = of_complex x\ unfolding of_complex_def by simp lemma scaleC_of_complex[simp]: \scaleC x (of_complex y) = of_complex (x * y)\ unfolding of_complex_def using scaleC_scaleC by blast lemma scaleC_1_apply[simp]: \(x *\<^sub>C 1) *\<^sub>V y = x *\<^sub>C y\ by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_apply_1_left[simp]: \1 *\<^sub>V y = y\ by (metis cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma of_complex_cblinfun_apply[simp]: \of_complex x *\<^sub>V y = x *\<^sub>C y\ unfolding of_complex_def by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_compose_1_left[simp]: \1 o\<^sub>C\<^sub>L x = x\ apply transfer by auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ apply transfer by auto lemma one_dim_iso_id_cblinfun: \one_dim_iso id_cblinfun = id_cblinfun\ by simp lemma one_dim_iso_id_cblinfun_eq_1: \one_dim_iso id_cblinfun = 1\ by simp lemma one_dim_iso_comp_distr[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a o\<^sub>C\<^sub>L one_dim_iso b\ by (smt (z3) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC) lemma one_dim_iso_comp_distr_times[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a * one_dim_iso b\ by (smt (verit, del_insts) mult.left_neutral mult_scaleC_left one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left) lemma one_dim_iso_adjoint[simp]: \one_dim_iso (A*) = (one_dim_iso A)*\ by (smt (z3) one_cblinfun_adj one_cinner_a_scaleC_one one_dim_iso_of_one one_dim_iso_scaleC scaleC_adj) lemma one_dim_iso_adjoint_complex[simp]: \one_dim_iso (A*) = cnj (one_dim_iso A)\ by (metis (mono_tags, lifting) one_cblinfun_adj one_dim_iso_idem one_dim_scaleC_1 scaleC_adj) lemma one_dim_cblinfun_compose_commute: \a o\<^sub>C\<^sub>L b = b o\<^sub>C\<^sub>L a\ for a b :: \('a::one_dim,'a) cblinfun\ by (simp add: one_dim_iso_inj) lemma one_cblinfun_apply_one[simp]: \1 *\<^sub>V 1 = 1\ by (simp add: one_cblinfun.rep_eq) subsection \Loewner order\ lift_definition heterogenous_cblinfun_id :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ is \if bounded_clinear (heterogenous_identity :: 'a::complex_normed_vector \ 'b::complex_normed_vector) then heterogenous_identity else (\_. 0)\ by auto lemma heterogenous_cblinfun_id_def'[simp]: "heterogenous_cblinfun_id = id_cblinfun" apply transfer by auto definition "heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'b::chilbert_space itself) \ unitary (heterogenous_cblinfun_id :: 'a \\<^sub>C\<^sub>L 'b) \ unitary (heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a)" lemma heterogenous_same_type_cblinfun[simp]: \heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'a::chilbert_space itself)\ unfolding heterogenous_same_type_cblinfun_def by auto instantiation cblinfun :: (chilbert_space, chilbert_space) ord begin definition less_eq_cblinfun :: \('a \\<^sub>C\<^sub>L 'b) \ ('a \\<^sub>C\<^sub>L 'b) \ bool\ - where less_eq_cblinfun_def_heterogenous: \less_eq_cblinfun A B = + 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 + unfolding less_eq_cblinfun_def_heterogenous by (auto simp del: less_eq_complex_def simp: cblinfun.diff_left cinner_diff_right) instantiation cblinfun :: (chilbert_space, chilbert_space) ordered_complex_vector begin instance proof intro_classes note less_eq_complex_def[simp del] fix x y z :: \'a \\<^sub>C\<^sub>L 'b\ fix a b :: complex define pos where \pos X \ (\\. cinner \ (X *\<^sub>V \) \ 0)\ for X :: \'b \\<^sub>C\<^sub>L 'b\ consider (unitary) \heterogenous_same_type_cblinfun TYPE('a) TYPE('b)\ \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B = pos ((B-A) o\<^sub>C\<^sub>L (heterogenous_cblinfun_id :: 'b\\<^sub>C\<^sub>L'a))\ | (trivial) \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B \ A = B\ apply atomize_elim by (auto simp: pos_def less_eq_cblinfun_def_heterogenous) note cases = this have [simp]: \pos 0\ unfolding pos_def by auto have pos_nondeg: \X = 0\ if \pos X\ and \pos (-X)\ for X apply (rule cblinfun_cinner_eqI, simp) using that by (metis (no_types, lifting) cblinfun.minus_left cinner_minus_right dual_order.antisym equation_minus_iff neg_le_0_iff_le pos_def) have pos_add: \pos (X+Y)\ if \pos X\ and \pos Y\ for X Y by (smt (z3) pos_def cblinfun.diff_left cinner_minus_right cinner_simps(3) diff_ge_0_iff_ge diff_minus_eq_add neg_le_0_iff_le order_trans that(1) that(2) uminus_cblinfun.rep_eq) have pos_scaleC: \pos (a *\<^sub>C X)\ if \a\0\ and \pos X\ for X a using that unfolding pos_def by (auto simp: cblinfun.scaleC_left) let ?id = \heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a\ show \x \ x\ apply (cases rule:cases) by auto show \(x < y) \ (x \ y \ \ y \ x)\ unfolding less_cblinfun_def by simp show \x \ z\ if \x \ y\ and \y \ z\ proof (cases rule:cases) case unitary define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ and \b = (z-y) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ with unitary that have \pos a\ and \pos b\ by auto then have \pos (a + b)\ by (rule pos_add) moreover have \a + b = (z - x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ unfolding a_def b_def by (metis (no_types, lifting) bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel ordered_field_class.sign_simps(2) ordered_field_class.sign_simps(8)) - ultimately show ?thesis + 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 + with that show ?thesis by auto next case trivial - with that show ?thesis + 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) + by (metis cblinfun_compose_scaleC_left complex_vector.scale_right_diff_distrib) next case trivial - with that show ?thesis + 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 + 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 del: less_eq_complex_def simp: reals_zero_comparable_iff less_eq_cblinfun_def) lemma positive_cblinfunI: \A \ 0\ if \\x. cinner x (A *\<^sub>V x) \ 0\ unfolding less_eq_cblinfun_def using that by auto (* Note: this does not require B to be a square operator *) lemma positive_cblinfun_squareI: \A = B* o\<^sub>C\<^sub>L B \ A \ 0\ apply (rule positive_cblinfunI) by (metis cblinfun_apply_cblinfun_compose cinner_adj_right cinner_ge_zero) lemma one_dim_loewner_order: \A \ B \ one_dim_iso A \ (one_dim_iso B :: complex)\ for A B :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ proof - note less_eq_complex_def[simp del] have A: \A = one_dim_iso A *\<^sub>C id_cblinfun\ by simp have B: \B = one_dim_iso B *\<^sub>C id_cblinfun\ by simp have \A \ B \ (\\. cinner \ (A \) \ cinner \ (B \))\ by (simp add: less_eq_cblinfun_def) also have \\ \ (\\::'a. one_dim_iso B * (\ \\<^sub>C \) \ one_dim_iso A * (\ \\<^sub>C \))\ apply (subst A, subst B) - by (metis (no_types, hide_lams) cinner_scaleC_right id_cblinfun_apply scaleC_cblinfun.rep_eq) + by (metis (no_types, opaque_lifting) cinner_scaleC_right id_cblinfun_apply scaleC_cblinfun.rep_eq) also have \\ \ one_dim_iso A \ (one_dim_iso B :: complex)\ by (auto intro!: mult_right_mono elim!: allE[where x=1]) finally show ?thesis by - qed lemma one_dim_positive: \A \ 0 \ one_dim_iso A \ (0::complex)\ for A :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ using one_dim_loewner_order[where B=0] by auto subsection \Embedding vectors to operators\ lift_definition vector_to_cblinfun :: \'a::complex_normed_vector \ 'b::one_dim \\<^sub>C\<^sub>L 'a\ is \\\ \. one_dim_iso \ *\<^sub>C \\ by (simp add: bounded_clinear_scaleC_const) -lemma vector_to_cblinfun_cblinfun_apply: - "vector_to_cblinfun (A *\<^sub>V \) = A o\<^sub>C\<^sub>L (vector_to_cblinfun \)" - apply transfer +lemma vector_to_cblinfun_cblinfun_apply: + "vector_to_cblinfun (A *\<^sub>V \) = A o\<^sub>C\<^sub>L (vector_to_cblinfun \)" + apply transfer unfolding comp_def bounded_clinear_def clinear_def Vector_Spaces.linear_def module_hom_def module_hom_axioms_def by simp lemma vector_to_cblinfun_add: \vector_to_cblinfun (x + y) = vector_to_cblinfun x + vector_to_cblinfun y\ apply transfer by (simp add: scaleC_add_right) lemma norm_vector_to_cblinfun[simp]: "norm (vector_to_cblinfun x) = norm x" proof transfer have "bounded_clinear (one_dim_iso::'a \ complex)" - by simp + by simp moreover have "onorm (one_dim_iso::'a \ complex) * norm x = norm x" for x :: 'b by simp ultimately show "onorm (\\. one_dim_iso (\::'a) *\<^sub>C x) = norm x" for x :: 'b by (subst onorm_scaleC_left) qed lemma bounded_clinear_vector_to_cblinfun[bounded_clinear]: "bounded_clinear vector_to_cblinfun" apply (rule bounded_clinearI[where K=1]) apply (transfer, simp add: scaleC_add_right) apply (transfer, simp add: mult.commute) by simp lemma vector_to_cblinfun_scaleC[simp]: "vector_to_cblinfun (a *\<^sub>C \) = a *\<^sub>C vector_to_cblinfun \" for a::complex proof (subst asm_rl [of "a *\<^sub>C \ = (a *\<^sub>C id_cblinfun) *\<^sub>V \"]) show "a *\<^sub>C \ = a *\<^sub>C id_cblinfun *\<^sub>V \" by (simp add: scaleC_cblinfun.rep_eq) show "vector_to_cblinfun (a *\<^sub>C id_cblinfun *\<^sub>V \) = a *\<^sub>C (vector_to_cblinfun \::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun_id_cblinfun_apply cblinfun_compose_scaleC_left vector_to_cblinfun_cblinfun_apply) qed lemma vector_to_cblinfun_apply_one_dim[simp]: shows "vector_to_cblinfun \ *\<^sub>V \ = one_dim_iso \ *\<^sub>C \" apply transfer by (rule refl) lemma vector_to_cblinfun_adj_apply[simp]: shows "vector_to_cblinfun \* *\<^sub>V \ = of_complex (cinner \ \)" - by (simp add: cinner_adj_right one_dim_iso_def one_dim_iso_inj) - -lemma vector_to_cblinfun_comp_one[simp]: - "(vector_to_cblinfun s :: 'a::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L 1 + by (simp add: cinner_adj_right one_dim_iso_def one_dim_iso_inj) + +lemma vector_to_cblinfun_comp_one[simp]: + "(vector_to_cblinfun s :: 'a::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L 1 = (vector_to_cblinfun s :: 'b::one_dim \\<^sub>C\<^sub>L _)" apply (transfer fixing: s) by fastforce lemma vector_to_cblinfun_0[simp]: "vector_to_cblinfun 0 = 0" by (metis cblinfun.zero_left cblinfun_compose_zero_left vector_to_cblinfun_cblinfun_apply) lemma image_vector_to_cblinfun[simp]: "vector_to_cblinfun x *\<^sub>S top = ccspan {x}" proof transfer show "closure (range (\\::'b. one_dim_iso \ *\<^sub>C x)) = closure (cspan {x})" for x :: 'a proof (rule arg_cong [where f = closure]) have "k *\<^sub>C x \ range (\\. one_dim_iso \ *\<^sub>C x)" for k by (smt (z3) id_apply one_dim_iso_id one_dim_iso_idem range_eqI) thus "range (\\. one_dim_iso (\::'b) *\<^sub>C x) = cspan {x}" unfolding complex_vector.span_singleton by auto qed qed lemma vector_to_cblinfun_adj_comp_vector_to_cblinfun[simp]: shows "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \ = cinner \ \ *\<^sub>C id_cblinfun" proof - have "one_dim_iso \ *\<^sub>C one_dim_iso (of_complex \\, \\) = \\, \\ *\<^sub>C one_dim_iso \" - for \ :: "'c::one_dim" + 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 \)" + = 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 \)" + hence "((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by (rule one_dim_iso_inj) thus ?thesis using cblinfun_eqI[where x = "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \" and y = "\\, \\ *\<^sub>C id_cblinfun"] by auto qed lemma isometry_vector_to_cblinfun[simp]: assumes "norm x = 1" shows "isometry (vector_to_cblinfun x)" using assms cnorm_eq_1 isometry_def by force subsection \Butterflies (rank-1 projectors)\ definition butterfly_def: "butterfly (s::'a::complex_normed_vector) (t::'b::chilbert_space) = vector_to_cblinfun s o\<^sub>C\<^sub>L (vector_to_cblinfun t :: complex \\<^sub>C\<^sub>L _)*" abbreviation "selfbutter s \ butterfly s s" lemma butterfly_add_left: \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (simp add: butterfly_def vector_to_cblinfun_add cbilinear_add_left bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma butterfly_add_right: \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (simp add: butterfly_def adj_plus vector_to_cblinfun_add cblinfun_compose_add_right) lemma butterfly_def_one_dim: "butterfly s t = (vector_to_cblinfun s :: 'c::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L (vector_to_cblinfun t :: 'c \\<^sub>C\<^sub>L _)*" (is "_ = ?rhs") for s :: "'a::complex_normed_vector" and t :: "'b::chilbert_space" proof - let ?isoAC = "1 :: 'c \\<^sub>C\<^sub>L complex" let ?isoCA = "1 :: complex \\<^sub>C\<^sub>L 'c" let ?vector = "vector_to_cblinfun :: _ \ ('c \\<^sub>C\<^sub>L _)" have "butterfly s t = (?vector s o\<^sub>C\<^sub>L ?isoCA) o\<^sub>C\<^sub>L (?vector t o\<^sub>C\<^sub>L ?isoCA)*" unfolding butterfly_def vector_to_cblinfun_comp_one by simp also have "\ = ?vector s o\<^sub>C\<^sub>L (?isoCA o\<^sub>C\<^sub>L ?isoCA*) o\<^sub>C\<^sub>L (?vector t)*" by (metis (no_types, lifting) cblinfun_compose_assoc adj_cblinfun_compose) also have "\ = ?rhs" by simp finally show ?thesis by simp qed lemma butterfly_comp_cblinfun: "butterfly \ \ o\<^sub>C\<^sub>L a = butterfly \ (a* *\<^sub>V \)" unfolding butterfly_def - by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) + by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma cblinfun_comp_butterfly: "a o\<^sub>C\<^sub>L butterfly \ \ = butterfly (a *\<^sub>V \) \" unfolding butterfly_def - by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) + by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma butterfly_apply[simp]: "butterfly \ \' *\<^sub>V \ = \\', \\ *\<^sub>C \" by (simp add: butterfly_def scaleC_cblinfun.rep_eq) lemma butterfly_scaleC_left[simp]: "butterfly (c *\<^sub>C \) \ = c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleC_right[simp]: "butterfly \ (c *\<^sub>C \) = cnj c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleR_left[simp]: "butterfly (r *\<^sub>R \) \ = r *\<^sub>C butterfly \ \" by (simp add: scaleR_scaleC) lemma butterfly_scaleR_right[simp]: "butterfly \ (r *\<^sub>R \) = r *\<^sub>C butterfly \ \" by (simp add: butterfly_scaleC_right scaleR_scaleC) lemma butterfly_adjoint[simp]: "(butterfly \ \)* = butterfly \ \" unfolding butterfly_def by auto lemma butterfly_comp_butterfly[simp]: "butterfly \1 \2 o\<^sub>C\<^sub>L butterfly \3 \4 = \\2, \3\ *\<^sub>C butterfly \1 \4" by (simp add: butterfly_comp_cblinfun) lemma butterfly_0_left[simp]: "butterfly 0 a = 0" by (simp add: butterfly_def) lemma butterfly_0_right[simp]: "butterfly a 0 = 0" by (simp add: butterfly_def) lemma norm_butterfly: "norm (butterfly \ \) = norm \ * norm \" proof (cases "\=0") case True then show ?thesis by simp next case False - show ?thesis + show ?thesis unfolding norm_cblinfun.rep_eq thm onormI[OF _ False] proof (rule onormI[OF _ False]) - fix x + fix x have "cmod \\, x\ * norm \ \ norm \ * norm \ * norm x" by (metis ab_semigroup_mult_class.mult_ac(1) complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult_left_mono norm_ge_zero) thus "norm (butterfly \ \ *\<^sub>V x) \ norm \ * norm \ * norm x" by (simp add: power2_eq_square) show "norm (butterfly \ \ *\<^sub>V \) = norm \ * norm \ * norm \" by (smt (z3) ab_semigroup_mult_class.mult_ac(1) butterfly_apply mult.commute norm_eq_sqrt_cinner norm_ge_zero norm_scaleC power2_eq_square real_sqrt_abs real_sqrt_eq_iff) qed qed lemma bounded_sesquilinear_butterfly[bounded_sesquilinear]: \bounded_sesquilinear (\(b::'b::chilbert_space) (a::'a::chilbert_space). butterfly a b)\ proof standard fix a a' :: 'a and b b' :: 'b and r :: complex show \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (rule butterfly_add_left) - show \butterfly a (b + b') = butterfly a b + butterfly a b'\ + 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: +lemma inj_selfbutter_upto_phase: assumes "selfbutter x = selfbutter y" shows "\c. cmod c = 1 \ x = c *\<^sub>C y" proof (cases "x = 0") case True from assms have "y = 0" using norm_butterfly by (metis True butterfly_0_left divisors_zero norm_eq_zero) with True show ?thesis using norm_one by fastforce next case False define c where "c = \y, x\ / \x, x\" have "\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x" by (simp add: butterfly_apply) also have "\ = selfbutter y *\<^sub>V x" using assms by simp also have "\ = \y, x\ *\<^sub>C y" by (simp add: butterfly_apply) finally have xcy: "x = c *\<^sub>C y" by (simp add: c_def ceq_vector_fraction_iff) have "cmod c * norm x = cmod c * norm y" using assms norm_butterfly by (smt (verit, ccfv_SIG) \\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x\ \selfbutter y *\<^sub>V x = \y, x\ *\<^sub>C y\ cinner_scaleC_right complex_vector.scale_left_commute complex_vector.scale_right_imp_eq mult_cancel_left norm_eq_sqrt_cinner norm_eq_zero scaleC_scaleC xcy) also have "cmod c * norm y = norm (c *\<^sub>C y)" by simp also have "\ = norm x" unfolding xcy[symmetric] by simp finally have c: "cmod c = 1" by (simp add: False) from c xcy show ?thesis by auto qed lemma butterfly_eq_proj: assumes "norm x = 1" shows "selfbutter x = proj x" proof - define B and \ :: "complex \\<^sub>C\<^sub>L 'a" where "B = selfbutter x" and "\ = vector_to_cblinfun x" then have B: "B = \ o\<^sub>C\<^sub>L \*" unfolding butterfly_def by simp - have \adj\: "\* o\<^sub>C\<^sub>L \ = id_cblinfun" + have \adj\: "\* o\<^sub>C\<^sub>L \ = id_cblinfun" using \_def assms isometry_def isometry_vector_to_cblinfun by blast have "B o\<^sub>C\<^sub>L B = \ o\<^sub>C\<^sub>L (\* o\<^sub>C\<^sub>L \) o\<^sub>C\<^sub>L \*" by (simp add: B cblinfun_assoc_left(1)) also have "\ = B" unfolding \adj\ by (simp add: B) finally have idem: "B o\<^sub>C\<^sub>L B = B". have herm: "B = B*" unfolding B by simp from idem herm have BProj: "B = Proj (B *\<^sub>S top)" by (rule Proj_on_own_range'[symmetric]) have "B *\<^sub>S top = ccspan {x}" by (simp add: B \_def assms cblinfun_compose_image range_adjoint_isometry) with BProj show "B = proj x" by simp qed lemma butterfly_is_Proj: \norm x = 1 \ is_Proj (selfbutter x)\ by (subst butterfly_eq_proj, simp_all) lemma cspan_butterfly_UNIV: assumes \cspan basisA = UNIV\ assumes \cspan basisB = UNIV\ assumes \is_ortho_set basisB\ assumes \\b. b \ basisB \ norm b = 1\ shows \cspan {butterfly a b| (a::'a::{complex_normed_vector}) (b::'b::{chilbert_space,cfinite_dim}). a \ basisA \ b \ basisB} = UNIV\ proof - have F: \\F\{butterfly a b |a b. a \ basisA \ b \ basisB}. \b'\basisB. F *\<^sub>V b' = (if b' = b then a else 0)\ if \a \ basisA\ and \b \ basisB\ for a b apply (rule bexI[where x=\butterfly a b\]) using assms that by (auto simp: is_ortho_set_def cnorm_eq_1) show ?thesis apply (rule cblinfun_cspan_UNIV[where basisA=basisB and basisB=basisA]) using assms apply auto[2] using F by (smt (verit, ccfv_SIG) image_iff) qed -lemma cindependent_butterfly: +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 + 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 + 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 + 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 + unfolding g using normA normB by (auto simp: butterfly_def scaleC_cblinfun.rep_eq cnorm_eq_1) finally show \f g = 0\ by simp qed lemma clinear_eq_butterflyI: fixes F G :: \('a::{chilbert_space,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_inner) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes \cspan basisA = UNIV\ \cspan basisB = UNIV\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes "\a b. a\basisA \ b\basisB \ F (butterfly a b) = G (butterfly a b)" assumes \\b. b\basisB \ norm b = 1\ shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_UNIV) using assms by auto subsection \Bifunctionals\ lift_definition bifunctional :: \'a::complex_normed_vector \\<^sub>C\<^sub>L (('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex)\ is \\x f. f *\<^sub>V x\ by (simp add: cblinfun.flip) lemma bifunctional_apply[simp]: \(bifunctional *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ by (transfer fixing: x f, simp) lemma bifunctional_isometric[simp]: \norm (bifunctional *\<^sub>V x) = norm x\ for x :: \'a::complex_inner\ proof - define f :: \'a \\<^sub>C\<^sub>L complex\ where \f = CBlinfun (\y. cinner x y)\ then have [simp]: \f *\<^sub>V y = cinner x y\ for y by (simp add: bounded_clinear_CBlinfun_apply bounded_clinear_cinner_right) then have [simp]: \norm f = norm x\ apply (auto intro!: norm_cblinfun_eqI[where x=x] simp: power2_norm_eq_cinner[symmetric]) apply (smt (verit, best) norm_eq_sqrt_cinner norm_ge_zero power2_norm_eq_cinner real_div_sqrt) using Cauchy_Schwarz_ineq2 by blast show ?thesis apply (auto intro!: norm_cblinfun_eqI[where x=f]) apply (metis norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) by (metis norm_cblinfun ordered_field_class.sign_simps(33)) qed lemma norm_bifunctional[simp]: \norm (bifunctional :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ proof - obtain x :: 'a where [simp]: \norm x = 1\ by (meson UNIV_not_singleton ex_norm1) show ?thesis by (auto intro!: norm_cblinfun_eqI[where x=x]) qed subsection \Banach-Steinhaus\ theorem cbanach_steinhaus: fixes F :: \'c \ 'a::cbanach \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\x. \M. \n. norm ((F n) *\<^sub>V x) \ M\ - shows \\M. \ n. norm (F n) \ M\ + shows \\M. \ n. norm (F n) \ M\ using cblinfun_blinfun_transfer[transfer_rule] apply (rule TrueI)? (* Deletes current facts *) proof (use assms in transfer) fix F :: \'c \ 'a \\<^sub>L 'b\ assume \(\x. \M. \n. norm (F n *\<^sub>v x) \ M)\ hence \\x. bounded (range (\n. blinfun_apply (F n) x))\ by (metis (no_types, lifting) boundedI rangeE) hence \bounded (range F)\ by (simp add: banach_steinhaus) thus \\M. \n. norm (F n) \ M\ by (simp add: bounded_iff) qed subsection \Riesz-representation theorem\ theorem riesz_frechet_representation_cblinfun_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ shows \\t. \x. f *\<^sub>V x = \t, x\\ apply transfer by (rule riesz_frechet_representation_existence) lemma riesz_frechet_representation_cblinfun_unique: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::complex_inner \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ assumes \\x. f *\<^sub>V x = \u, x\\ shows \t = u\ using assms by (rule riesz_frechet_representation_unique) theorem riesz_frechet_representation_cblinfun_norm: includes notation_norm fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ shows \\f\ = \t\\ - using assms + using assms proof transfer fix f::\'a \ complex\ and t - assume \bounded_clinear f\ and \\x. f x = \t, x\\ - from \\x. f x = \t, x\\ + assume \bounded_clinear f\ and \\x. f x = \t, x\\ + from \\x. f x = \t, x\\ have \(norm (f x)) / (norm x) \ norm t\ for x proof(cases \norm x = 0\) case True thus ?thesis by simp next case False have \norm (f x) = norm (\t, x\)\ using \\x. f x = \t, x\\ by simp also have \norm \t, x\ \ norm t * norm x\ by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) finally have \norm (f x) \ norm t * norm x\ by blast thus ?thesis - by (metis False linordered_field_class.divide_right_mono nonzero_mult_div_cancel_right norm_ge_zero) + 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 + by simp next case False have \f t = \t, t\\ using \\x. f x = \t, x\\ by blast also have \\ = (norm t)^2\ by (meson cnorm_eq_square) also have \\ = (norm t)*(norm t)\ by (simp add: power2_eq_square) finally have \f t = (norm t)*(norm t)\ by blast thus ?thesis by (metis False Re_complex_of_real \\x. f x = cinner t x\ cinner_ge_zero complex_of_real_cmod nonzero_divide_eq_eq) qed ultimately have \Sup {(norm (f x)) / (norm x)| x. True} = norm t\ - by (smt cSup_eq_maximum mem_Collect_eq) + 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) + by (simp add: full_SetCompr_eq) ultimately show \onorm f = norm t\ - by (simp add: onorm_def) + by (simp add: onorm_def) qed subsection \Extension of complex bounded operators\ -definition cblinfun_extension where +definition cblinfun_extension where "cblinfun_extension S \ = (SOME B. \x\S. B *\<^sub>V x = \ x)" -definition cblinfun_extension_exists where +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" + 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) + 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" + 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) + by (simp add: complex_vector.construct_basis) finally show?thesis by blast qed - thus ?thesis + thus ?thesis unfolding cblinfun_extension_exists_def by blast qed lemma cblinfun_extension_exists_bounded_dense: fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes \closure S = UNIV\ assumes f_add: \\x y. x \ S \ y \ S \ f (x + y) = f x + f y\ assumes f_scale: \\c x y. x \ S \ f (c *\<^sub>C x) = c *\<^sub>C f x\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows \cblinfun_extension_exists S f\ proof - obtain B where bounded: \\x. x \ S \ norm (f x) \ B * norm x\ and \B > 0\ - using bounded by (smt (z3) mult_mono norm_ge_zero) + using bounded by (smt (z3) mult_mono norm_ge_zero) have \\xi. (xi \ x) \ (\i. xi i \ S)\ for x using assms(2) closure_sequential by blast then obtain seq :: \'a \ nat \ 'a\ where seq_lim: \seq x \ x\ and seq_S: \seq x i \ S\ for x i apply (atomize_elim, subst all_conj_distrib[symmetric]) apply (rule choice) by auto define g where \g x = lim (\i. f (seq x i))\ for x have \Cauchy (\i. f (seq x i))\ for x proof (rule CauchyI) fix e :: real assume \e > 0\ have \Cauchy (seq x)\ using LIMSEQ_imp_Cauchy seq_lim by blast then obtain M where less_eB: \norm (seq x m - seq x n) < e/B\ if \n \ M\ and \m \ M\ for n m apply atomize_elim by (meson CauchyD \0 < B\ \0 < e\ linordered_field_class.divide_pos_pos) have \norm (f (seq x m) - f (seq x n)) < e\ if \n \ M\ and \m \ M\ for n m proof - have \norm (f (seq x m) - f (seq x n)) = norm (f (seq x m - seq x n))\ using f_add f_scale seq_S - by (metis add_diff_cancel assms(1) complex_vector.subspace_diff diff_add_cancel) + 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 + using bounded using \xi _ \ S\ by blast qed have \bounded_clinear g\ using g_add g_scale apply (rule bounded_clinearI[where K=B]) using g_bounded by (simp add: ordered_field_class.sign_simps(5)) then have [simp]: \CBlinfun g *\<^sub>V x = g x\ for x by (subst CBlinfun_inverse, auto) show \cblinfun_extension_exists S f\ apply (rule cblinfun_extension_existsI[where B=\CBlinfun g\]) by auto qed lemma cblinfun_extension_apply: assumes "cblinfun_extension_exists S f" and "v \ S" shows "(cblinfun_extension S f) *\<^sub>V v = f v" by (smt assms cblinfun_extension_def cblinfun_extension_exists_def tfl_some) subsection \Notation\ bundle cblinfun_notation begin notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) notation cblinfun_apply (infixr "*\<^sub>V" 70) notation cblinfun_image (infixr "*\<^sub>S" 70) notation adj ("_*" [99] 100) end bundle no_cblinfun_notation begin no_notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) no_notation cblinfun_apply (infixr "*\<^sub>V" 70) no_notation cblinfun_image (infixr "*\<^sub>S" 70) no_notation adj ("_*" [99] 100) end bundle blinfun_notation begin notation blinfun_apply (infixr "*\<^sub>V" 70) end bundle no_blinfun_notation begin no_notation blinfun_apply (infixr "*\<^sub>V" 70) end unbundle no_cblinfun_notation 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,2230 +1,2230 @@ (* 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 + imports Complex_Vector_Spaces - "HOL-Analysis.Infinite_Set_Sum" + "HOL-Analysis.Infinite_Set_Sum" Complex_Inner_Product0 begin subsection \Complex inner product spaces\ bundle cinner_bracket_notation begin notation cinner ("\_, _\") end unbundle cinner_bracket_notation bundle no_cinner_bracket_notation begin no_notation cinner ("\_, _\") end lemma cinner_real: "cinner x x \ \" by (meson cinner_ge_zero reals_zero_comparable_iff) 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 , 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\ text \This is a useful rule for establishing the equality of vectors\ lemma cinner_extensionality: assumes \\\. \\, \\ = \\, \\\ 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, y\\ \ \Shown in the proof of Corollary 1.5 in @{cite conway2013course}\ proof - have \\x , y\ + \y , x\ = \x , y\ + cnj \x , y\\ by simp hence \\x , y\ + \y , x\ = 2 * Re \x , y\ \ using complex_add_cnj by presburger have \\x + y\^2 = \x+y, x+y\\ - by (simp add: cdot_square_norm) + by (simp add: cdot_square_norm) hence \\x + y\^2 = \x , x\ + \x , y\ + \y , x\ + \y , y\\ by (simp add: cinner_add_left cinner_add_right) thus ?thesis using \\x , y\ + \y , x\ = 2 * Re \x , 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 + includes notation_norm shows \\x - y\^2 = \x\^2 + \y\^2 - 2 * Re \x, y\\ proof- have \\x + (-y)\^2 = \x\^2 + \-y\^2 + 2 * Re \x , (-y)\\ using polar_identity by blast hence \\x - y\^2 = \x\^2 + \y\^2 - 2*Re \x , y\\ by simp - thus ?thesis + 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}\ + \ \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 , y\ = 0 \ \ x + y \^2 = \ x \^2 + \ y \^2\ - \ \Shown in the proof of Theorem 2.2 in @{cite conway2013course}\ + shows \\x , 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, 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 + by auto next case (insert x F) have r1: "\f x, f a\ = 0" if "a \ F" for a - using that insert.hyps(2) insert.prems by auto + using that insert.hyps(2) insert.prems by auto have "sum f F = (\a\F. f a)" by simp hence s4: "\f x, sum f F\ = \f x, (\a\F. f a)\" by simp also have s3: "\ = (\a\F. \f x, 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, sum f F\ = 0" - using s2 s3 by auto + 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, 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) + 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) + 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) + 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) + by (smt M2 norm_not_less_zero) have \\N. \n \ N. \m \ N. norm ( (\ i. \ x i, y i \) n - (\ i. \ x i, 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) + 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) + 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, y n \ - \ x m, y m \ ) < e\ if \n \ N\ and \m \ N\ for n m proof - have \\ x n, y n \ - \ x m, y m \ = (\ x n, y n \ - \ x m, y n \) + (\ x m, y n \ - \ x m, y m \)\ by simp hence y1: \norm (\ x n, y n \ - \ x m, y m \) \ norm (\ x n, y n \ - \ x m, y n \) + norm (\ x m, y n \ - \ x m, y m \)\ by (metis norm_triangle_ineq) have \\ x n, y n \ - \ x m, y n \ = \ x n - x m, y n \\ by (simp add: cinner_diff_left) hence \norm (\ x n, y n \ - \ x m, y n \) = norm \ x n - x m, y n \\ by simp moreover have \norm \ x n - x m, 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) + 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, y n \ - \ x m, 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, y n \ - \ x m, y n \) < e/2\ - by simp + by simp hence y2: \norm (\ x n, y n \ - \ x m, y n \) < e/2\ - by blast + by blast have \\ x m, y n \ - \ x m, y m \ = \ x m, y n - y m \\ by (simp add: cinner_diff_right) hence \norm (\ x m, y n \ - \ x m, y m \) = norm \ x m, y n - y m \\ by simp moreover have \norm \ x m, 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 + using \N \ m\ \N \ n\ \N2 \ N\ N2_def by auto ultimately have \norm (\ x m, y n \ - \ x m, 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, y n \ - \ x m, y m \) < e/2\ by simp hence y3: \norm (\ x m, y n \ - \ x m, y m \) < e/2\ by blast show \norm ( \ x n, y n \ - \ x m, 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: +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 subsection \Orthogonality\ -definition "orthogonal_complement S = {x| x. \y\S. cinner x y = 0}" +definition "orthogonal_complement S = {x| x. \y\S. cinner x y = 0}" lemma orthogonal_complement_orthoI: \x \ orthogonal_complement M \ y \ M \ \ x, y \ = 0\ unfolding orthogonal_complement_def by auto lemma orthogonal_complement_orthoI': \x \ M \ y \ orthogonal_complement M \ \ x, y \ = 0\ by (metis cinner_commute' complex_cnj_zero orthogonal_complement_orthoI) lemma orthogonal_complementI: \(\x. x \ M \ \ y, 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, 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 orthogonal_complement_closed_subspace[simp]: +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\ + 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. \ y , an n \ = 0\ using orthogonal_complement_orthoI' by (simp add: orthogonal_complement_orthoI' ortho) moreover have \isCont (\ x. \ y , 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 , v \) (an n)) \ (\ v. \ y , v \) a\ for y using isCont_tendsto_compose by (simp add: isCont_tendsto_compose lim) hence \\ y\A. (\ n. \ y , an n \ ) \ \ y , a \\ by simp - hence \\ y\A. (\ n. 0 ) \ \ y , a \\ - using \\ y \ A. \ n. \ y , an n \ = 0\ + hence \\ y\A. (\ n. 0 ) \ \ y , a \\ + using \\ y \ A. \ n. \ y , an n \ = 0\ by fastforce - hence \\ y \ A. \ y , a \ = 0\ + hence \\ y \ A. \ y , a \ = 0\ 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 "\ x, x \ = 0" 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 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 text \Orthogonal set\ definition is_ortho_set :: "'a::complex_inner set \ bool" where \is_ortho_set S = ((\x\S. \y\S. x \ y \ \x, y\ = 0) \ 0 \ S)\ 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: \\ y, x \ = 0\ + have s1: \\ y, x \ = 0\ if a1: "x \ (orthogonal_complement A)" - and a2: \y \ closure A\ + and a2: \y \ closure A\ for x y proof- have \\ y \ A. \ y , x \ = 0\ 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 , x \) y\ by simp hence \(\ n. \ yy n , x \) \ \ y , x \\ using \yy \ y\ isCont_tendsto_compose by fastforce hence \(\ n. 0) \ \ y , x \\ using \\ y \ A. \ y , x \ = 0\ \\ n. yy n \ A\ by simp - thus ?thesis + 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)\ + 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: +lemma is_orthogonal_closure: assumes \\s. s \ S \ is_orthogonal a s\ - assumes \x \ closure 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 "\a, x\ = 0" 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: "\a, i\ = 0" if "i\t" for i using b2 a1 that by blast have "\a, x\ = \a, (\i\t. r i *\<^sub>C i)\" - by (simp add: b3) + by (simp add: b3) also have "\ = (\i\t. r i *\<^sub>C \a, 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) + 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). \ x, y \ = 0\ 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" + 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 "\v, v'\ = 0" 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, v'\) = 0" by simp have "\v, (\v'\t. u v' *\<^sub>C v')\ = (\v'\t. u v' * \v, v'\)" using b1 - by (metis (mono_tags, lifting) cinner_scaleC_right cinner_sum_right sum.cong) + by (metis (mono_tags, lifting) cinner_scaleC_right cinner_sum_right sum.cong) also have "\ = u v * \v, v\ + (\v'\t-{v}. u v' * \v, v'\)" by (meson b1 b4 sum.remove) also have "\ = u v * \v, v\" using sum0 by simp finally have "\v, (\v'\t. u v' *\<^sub>C v')\ = u v * \v, v\" by blast hence "u v * \v, v\ = 0" using b3 by simp moreover have "\v, v\ \ 0" - using assms is_ortho_set_def b2 b4 by auto + 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, 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) + 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, 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, a\)\ - using norm_eq_sqrt_cinner by auto + using norm_eq_sqrt_cinner by auto ultimately have \sqrt (norm \a, a\) = 1\ by simp hence \norm \a, a\ = 1\ using real_sqrt_eq_1_iff by blast moreover have \\a, a\ \ \\ - by (simp add: cinner_real) + by (simp add: cinner_real) moreover have \\a, a\ \ 0\ using cinner_ge_zero by blast ultimately have w1: \\a, a\ = 1\ by (metis \0 \ \a, a\\ \cmod \a, a\ = 1\ complex_of_real_cmod of_real_1) have \r t * \a, 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, t\) = 0\ - by (simp add: \\t. t \ T - {a} \ r t * \a, t\ = 0\) + by (simp add: \\t. t \ T - {a} \ r t * \a, t\ = 0\) have \\a, x\ = \a, (\ t\T. r t *\<^sub>C t)\\ using \x = (\ a\T. r a *\<^sub>C a)\ by simp also have \\ = (\ t\T. \a, r t *\<^sub>C t\)\ using cinner_sum_right by blast also have \\ = (\ t\T. r t * \a, t\)\ - by simp + by simp also have \\ = r a * \a, a\ + (\ t\T-{a}. r t * \a, t\)\ using \a \ T\ by (meson \finite T\ sum.remove) also have \\ = r a * \a, a\\ using s1 by simp also have \\ = r a\ by (simp add: w1) finally show ?thesis by auto qed - thus ?thesis + thus ?thesis using \x = (\ a\T. r a *\<^sub>C a)\ - by fastforce + by fastforce qed subsection \Projections\ lemma smallest_norm_exists: \ \Theorem 2.5 in @{cite conway2013course} (inside the proof)\ includes notation_norm fixes M :: \'a::chilbert_space set\ assumes q1: \convex M\ and q2: \closed M\ and q3: \M \ {}\ shows \\k. is_arg_min (\ x. \x\) (\ t. t \ M) k\ proof- - define d where \d = Inf { \x\^2 | x. x \ M }\ + 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\ + by fastforce + have \d \ \x\^2\ if a1: "x \ M" for x proof- have "\v. (\w. Re (\v , v\ ) = \w\\<^sup>2 \ w \ M) \ v \ M" by (metis (no_types) power2_norm_eq_cinner') hence "Re (\x , 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 + 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\ + 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 + 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\) + 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\ + 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\ + = (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 \... + 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 \... + 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 \... + 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 \... + also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) + 2*d ) - d\ by simp - also have \... + also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + (1/2)*(2*d) - d\ by (simp add: distrib_left) - also have \... + also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + d - d\ by simp - also have \... + 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) + 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 + 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" + 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 + 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" + if f1: "n \ N" and f2: "m \ N" for m n::nat proof- - have \1/(n+1) \ 1/(N+1)\ + have \1/(n+1) \ 1/(N+1)\ by (simp add: f1 linordered_field_class.frac_le) - moreover have \1/(m+1) \ 1/(N+1)\ + 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\ + 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" + 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 + 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\ + 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\ + 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\ + 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\ + 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\ + 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\ + have \r \ M\ using \is_arg_min (\x. \x\) (\ t. t \ M) r\ by (simp add: is_arg_min_def) - moreover have \s \ M\ + 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\\ + 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 + 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 + \ \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 + \ \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 + 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) \ + 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 , 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\ + 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 , f \)\ by (simp add: polar_identity_minus) finally have \\ (h - k) \^2 \ \ (h - k) \^2 + \ f \^2 - 2 * Re (\ h - k , f \)\ by simp thus ?thesis by simp qed have q4: \\ c > 0. 2 * Re (\ h - k , f \) \ c\ if \\c>0. 2 * Re (\h - k , f\ ) \ c * \f\\<^sup>2\ for f proof (cases \\ f \^2 > 0\) case True hence \\ c > 0. 2 * Re (\ h - k , f \) \ (c/\ f \^2)*\ f \^2\ using that linordered_field_class.divide_pos_pos by blast - thus ?thesis + thus ?thesis using True by auto next case False - hence \\ f \^2 = 0\ + hence \\ f \^2 = 0\ by simp - thus ?thesis + thus ?thesis by auto qed - have q3: \\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ 0\ + have q3: \\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ 0\ if a3: \\f. f \ M \ (\c>0. 2 * Re \h - k , 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 , f \) \ c*\ f \^2\ - by (simp add: that ) - thus ?thesis + 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 + using is_arg_min_def that by fastforce hence \\ f. f \ M \ 2 * Re (\ h - k , f \) \ \ f \^2\ using r1 - by (simp add: that) - have \\ f. f \ M \ + by (simp add: that) + have \\ f. f \ M \ (\ c::real. 2 * Re (\ h - k , 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, f\ \ \f\\<^sup>2\) hence \\ f. f \ M \ (\ c::real. c * (2 * Re (\ h - k , 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 , f \)) \ \c\^2*\ f \^2)\ by (simp add: power_mult_distrib) - hence \\ f. f \ M \ + hence \\ f. f \ M \ (\ c::real. c * (2 * Re (\ h - k , f \)) \ c^2*\ f \^2)\ by auto - hence \\ f. f \ M \ + hence \\ f. f \ M \ (\ c::real. c > 0 \ c * (2 * Re (\ h - k , f \)) \ c^2*\ f \^2)\ by simp - hence \\ f. f \ M \ + hence \\ f. f \ M \ (\ c::real. c > 0 \ c*(2 * Re (\ h - k , f \)) \ c*(c*\ f \^2))\ by (simp add: power2_eq_square) - hence q4: \\ f. f \ M \ + hence q4: \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ c*\ f \^2)\ - by simp + by simp have \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ 0)\ using q3 - by (simp add: q4 that) - hence \\ f. f \ M \ + by (simp add: q4 that) + hence \\ f. f \ M \ (\ c::real. c > 0 \ (2 * Re (\ h - k , (-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 , f \)) \ 0)\ by simp - hence \\ f. f \ M \ + hence \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ 0)\ by simp - hence \\ f. f \ M \ + hence \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) = 0)\ - using \\ f. f \ M \ + using \\ f. f \ M \ (\ c::real. c > 0 \ (2 * Re (\ h - k , f \)) \ 0)\ by fastforce - have \\ f. f \ M \ + have \\ f. f \ M \ ((1::real) > 0 \ 2 * Re (\ h - k , f \) = 0)\ using \\f. f \ M \ (\c>0. 2 * Re (\h - k , f\ ) = 0)\ by blast hence \\ f. f \ M \ 2 * Re (\ h - k , f \) = 0\ by simp - hence \\ f. f \ M \ Re (\ h - k , f \) = 0\ - by simp + hence \\ f. f \ M \ Re (\ h - k , f \) = 0\ + by simp have \\ f. f \ M \ Re (\ h - k , (Complex 0 (-1)) *\<^sub>C f \) = 0\ using assms complex_vector.subspace_def \csubspace M\ by (metis \\f. f \ M \ Re \h - k, f\ = 0\) hence \\ f. f \ M \ Re ( (Complex 0 (-1))*(\ h - k , f \) ) = 0\ by simp - hence \\ f. f \ M \ Im (\ h - k , f \) = 0\ - using Complex_eq_neg_1 Re_i_times cinner_scaleC_right complex_of_real_def by auto + hence \\ f. f \ M \ Im (\ h - k , 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 , f \) = 0\ using complex_eq_iff by (simp add: \\f. f \ M \ Im \h - k, f\ = 0\ \\f. f \ M \ Re \h - k, f\ = 0\) hence \h - k \ orthogonal_complement M \ k \ M\ - by (simp add: \k \ M\ orthogonal_complementI) + 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 , f \ = 0\ + have \\ h - k , 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\ + hence \h - k \ orthogonal_complement M\ by (simp add: orthogonal_complement_def) thus ?thesis - using \k \ M\ by auto + using \k \ M\ by auto qed - have q1: \dist h k \ dist h f \ + 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, k - f \ = 0\ - by (metis (no_types, lifting) that + 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, 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 + 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 + by blast thus ?thesis - by (metis (no_types, lifting) dist_commute is_arg_min_linorder q1 that) + by (metis (no_types, lifting) dist_commute is_arg_min_linorder q1 that) qed show ?thesis - using w1 w2 by blast + using w1 w2 by blast qed corollary orthog_proj_exists: - fixes M :: \'a::chilbert_space set\ + 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) + by (simp add: assms smallest_dist_is_ortho) qed corollary orthog_proj_unique: - fixes M :: \'a::complex_inner set\ + 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\ + 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 a1: "closed_csubspace M" and a2: "x \ M" shows "(projection M) x = x" using is_projection_on_fixes_image \ \Theorem 2.7 in @{cite conway2013course}\ by (simp add: a1 a2 complex_vector.subspace_0 projection_eqI) 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 + using assms is_projection_on_iff_orthog by blast hence \\ k \ M. \ h - \ h , k \ = 0\ - using orthogonal_complement_orthoI by blast + 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 \\ h - \ h , \ h \ = 0\ 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 + 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 + 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 + 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 + using a2 by auto hence \\ (x + y) - ( (\ x) + (\ y) ) \ M\ - by (simp add: complex_vector.subspace_add complex_vector.subspace_diff f1) + 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) + 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) + 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 + 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\) + 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 + 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)\ + 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 + 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})" + 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 + ultimately show ?thesis by blast qed -\ \Theorem 2.7 in @{cite conway2013course}\ +\ \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 , x\ / \t , 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 , x\ / \t , 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 + using cc ortho inspan by auto qed lemma projection_rank1: fixes t x :: \'a::complex_inner\ shows \projection (cspan {t}) x = (\t , x\ / \t , 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" + 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\ + 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 + using b2 b3 by blast hence b4: \( id - (projection M) ) -` {0} = M\ by blast - have b1: "orthogonal_complement (orthogonal_complement M) + 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) + by (simp add: b4) finally show ?thesis by blast qed 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_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 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 orthogonal_complement_UNIV[simp]: +lemma orthogonal_complement_UNIV[simp]: "orthogonal_complement UNIV = {0}" by (metis Int_UNIV_left complex_vector.subspace_UNIV complex_vector.subspace_def orthogonal_complement_zero_intersection) lemma orthogonal_complement_zero[simp]: "orthogonal_complement {0} = UNIV" unfolding orthogonal_complement_def by auto -lemma de_morgan_orthogonal_complement_plus: +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)" + 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 + using that by blast hence t1: \\z \ (A + B). \ z , x \ = 0\ - by (simp add: orthogonal_complement_orthoI') + 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 , x \ = 0\ + hence \\z \ A. \ z , 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 + using \0 \ A\ subset_iff set_zero_plus2 by blast hence \\ z \ B. \ z , x \ = 0\ - using t1 by auto + using t1 by auto hence \x \ (orthogonal_complement B)\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) - thus ?thesis + 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)\ + have \x \ (orthogonal_complement A)\ using v1 by blast hence \\y\ A. \ y , x \ = 0\ by (simp add: orthogonal_complement_orthoI') - have \x \ (orthogonal_complement B)\ - using v1 + have \x \ (orthogonal_complement B)\ + using v1 by blast hence \\ y\ B. \ y , x \ = 0\ by (simp add: orthogonal_complement_orthoI') have \\ a\A. \ b\B. \ a+b , x \ = 0\ by (simp add: \\y\A. \y , x\ = 0\ \\y\B. \y , x\ = 0\ cinner_add_left) hence \\ y \ (A + B). \ y , 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 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 + 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, x\\ proof(cases \\ x. f x = 0\) case True thus ?thesis - by (metis cinner_zero_left) + 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 , x\/\t , t\) *\<^sub>C t\ for x apply (subst spant) by (rule projection_rank1) hence \f (projection (orthogonal_complement (f -` {0})) x) = ((\t , x\)/(\t , 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 , t\) *\<^sub>C t) , 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 , t\) *\<^sub>C t, x\\ for x - by (simp add: l2) + 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, x\\ assumes \\x. f x = \u, 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, y\ = \x, 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\ + define g :: \'a \ 'b \ complex\ where \g x y = \x , 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 + 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, y\\ using riesz_frechet_representation_existence by blast then obtain F where \\x. \y. g x y = \F x, 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\ proof (rule ext) fix x - { + { fix y have \cinner (F1 x - F2 x) y = cinner (F1 x) y - cinner (F2 x) y\ by (simp add: cinner_diff_left) also have \\ = cinner x (G y) - cinner x (G y)\ by (metis assms(1) assms(2) is_cadjoint_def) also have \\ = 0\ by simp finally have \cinner (F1 x - F2 x) y = 0\ by - } then show \F1 x = F2 x\ by fastforce qed lemma cadjoint_univ_prop: fixes G :: "'b::chilbert_space \ 'a::complex_inner" assumes a1: \bounded_clinear G\ shows \\x. \y. \cadjoint G x, y\ = \x, 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. \y. \x, cadjoint G y\ = \G x, 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, y\ = \x, 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 + include notation_norm have b1: \\(A\<^sup>\) x, y\ = \x , A y\\ for x y using cadjoint_univ_prop a1 by auto have \\(A\<^sup>\) (x1 + x2) - ((A\<^sup>\) x1 + (A\<^sup>\) x2) , y\ = 0\ for x1 x2 y - by (simp add: b1 cinner_diff_left cinner_add_left) + 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: \\(A\<^sup>\) (r *\<^sub>C x) - (r *\<^sub>C (A\<^sup>\) x ), y\ = 0\ 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, (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, (A\<^sup>\) x\ \\ for x by (metis abs_pos cinner_ge_zero) hence \\ (A\<^sup>\) x \^2 = \ \x, A ((A\<^sup>\) x)\ \\ for x by (simp add: b1) moreover have \\\x , A ((A\<^sup>\) x)\\ \ \x\ * \A ((A\<^sup>\) x)\\ for x by (simp add: abs_complex_def complex_inner_class.Cauchy_Schwarz_ineq2) 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, hide_lams) bounded_clinear.bounded linear mult_nonneg_nonpos - mult_zero_right norm_ge_zero order.trans semiring_normalization_rules(7)) + 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) + 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\) + 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)) + 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: \(id::'a::complex_inner\'a)\<^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, y \ = \x, (\ 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) + 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 \ \\ proof - have \cinner (x - \ x) y = 0\ if \y\M\ for x y using a1 a2 is_projection_on_iff_orthog orthogonal_complement_orthoI that by blast then have \cinner x y = cinner (\ x) y\ if \y\M\ for x y by (metis cinner_diff_left eq_iff_diff_eq_0 that) moreover have \cinner x y = cinner x (\ y)\ if \y\M\ for x y using a1 is_projection_on_fixes_image that by fastforce ultimately have 1: \cinner (\ x) y = cinner x (\ y)\ if \y\M\ for x y using that by metis have \cinner (\ x) y = 0\ if \y \ orthogonal_complement M\ for x y by (meson a1 is_projection_on_in_image orthogonal_complement_orthoI' that) also have \0 = cinner x (\ y)\ if \y \ orthogonal_complement M\ for x y by (metis a1 a2 cinner_zero_right closed_csubspace.subspace complex_vector.subspace_0 diff_zero is_projection_on_eqI that) finally have 2: \cinner (\ x) y = cinner x (\ y)\ if \y \ orthogonal_complement M\ for x y using that by simp from 1 2 have \cinner (\ x) y = cinner x (\ y)\ for x y by (smt (verit, ccfv_threshold) a1 a2 cinner_commute cinner_diff_left eq_iff_diff_eq_0 is_projection_on_iff_orthog orthogonal_complement_orthoI) then show ?thesis by (simp add: is_cadjoint_def) qed 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) + by (metis closed_csubspace.closed closed_csubspace.subspace csubspace_is_convex empty_iff orthog_proj_exists projection_is_projection_on) instance ccsubspace :: (chilbert_space) complete_orthomodular_lattice proof show "inf x (- x) = bot" for x :: "'a ccsubspace" apply transfer by (simp add: closed_csubspace_def complex_vector.subspace_0 orthogonal_complement_zero_intersection) have \t \ x +\<^sub>M orthogonal_complement x\ if a1: \closed_csubspace x\ for t::'a and x proof- have e1: \t = (projection x) t + (projection (orthogonal_complement x)) t\ by (simp add: that) have e2: \(projection x) t \ x\ by (metis closed_csubspace.closed closed_csubspace.subspace csubspace_is_convex empty_iff orthog_proj_exists projection_in_image that) have e3: \(projection (orthogonal_complement x)) t \ orthogonal_complement x\ by (metis add_diff_cancel_left' e1 orthogonal_complementI projection_orthogonal that) have "orthogonal_complement x \ x +\<^sub>M orthogonal_complement x" by (simp add: closed_sum_right_subset complex_vector.subspace_0 that) thus ?thesis - using \closed_csubspace x\ + using \closed_csubspace x\ \projection (orthogonal_complement x) t \ orthogonal_complement x\ \projection x t \ x\ - \t = projection x t + projection (orthogonal_complement x) t\ in_mono + \t = projection x t + projection (orthogonal_complement x) t\ in_mono closed_sum_left_subset complex_vector.subspace_def - by (metis closed_csubspace.subspace closed_subspace_closed_sum orthogonal_complement_closed_subspace) - qed + by (metis closed_csubspace.subspace closed_subspace_closed_sum orthogonal_complement_closed_subspace) + qed hence b1: \x +\<^sub>M orthogonal_complement x = UNIV\ if a1: \closed_csubspace x\ for x::\'a set\ using that by blast show "sup x (- x) = top" for x :: "'a ccsubspace" apply transfer using b1 by auto show "- (- x) = x" for x :: "'a ccsubspace" apply transfer by (simp) show "- y \ - x" if "x \ y" for x :: "'a ccsubspace" and y :: "'a ccsubspace" using that apply transfer - by simp + by simp have c1: "x +\<^sub>M orthogonal_complement x \ y \ y" if "closed_csubspace x" and "closed_csubspace y" and "x \ y" for x :: "'a set" and y :: "'a set" using that - by (simp add: closed_sum_is_sup) + by (simp add: closed_sum_is_sup) have c2: \u \ x +\<^sub>M ((orthogonal_complement x) \ y)\ if a1: "closed_csubspace x" and a2: "closed_csubspace y" and a3: "x \ y" and x1: \u \ y\ for x :: "'a set" and y :: "'a set" and u proof- have d4: \(projection x) u \ x\ by (metis a1 closed_csubspace_def csubspace_is_convex equals0D orthog_proj_exists projection_in_image) hence d2: \(projection x) u \ y\ using a3 by auto have d1: \csubspace y\ by (simp add: a2) have \u - (projection x) u \ orthogonal_complement x\ by (simp add: a1 orthogonal_complementI projection_orthogonal) moreover have \u - (projection x) u \ y\ - by (simp add: d1 d2 complex_vector.subspace_diff x1) + by (simp add: d1 d2 complex_vector.subspace_diff x1) ultimately have d3: \u - (projection x) u \ ((orthogonal_complement x) \ y)\ by simp hence \\ v \ ((orthogonal_complement x) \ y). u = (projection x) u + v\ by (metis d3 diff_add_cancel ordered_field_class.sign_simps(2)) then obtain v where \v \ ((orthogonal_complement x) \ y)\ and \u = (projection x) u + v\ by blast hence \u \ x + ((orthogonal_complement x) \ y)\ by (metis d4 set_plus_intro) thus ?thesis unfolding closed_sum_def - using closure_subset by blast + using closure_subset by blast qed have c3: "y \ x +\<^sub>M ((orthogonal_complement x) \ y)" if a1: "closed_csubspace x" and a2: "closed_csubspace y" and a3: "x \ y" - for x y :: "'a set" - using c2 a1 a2 a3 by auto + for x y :: "'a set" + using c2 a1 a2 a3 by auto show "sup x (inf (- x) y) = y" if "x \ y" for x y :: "'a ccsubspace" using that apply transfer using c1 c3 by (simp add: subset_antisym) show "x - y = inf x (- y)" for x y :: "'a ccsubspace" apply transfer by simp qed 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 \ \a, s\ = 0" 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 \ \a, s\ = 0" 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: assumes a1: "\s. s \ S \ \a, s\ = 0" and a2: "finite (S::'a::chilbert_space set)" 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) + 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 show "convergent X" if "Cauchy X" for X :: "nat \ 'a" proof- have \finite (set canonical_basis)\ by simp have \Cauchy (\ n. \ t, X n \)\ for t by (simp add: bounded_clinear.Cauchy bounded_clinear_cinner_right that) hence \convergent (\ n. \ t, X n \)\ for t - by (simp add: Cauchy_convergent_iff) + by (simp add: Cauchy_convergent_iff) hence \\ t\set canonical_basis. \ L. (\ n. \ t, X n \) \ L\ by (simp add: convergentD) hence \\ L. \ t\set canonical_basis. (\ n. \ t, X n \) \ L t\ by metis then obtain L where \\ t. t\set canonical_basis \ (\ n. \ t, X n \) \ L t\ by blast define l where \l = (\t\set canonical_basis. L t *\<^sub>C t)\ have x1: \X n = (\t\set canonical_basis. \ t, X n \ *\<^sub>C t)\ for n using onb_expansion_finite[where T = "set canonical_basis" and x = "X n"] - \finite (set canonical_basis)\ + \finite (set canonical_basis)\ by (smt is_generator_set is_normal is_orthonormal) - have \(\ n. \ t, X n \ *\<^sub>C t) \ L t *\<^sub>C t\ + have \(\ n. \ t, X n \ *\<^sub>C t) \ L t *\<^sub>C t\ if r1: \t\set canonical_basis\ for t proof- have \(\ n. \ t, X n \) \ L t\ using r1 \\ t. t\set canonical_basis \ (\ n. \ t, X n \) \ L t\ by blast define f where \f x = x *\<^sub>C t\ for x have \isCont f r\ for r unfolding f_def by (simp add: bounded_clinear_scaleC_left clinear_continuous_at) hence \(\ n. f \ t, X n \) \ f (L t)\ using \(\n. \t, X n\) \ L t\ isCont_tendsto_compose by blast hence \(\ n. \ t, X n \ *\<^sub>C t) \ L t *\<^sub>C t\ unfolding f_def by simp thus ?thesis by blast qed hence \(\ n. (\t\set canonical_basis. \ t, X n \ *\<^sub>C t)) \ (\t\set canonical_basis. L t *\<^sub>C t)\ using \finite (set canonical_basis)\ tendsto_sum[where I = "set canonical_basis" and f = "\ t. \ n. \t, X n\ *\<^sub>C t" and a = "\ t. L t *\<^sub>C t"] - by auto + by auto hence x2: \(\ n. (\t\set canonical_basis. \ t, X n \ *\<^sub>C t)) \ l\ - using l_def by blast + using l_def by blast have \X \ l\ using x1 x2 by simp - thus ?thesis + thus ?thesis unfolding convergent_def by blast qed 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.. end diff --git a/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy b/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy --- a/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy +++ b/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy @@ -1,548 +1,548 @@ (* Based on HOL/Real_Vector_Spaces.thy by Brian Huffman Adapted to the complex case by Dominique Unruh *) section \\Complex_Inner_Product0\ -- Inner Product Spaces and Gradient Derivative\ theory Complex_Inner_Product0 imports Complex_Main Complex_Vector_Spaces "HOL-Analysis.Inner_Product" "Complex_Bounded_Operators.Extra_Ordered_Fields" begin subsection \Complex inner product spaces\ text \ Temporarily relax type constraints for \<^term>\open\, \<^term>\uniformity\, \<^term>\dist\, and \<^term>\norm\. \ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::open set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::dist \ 'a \ real\)\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::norm \ real\)\ class complex_inner = complex_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + fixes cinner :: "'a \ 'a \ complex" assumes cinner_commute: "cinner x y = cnj (cinner y x)" and cinner_add_left: "cinner (x + y) z = cinner x z + cinner y z" and cinner_scaleC_left [simp]: "cinner (scaleC r x) y = (cnj r) * (cinner x y)" and cinner_ge_zero [simp]: "0 \ cinner x x" and cinner_eq_zero_iff [simp]: "cinner x x = 0 \ x = 0" and norm_eq_sqrt_cinner: "norm x = sqrt (cmod (cinner x x))" begin lemma cinner_zero_left [simp]: "cinner 0 x = 0" using cinner_add_left [of 0 0 x] by simp lemma cinner_minus_left [simp]: "cinner (- x) y = - cinner x y" using cinner_add_left [of x "- x" y] by (simp add: group_add_class.add_eq_0_iff) lemma cinner_diff_left: "cinner (x - y) z = cinner x z - cinner y z" using cinner_add_left [of x "- y" z] by simp lemma cinner_sum_left: "cinner (\x\A. f x) y = (\x\A. cinner (f x) y)" by (cases "finite A", induct set: finite, simp_all add: cinner_add_left) lemma call_zero_iff [simp]: "(\u. cinner x u = 0) \ (x = 0)" by auto (use cinner_eq_zero_iff in blast) text \Transfer distributivity rules to right argument.\ lemma cinner_add_right: "cinner x (y + z) = cinner x y + cinner x z" using cinner_add_left [of y z x] by (metis complex_cnj_add local.cinner_commute) lemma cinner_scaleC_right [simp]: "cinner x (scaleC r y) = r * (cinner x y)" using cinner_scaleC_left [of r y x] by (metis complex_cnj_cnj complex_cnj_mult local.cinner_commute) lemma cinner_zero_right [simp]: "cinner x 0 = 0" using cinner_zero_left [of x] - by (metis (mono_tags, hide_lams) complex_cnj_zero local.cinner_commute) + by (metis (mono_tags, opaque_lifting) complex_cnj_zero local.cinner_commute) lemma cinner_minus_right [simp]: "cinner x (- y) = - cinner x y" using cinner_minus_left [of y x] by (metis complex_cnj_minus local.cinner_commute) lemma cinner_diff_right: "cinner x (y - z) = cinner x y - cinner x z" using cinner_diff_left [of y z x] by (metis complex_cnj_diff local.cinner_commute) lemma cinner_sum_right: "cinner x (\y\A. f y) = (\y\A. cinner x (f y))" proof (subst cinner_commute) - have "(\y\A. cinner (f y) x) = (\y\A. cinner (f y) x)" - by blast + have "(\y\A. cinner (f y) x) = (\y\A. cinner (f y) x)" + by blast hence "cnj (\y\A. cinner (f y) x) = cnj (\y\A. (cinner (f y) x))" by simp hence "cnj (cinner (sum f A) x) = (\y\A. cnj (cinner (f y) x))" by (simp add: cinner_sum_left) thus "cnj (cinner (sum f A) x) = (\y\A. (cinner x (f y)))" - by (subst (2) cinner_commute) + by (subst (2) cinner_commute) qed lemmas cinner_add [algebra_simps] = cinner_add_left cinner_add_right lemmas cinner_diff [algebra_simps] = cinner_diff_left cinner_diff_right lemmas cinner_scaleC = cinner_scaleC_left cinner_scaleC_right (* text \Legacy theorem names\ lemmas cinner_left_distrib = cinner_add_left lemmas cinner_right_distrib = cinner_add_right lemmas cinner_distrib = cinner_left_distrib cinner_right_distrib *) lemma cinner_gt_zero_iff [simp]: "0 < cinner x x \ x \ 0" by (smt (verit) less_irrefl local.cinner_eq_zero_iff local.cinner_ge_zero order.not_eq_order_implies_strict) (* In Inner_Product, we have lemma power2_norm_eq_cinner: "(norm x)\<^sup>2 = cinner x x" The following are two ways of inserting the conversions between real and complex into this: *) lemma power2_norm_eq_cinner: shows "(complex_of_real (norm x))\<^sup>2 = (cinner x x)" by (smt (verit, del_insts) Im_complex_of_real Re_complex_of_real cinner_gt_zero_iff cinner_zero_right cmod_def complex_eq_0 complex_eq_iff less_complex_def local.norm_eq_sqrt_cinner of_real_power real_sqrt_abs real_sqrt_pow2_iff zero_complex.sel(1)) lemma power2_norm_eq_cinner': shows "(norm x)\<^sup>2 = Re (cinner x x)" by (metis Re_complex_of_real of_real_power power2_norm_eq_cinner) text \Identities involving real multiplication and division.\ lemma cinner_mult_left: "cinner (of_complex m * a) b = cnj m * (cinner a b)" by (simp add: of_complex_def) lemma cinner_mult_right: "cinner a (of_complex m * b) = m * (cinner a b)" by (metis complex_inner_class.cinner_scaleC_right scaleC_conv_of_complex) lemma cinner_mult_left': "cinner (a * of_complex m) b = cnj m * (cinner a b)" by (metis cinner_mult_left mult.right_neutral mult_scaleC_right scaleC_conv_of_complex) lemma cinner_mult_right': "cinner a (b * of_complex m) = (cinner a b) * m" by (simp add: complex_inner_class.cinner_scaleC_right of_complex_def) (* In Inner_Product, we have lemma Cauchy_Schwarz_ineq: "(cinner x y)\<^sup>2 \ cinner x x * cinner y y" The following are two ways of inserting the conversions between real and complex into this: *) lemma Cauchy_Schwarz_ineq: "(cinner x y) * (cinner y x) \ cinner x x * cinner y y" proof (cases) assume "y = 0" thus ?thesis by simp next assume y: "y \ 0" have [simp]: "cnj (cinner y y) = cinner y y" for y by (metis cinner_commute) define r where "r = cnj (cinner x y) / cinner y y" have "0 \ cinner (x - scaleC r y) (x - scaleC r y)" by (rule cinner_ge_zero) also have "\ = cinner x x - r * cinner x y - cnj r * cinner y x + r * cnj r * cinner y y" unfolding cinner_diff_left cinner_diff_right cinner_scaleC_left cinner_scaleC_right by (smt (z3) cancel_comm_monoid_add_class.diff_cancel cancel_comm_monoid_add_class.diff_zero complex_cnj_divide group_add_class.diff_add_cancel local.cinner_commute local.cinner_eq_zero_iff local.cinner_scaleC_left mult.assoc mult.commute mult_eq_0_iff nonzero_eq_divide_eq r_def y) also have "\ = cinner x x - cinner y x * cnj r" unfolding r_def by auto also have "\ = cinner x x - cinner x y * cnj (cinner x y) / cinner y y" unfolding r_def by (metis complex_cnj_divide local.cinner_commute mult.commute times_divide_eq_left) finally have "0 \ cinner x x - cinner x y * cnj (cinner x y) / cinner y y" . hence "cinner x y * cnj (cinner x y) / cinner y y \ cinner x x" by (simp add: le_diff_eq) thus "cinner x y * cinner y x \ cinner x x * cinner y y" by (metis cinner_gt_zero_iff local.cinner_commute nice_ordered_field_class.pos_divide_le_eq y) qed lemma Cauchy_Schwarz_ineq2: shows "norm (cinner x y) \ norm x * norm y" proof (rule power2_le_imp_le) have "(norm (cinner x y))^2 = Re (cinner x y * cinner y x)" by (metis (full_types) Re_complex_of_real complex_norm_square local.cinner_commute) also have "\ \ Re (cinner x x * cinner y y)" using Cauchy_Schwarz_ineq by (rule Re_mono) also have "\ = Re (complex_of_real ((norm x)^2) * complex_of_real ((norm y)^2))" by (simp add: power2_norm_eq_cinner) also have "\ = (norm x * norm y)\<^sup>2" by (simp add: power_mult_distrib) finally show "(cmod (cinner x y))^2 \ (norm x * norm y)\<^sup>2" . show "0 \ norm x * norm y" by (simp add: local.norm_eq_sqrt_cinner) qed (* The following variant does not hold in the complex case: *) (* lemma norm_cauchy_schwarz: "cinner x y \ norm x * norm y" using Cauchy_Schwarz_ineq2 [of x y] by auto *) subclass complex_normed_vector proof fix a :: complex and r :: real and x y :: 'a show "norm x = 0 \ x = 0" unfolding norm_eq_sqrt_cinner by simp show "norm (x + y) \ norm x + norm y" proof (rule power2_le_imp_le) have "Re (cinner x y) \ cmod (cinner x y)" if "\x. Re x \ cmod x" and "\x y. x \ y \ complex_of_real x \ complex_of_real y" using that by simp hence a1: "2 * Re (cinner x y) \ 2 * cmod (cinner x y)" if "\x. Re x \ cmod x" and "\x y. x \ y \ complex_of_real x \ complex_of_real y" using that by simp have "cinner x y + cinner y x = complex_of_real (2 * Re (cinner x y))" by (metis complex_add_cnj local.cinner_commute) also have "\ \ complex_of_real (2 * cmod (cinner x y))" using complex_Re_le_cmod complex_of_real_mono a1 - by blast + by blast also have "\ = 2 * abs (cinner x y)" unfolding abs_complex_def by simp also have "\ \ 2 * complex_of_real (norm x) * complex_of_real (norm y)" using Cauchy_Schwarz_ineq2 unfolding abs_complex_def by auto - finally have xyyx: "cinner x y + cinner y x \ complex_of_real (2 * norm x * norm y)" + finally have xyyx: "cinner x y + cinner y x \ complex_of_real (2 * norm x * norm y)" by auto have "complex_of_real ((norm (x + y))\<^sup>2) = cinner (x+y) (x+y)" by (simp add: power2_norm_eq_cinner) also have "\ = cinner x x + cinner x y + cinner y x + cinner y y" by (simp add: cinner_add) also have "\ = complex_of_real ((norm x)\<^sup>2) + complex_of_real ((norm y)\<^sup>2) + cinner x y + cinner y x" by (simp add: power2_norm_eq_cinner) also have "\ \ complex_of_real ((norm x)\<^sup>2) + complex_of_real ((norm y)\<^sup>2) + complex_of_real (2 * norm x * norm y)" using xyyx by auto also have "\ = complex_of_real ((norm x + norm y)\<^sup>2)" unfolding power2_sum by auto finally show "(norm (x + y))\<^sup>2 \ (norm x + norm y)\<^sup>2" using complex_of_real_mono_iff by blast show "0 \ norm x + norm y" unfolding norm_eq_sqrt_cinner by simp qed show norm_scaleC: "norm (a *\<^sub>C x) = cmod a * norm x" for a proof (rule power2_eq_imp_eq) show "(norm (a *\<^sub>C x))\<^sup>2 = (cmod a * norm x)\<^sup>2" by (simp_all add: norm_eq_sqrt_cinner norm_mult power2_eq_square) show "0 \ norm (a *\<^sub>C x)" - by (simp_all add: norm_eq_sqrt_cinner) + by (simp_all add: norm_eq_sqrt_cinner) show "0 \ cmod a * norm x" - by (simp_all add: norm_eq_sqrt_cinner) + by (simp_all add: norm_eq_sqrt_cinner) qed show "norm (r *\<^sub>R x) = \r\ * norm x" unfolding scaleR_scaleC norm_scaleC by auto qed end (* Does not hold in the complex case *) (* lemma csquare_bound_lemma: fixes x :: complex shows "x < (1 + x) * (1 + x)" *) lemma csquare_continuous: fixes e :: real shows "e > 0 \ \d. 0 < d \ (\y. cmod (y - x) < d \ cmod (y * y - x * x) < e)" using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] by (force simp add: power2_eq_square) lemma cnorm_le: "norm x \ norm y \ cinner x x \ cinner y y" by (smt (verit) complex_of_real_mono_iff norm_eq_sqrt_cinner norm_ge_zero of_real_power power2_norm_eq_cinner real_sqrt_le_mono real_sqrt_pow2) lemma cnorm_lt: "norm x < norm y \ cinner x x < cinner y y" by (meson cnorm_le less_le_not_le) lemma cnorm_eq: "norm x = norm y \ cinner x x = cinner y y" by (metis norm_eq_sqrt_cinner power2_norm_eq_cinner) lemma cnorm_eq_1: "norm x = 1 \ cinner x x = 1" by (metis cinner_ge_zero complex_of_real_cmod norm_eq_sqrt_cinner norm_one of_real_1 real_sqrt_eq_iff real_sqrt_one) lemma cinner_divide_left: fixes a :: "'a :: {complex_inner,complex_div_algebra}" shows "cinner (a / of_complex m) b = (cinner a b) / cnj m" by (metis cinner_mult_left' complex_cnj_inverse divide_inverse of_complex_inverse ordered_field_class.sign_simps(33)) lemma cinner_divide_right: fixes a :: "'a :: {complex_inner,complex_div_algebra}" shows "cinner a (b / of_complex m) = (cinner a b) / m" by (metis cinner_mult_right' divide_inverse of_complex_inverse) text \ Re-enable constraints for \<^term>\open\, \<^term>\uniformity\, \<^term>\dist\, and \<^term>\norm\. \ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniform_space \ 'a) filter\)\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::real_normed_vector \ real\)\ lemma bounded_sesquilinear_cinner: "bounded_sesquilinear (cinner::'a::complex_inner \ 'a \ complex)" proof fix x y z :: 'a and r :: complex show "cinner (x + y) z = cinner x z + cinner y z" by (rule cinner_add_left) show "cinner x (y + z) = cinner x y + cinner x z" by (rule cinner_add_right) show "cinner (scaleC r x) y = scaleC (cnj r) (cinner x y)" unfolding complex_scaleC_def by (rule cinner_scaleC_left) show "cinner x (scaleC r y) = scaleC r (cinner x y)" unfolding complex_scaleC_def by (rule cinner_scaleC_right) have "\x y::'a. norm (cinner x y) \ norm x * norm y * 1" by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) thus "\K. \x y::'a. norm (cinner x y) \ norm x * norm y * K" by metis qed lemmas tendsto_cinner [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas isCont_cinner [simp] = bounded_bilinear.isCont [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas has_derivative_cinner [derivative_intros] = bounded_bilinear.FDERIV [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas bounded_antilinear_cinner_left = bounded_sesquilinear.bounded_antilinear_left [OF bounded_sesquilinear_cinner] lemmas bounded_clinear_cinner_right = bounded_sesquilinear.bounded_clinear_right [OF bounded_sesquilinear_cinner] lemmas bounded_antilinear_cinner_left_comp = bounded_antilinear_cinner_left[THEN bounded_antilinear_o_bounded_clinear] lemmas bounded_clinear_cinner_right_comp = bounded_clinear_cinner_right[THEN bounded_clinear_compose] lemmas has_derivative_cinner_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_clinear_cinner_right[THEN bounded_clinear.bounded_linear]] lemmas has_derivative_cinner_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_antilinear_cinner_left[THEN bounded_antilinear.bounded_linear]] lemma differentiable_cinner [simp]: "f differentiable (at x within s) \ g differentiable at x within s \ (\x. cinner (f x) (g x)) differentiable at x within s" unfolding differentiable_def by (blast intro: has_derivative_cinner) subsection \Class instances\ instantiation complex :: complex_inner begin definition cinner_complex_def [simp]: "cinner x y = cnj x * y" instance proof fix x y z r :: complex show "cinner x y = cnj (cinner y x)" unfolding cinner_complex_def by auto show "cinner (x + y) z = cinner x z + cinner y z" unfolding cinner_complex_def by (simp add: ring_class.ring_distribs(2)) show "cinner (scaleC r x) y = cnj r * cinner x y" unfolding cinner_complex_def complex_scaleC_def by simp show "0 \ cinner x x" by simp show "cinner x x = 0 \ x = 0" unfolding cinner_complex_def by simp have "cmod (Complex x1 x2) = sqrt (cmod (cinner (Complex x1 x2) (Complex x1 x2)))" for x1 x2 unfolding cinner_complex_def complex_cnj complex_mult complex_norm by (simp add: power2_eq_square) thus "norm x = sqrt (cmod (cinner x x))" - by (cases x, hypsubst_thin) + by (cases x, hypsubst_thin) qed end lemma shows complex_inner_1_left[simp]: "cinner 1 x = x" and complex_inner_1_right[simp]: "cinner x 1 = cnj x" by simp_all (* No analogous to \instantiation complex :: real_inner\ or to lemma complex_inner_1 [simp]: "inner 1 x = Re x" lemma complex_inner_1_right [simp]: "inner x 1 = Re x" lemma complex_inner_i_left [simp]: "inner \ x = Im x" lemma complex_inner_i_right [simp]: "inner x \ = Im x" *) lemma cdot_square_norm: "cinner x x = complex_of_real ((norm x)\<^sup>2)" by (metis Im_complex_of_real Re_complex_of_real cinner_ge_zero complex_eq_iff less_eq_complex_def power2_norm_eq_cinner' zero_complex.simps(2)) lemma cnorm_eq_square: "norm x = a \ 0 \ a \ cinner x x = complex_of_real (a\<^sup>2)" by (metis cdot_square_norm norm_ge_zero of_real_eq_iff power2_eq_iff_nonneg) lemma cnorm_le_square: "norm x \ a \ 0 \ a \ cinner x x \ complex_of_real (a\<^sup>2)" by (smt (verit) cdot_square_norm complex_of_real_mono_iff norm_ge_zero power2_le_imp_le) lemma cnorm_ge_square: "norm x \ a \ a \ 0 \ cinner x x \ complex_of_real (a\<^sup>2)" by (smt (verit, best) antisym_conv cnorm_eq_square cnorm_le_square complex_of_real_nn_iff nn_comparable zero_le_power2) lemma norm_lt_square: "norm x < a \ 0 < a \ cinner x x < complex_of_real (a\<^sup>2)" by (meson cnorm_ge_square cnorm_le_square less_le_not_le) lemma norm_gt_square: "norm x > a \ a < 0 \ cinner x x > complex_of_real (a\<^sup>2)" by (smt (verit, ccfv_SIG) cdot_square_norm complex_of_real_strict_mono_iff norm_ge_zero power2_eq_imp_eq power_mono) text\Dot product in terms of the norm rather than conversely.\ lemmas cinner_simps = cinner_add_left cinner_add_right cinner_diff_right cinner_diff_left cinner_scaleC_left cinner_scaleC_right (* Analogue to both dot_norm and dot_norm_neg *) lemma cdot_norm: "cinner x y = ((norm (x+y))\<^sup>2 - (norm (x-y))\<^sup>2 - \ * (norm (x + \ *\<^sub>C y))\<^sup>2 + \ * (norm (x - \ *\<^sub>C y))\<^sup>2) / 4" unfolding power2_norm_eq_cinner - by (simp add: power2_norm_eq_cinner cinner_add_left cinner_add_right + by (simp add: power2_norm_eq_cinner cinner_add_left cinner_add_right cinner_diff_left cinner_diff_right ring_distribs) -lemma of_complex_inner_1 [simp]: +lemma of_complex_inner_1 [simp]: "cinner (of_complex x) (1 :: 'a :: {complex_inner, complex_normed_algebra_1}) = cnj x" by (metis Complex_Inner_Product0.complex_inner_1_right cinner_complex_def cinner_mult_left complex_cnj_one norm_one of_complex_def power2_norm_eq_cinner scaleC_conv_of_complex) -lemma summable_of_complex_iff: +lemma summable_of_complex_iff: "summable (\x. of_complex (f x) :: 'a :: {complex_normed_algebra_1,complex_inner}) \ summable f" proof assume *: "summable (\x. of_complex (f x) :: 'a)" have "bounded_clinear (cinner (1::'a))" by (rule bounded_clinear_cinner_right) then interpret bounded_linear "\x::'a. cinner 1 x" by (rule bounded_clinear.bounded_linear) from summable [OF *] show "summable f" apply (subst (asm) cinner_commute) by simp next assume sum: "summable f" thus "summable (\x. of_complex (f x) :: 'a)" by (rule summable_of_complex) qed subsection \Gradient derivative\ definition\<^marker>\tag important\ cgderiv :: "['a::complex_inner \ complex, 'a, 'a] \ bool" ("(cGDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where (* Must be "cinner D" not "\h. cinner h D", otherwise not even "cGDERIV id x :> 1" holds *) "cGDERIV f x :> D \ FDERIV f x :> cinner D" lemma cgderiv_deriv [simp]: "cGDERIV f x :> D \ DERIV f x :> cnj D" by (simp only: cgderiv_def has_field_derivative_def cinner_complex_def[THEN ext]) lemma cGDERIV_DERIV_compose: assumes "cGDERIV f x :> df" and "DERIV g (f x) :> cnj dg" shows "cGDERIV (\x. g (f x)) x :> scaleC dg df" proof (insert assms) show "cGDERIV (\x. g (f x)) x :> dg *\<^sub>C df" if "cGDERIV f x :> df" and "(g has_field_derivative cnj dg) (at (f x))" unfolding cgderiv_def has_field_derivative_def cinner_scaleC_left complex_cnj_cnj using that - by (simp add: cgderiv_def has_derivative_compose has_field_derivative_imp_has_derivative) + by (simp add: cgderiv_def has_derivative_compose has_field_derivative_imp_has_derivative) qed (* Not specific to complex/real *) (* lemma has_derivative_subst: "\FDERIV f x :> df; df = d\ \ FDERIV f x :> d" *) lemma cGDERIV_subst: "\cGDERIV f x :> df; df = d\ \ cGDERIV f x :> d" by simp lemma cGDERIV_const: "cGDERIV (\x. k) x :> 0" unfolding cgderiv_def cinner_zero_left[THEN ext] by (rule has_derivative_const) lemma cGDERIV_add: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x + g x) x :> df + dg" unfolding cgderiv_def cinner_add_left[THEN ext] by (rule has_derivative_add) lemma cGDERIV_minus: "cGDERIV f x :> df \ cGDERIV (\x. - f x) x :> - df" unfolding cgderiv_def cinner_minus_left[THEN ext] by (rule has_derivative_minus) lemma cGDERIV_diff: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x - g x) x :> df - dg" unfolding cgderiv_def cinner_diff_left by (rule has_derivative_diff) lemma cGDERIV_scaleC: "\DERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. scaleC (f x) (g x)) x :> (scaleC (cnj (f x)) dg + scaleC (cnj df) (cnj (g x)))" unfolding cgderiv_def has_field_derivative_def cinner_add_left cinner_scaleC_left apply (rule has_derivative_subst) apply (erule (1) has_derivative_scaleC) by (simp add: ac_simps) lemma GDERIV_mult: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x * g x) x :> cnj (f x) *\<^sub>C dg + cnj (g x) *\<^sub>C df" unfolding cgderiv_def apply (rule has_derivative_subst) apply (erule (1) has_derivative_mult) apply (rule ext) by (simp add: cinner_add ac_simps) lemma cGDERIV_inverse: "\cGDERIV f x :> df; f x \ 0\ \ cGDERIV (\x. inverse (f x)) x :> - cnj ((inverse (f x))\<^sup>2) *\<^sub>C df" by (metis DERIV_inverse cGDERIV_DERIV_compose complex_cnj_cnj complex_cnj_minus numerals(2)) (* Don't know if this holds: *) (* lemma cGDERIV_norm: - assumes "x \ 0" shows "cGDERIV (\x. norm x) x :> sgn x" + assumes "x \ 0" shows "cGDERIV (\x. norm x) x :> sgn x" *) lemma has_derivative_norm[derivative_intros]: fixes x :: "'a::complex_inner" - assumes "x \ 0" + assumes "x \ 0" shows "(norm has_derivative (\h. Re (cinner (sgn x) h))) (at x)" thm has_derivative_norm proof - have Re_pos: "0 < Re (cinner x x)" - using assms + using assms by (metis Re_strict_mono cinner_gt_zero_iff zero_complex.simps(1)) - have Re_plus_Re: "Re (cinner x y) + Re (cinner y x) = 2 * Re (cinner x y)" + have Re_plus_Re: "Re (cinner x y) + Re (cinner y x) = 2 * Re (cinner x y)" for x y :: 'a by (metis cinner_commute cnj.simps(1) mult_2_right semiring_normalization_rules(7)) have norm: "norm x = sqrt (Re (cinner x x))" for x :: 'a apply (subst norm_eq_sqrt_cinner, subst cmod_Re) using cinner_ge_zero by auto have v2:"((\x. sqrt (Re (cinner x x))) has_derivative - (\xa. (Re (cinner x xa) + Re (cinner xa x)) * (inverse (sqrt (Re (cinner x x))) / 2))) (at x)" + (\xa. (Re (cinner x xa) + Re (cinner xa x)) * (inverse (sqrt (Re (cinner x x))) / 2))) (at x)" by (rule derivative_eq_intros | simp add: Re_pos)+ have v1: "((\x. sqrt (Re (cinner x x))) has_derivative (\y. Re (cinner x y) / sqrt (Re (cinner x x)))) (at x)" if "((\x. sqrt (Re (cinner x x))) has_derivative (\xa. Re (cinner x xa) * inverse (sqrt (Re (cinner x x))))) (at x)" using that apply (subst divide_real_def) by simp have \(norm has_derivative (\y. Re (cinner x y) / norm x)) (at x)\ using v2 apply (auto simp: Re_plus_Re norm [abs_def]) using v1 by blast then show ?thesis by (auto simp: power2_eq_square sgn_div_norm scaleR_scaleC) qed bundle cinner_syntax begin notation cinner (infix "\\<^sub>C" 70) end bundle no_cinner_syntax begin no_notation cinner (infix "\\<^sub>C" 70) end end diff --git a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy --- a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy +++ b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy @@ -1,2820 +1,2820 @@ section \\Complex_Vector_Spaces\ -- Complex Vector Spaces\ (* -Authors: +Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Vector_Spaces imports "HOL-Analysis.Elementary_Topology" "HOL-Analysis.Operator_Norm" "HOL-Analysis.Elementary_Normed_Spaces" "HOL-Library.Set_Algebras" "HOL-Analysis.Starlike" "HOL-Types_To_Sets.Types_To_Sets" "Complex_Bounded_Operators.Extra_Vector_Spaces" "Complex_Bounded_Operators.Extra_Ordered_Fields" "Complex_Bounded_Operators.Extra_Lattice" "Complex_Bounded_Operators.Extra_General" Complex_Vector_Spaces0 begin bundle notation_norm begin notation norm ("\_\") end subsection \Misc\ lemma (in scaleC) scaleC_real: assumes "r\\" shows "r *\<^sub>C x = Re r *\<^sub>R x" unfolding scaleR_scaleC using assms by simp lemma of_complex_of_real_eq [simp]: "of_complex (of_real n) = of_real n" unfolding of_complex_def of_real_def unfolding scaleR_scaleC by simp lemma Complexs_of_real [simp]: "of_real r \ \" - unfolding Complexs_def of_real_def of_complex_def + unfolding Complexs_def of_real_def of_complex_def apply (subst scaleR_scaleC) by simp lemma Reals_in_Complexs: "\ \ \" unfolding Reals_def by auto lemma (in clinear) "linear f" apply standard by (simp_all add: add scaleC scaleR_scaleC) lemma (in bounded_clinear) bounded_linear: "bounded_linear f" by (simp add: add bounded bounded_linear.intro bounded_linear_axioms.intro linearI scaleC scaleR_scaleC) lemma clinear_times: "clinear (\x. c * x)" for c :: "'a::complex_algebra" by (auto simp: clinearI distrib_left) lemma (in clinear) linear: shows \linear f\ by (simp add: add linearI scaleC scaleR_scaleC) lemma bounded_clinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_clinear f" using assms by (auto intro!: exI bounded_clinear.intro clinearI simp: bounded_clinear_axioms_def) lemma bounded_clinear_id[simp]: \bounded_clinear id\ by (simp add: id_def) (* The following would be a natural inclusion of locales, but unfortunately it leads to name conflicts upon interpretation of bounded_cbilinear *) (* sublocale bounded_cbilinear \ bounded_bilinear by (rule bounded_bilinear) *) definition cbilinear :: \('a::complex_vector \ 'b::complex_vector \ 'c::complex_vector) \ bool\ where \cbilinear = (\ f. (\ y. clinear (\ x. f x y)) \ (\ x. clinear (\ y. f x y)) )\ lemma cbilinear_add_left: assumes \cbilinear f\ shows \f (a + b) c = f a c + f b c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_add_right: assumes \cbilinear f\ shows \f a (b + c) = f a b + f a c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_times: fixes g' :: \'a::complex_vector \ complex\ and g :: \'b::complex_vector \ complex\ assumes \\ x y. h x y = (g' x)*(g y)\ and \clinear g\ and \clinear g'\ shows \cbilinear h\ proof - have w1: "h (b1 + b2) y = h b1 y + h b2 y" for b1 :: 'a and b2 :: 'a and y proof- have \h (b1 + b2) y = g' (b1 + b2) * g y\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = (g' b1 + g' b2) * g y\ using \clinear g'\ unfolding clinear_def by (simp add: assms(3) complex_vector.linear_add) also have \\ = g' b1 * g y + g' b2 * g y\ by (simp add: ring_class.ring_distribs(2)) also have \\ = h b1 y + h b2 y\ - using assms(1) by auto + 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)) + 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 + 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)) + 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) + 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) + by (simp add: t1 t2) qed lemma csubspace_is_subspace: "csubspace A \ subspace A" - apply (rule subspaceI) + 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: +lemma cindependent_implies_independent: assumes "cindependent (S::'a::complex_vector set)" shows "independent S" using assms unfolding dependent_def complex_vector.dependent_def using span_subset_cspan by blast lemma cspan_singleton: "cspan {x} = {\ *\<^sub>C x| \. True}" proof - have \cspan {x} = {y. y\cspan {x}}\ by auto also have \\ = {\ *\<^sub>C x| \. True}\ apply (subst complex_vector.span_breakdown_eq) by auto finally show ?thesis by - qed lemma cspan_as_span: "cspan (B::'a::complex_vector set) = span (B \ scaleC \ ` B)" proof auto let ?cspan = complex_vector.span let ?rspan = real_vector.span fix \ assume cspan: "\ \ ?cspan B" have "\B' r. finite B' \ B' \ B \ \ = (\b\B'. r b *\<^sub>C b)" using complex_vector.span_explicit[of B] cspan by auto then obtain B' r where "finite B'" and "B' \ B" and \_explicit: "\ = (\b\B'. r b *\<^sub>C b)" - by atomize_elim + 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 + 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) + by (simp add: sum.cartesian_product) also have "\ \ ?rspan R" unfolding R_def using x2 - by (rule real_vector.span_sum) + 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 + 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 + 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) + 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)\ + 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\ + 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: +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 + 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) = \" + 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 + unfolding inj_def + by simp have l0: "inj ((*\<^sub>C) (- \)::'a \'a)" - unfolding inj_def - by simp + 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 + 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 + 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 + 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 + 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'\ + 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) + 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) + 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) + 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 + 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 + 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) + by (meson g4 subset_imageE) have "inj ((*\<^sub>C) c::'a\_)" unfolding inj_def - using a3 by auto + using a3 by auto hence "finite S'" - using S_S' finite_imageD g3 subset_inj_on by blast + 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 + 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 + by blast hence "inverse c *\<^sub>C y \ S'" - using that(1) by blast + 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 + 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 + by blast thus "u y = 0" unfolding v_def - by (simp add: a3) + by (simp add: a3) qed thus ?thesis using complex_vector.dependent_explicit - by (simp add: complex_vector.dependent_explicit ) + by (simp add: complex_vector.dependent_explicit ) qed subsection \Antilinear maps and friends\ locale antilinear = additive f for f :: "'a::complex_vector \ 'b::complex_vector" + assumes scaleC: "f (scaleC r x) = cnj r *\<^sub>C f x" sublocale antilinear \ linear proof (rule linearI) show "f (b1 + b2) = f b1 + f b2" for b1 :: 'a and b2 :: 'a - by (simp add: add) + 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) + unfolding scaleR_scaleC by (subst scaleC, simp) qed lemma antilinear_imp_scaleC: fixes D :: "complex \ 'a::complex_vector" assumes "antilinear D" obtains d where "D = (\x. cnj x *\<^sub>C d)" proof - interpret clinear "D o cnj" apply standard apply auto apply (simp add: additive.add assms antilinear.axioms(1)) using assms antilinear.scaleC by fastforce obtain d where "D o cnj = (\x. x *\<^sub>C d)" using clinear_axioms complex_vector.linear_imp_scale by blast then have \D = (\x. cnj x *\<^sub>C d)\ by (metis comp_apply complex_cnj_cnj) then show ?thesis by (rule that) qed corollary complex_antilinearD: fixes f :: "complex \ complex" assumes "antilinear f" obtains c where "f = (\x. c * cnj x)" by (rule antilinear_imp_scaleC [OF assms]) (force simp: scaleC_conv_of_complex) lemma antilinearI: assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>C x) = cnj c *\<^sub>C f x" shows "antilinear f" by standard (rule assms)+ lemma antilinear_o_antilinear: "antilinear f \ antilinear g \ clinear (g o f)" apply (rule clinearI) apply (simp add: additive.add antilinear_def) by (simp add: antilinear.scaleC) lemma clinear_o_antilinear: "antilinear f \ clinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) lemma antilinear_o_clinear: "clinear f \ antilinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) locale bounded_antilinear = antilinear f for f :: "'a::complex_normed_vector \ 'b::complex_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" lemma bounded_antilinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = cnj r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_antilinear f" using assms by (auto intro!: exI bounded_antilinear.intro antilinearI simp: bounded_antilinear_axioms_def) sublocale bounded_antilinear \ bounded_linear apply standard by (fact bounded) lemma (in bounded_antilinear) bounded_linear: "bounded_linear f" by (fact bounded_linear) lemma (in bounded_antilinear) antilinear: "antilinear f" by (fact antilinear_axioms) lemma bounded_antilinear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleC r x) = scaleC (cnj r) (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_antilinear f" by standard (blast intro: assms)+ lemma bounded_antilinear_0[simp]: \bounded_antilinear (\_. 0)\ by (rule bounded_antilinear_intro[where K=0], auto) lemma cnj_bounded_antilinear[simp]: "bounded_antilinear cnj" apply (rule bounded_antilinear_intro [where K = 1]) by auto lemma bounded_antilinear_o_bounded_antilinear: assumes "bounded_antilinear f" and "bounded_antilinear g" shows "bounded_clinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_antilinear g by fact fix b1 b2 b r show "f (g (b1 + b2)) = f (g b1) + f (g b2)" by (simp add: f.add g.add) show "f (g (r *\<^sub>C b)) = r *\<^sub>C f (g b)" by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_antilinear_o_bounded_clinear: assumes "bounded_antilinear f" and "bounded_clinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_clinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_clinear_o_bounded_antilinear: assumes "bounded_clinear f" and "bounded_antilinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_clinear f by fact interpret g: bounded_antilinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x using f.scaleC g.scaleC by fastforce have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bij_clinear_imp_inv_clinear: "clinear (inv f)" if a1: "clinear f" and a2: "bij f" proof fix b1 b2 r b show "inv f (b1 + b2) = inv f b1 + inv f b2" by (simp add: a1 a2 bij_is_inj bij_is_surj complex_vector.linear_add inv_f_eq surj_f_inv_f) show "inv f (r *\<^sub>C b) = r *\<^sub>C inv f b" using that - by (smt bij_inv_eq_iff clinear_def complex_vector.linear_scale) + by (smt bij_inv_eq_iff clinear_def complex_vector.linear_scale) qed locale bounded_sesquilinear = - fixes + fixes prod :: "'a::complex_normed_vector \ 'b::complex_normed_vector \ 'c::complex_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleC_left: "prod (r *\<^sub>C a) b = (cnj r) *\<^sub>C (prod a b)" and scaleC_right: "prod a (r *\<^sub>C b) = r *\<^sub>C (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" sublocale bounded_sesquilinear \ bounded_bilinear apply standard by (auto simp: add_left add_right scaleC_left scaleC_right bounded scaleR_scaleC) -lemma (in bounded_sesquilinear) bounded_bilinear[simp]: "bounded_bilinear prod" +lemma (in bounded_sesquilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by (fact bounded_bilinear_axioms) lemma (in bounded_sesquilinear) bounded_antilinear_left: "bounded_antilinear (\a. prod a b)" apply standard apply (auto simp add: scaleC_left add_left) by (metis ab_semigroup_mult_class.mult_ac(1) bounded) lemma (in bounded_sesquilinear) bounded_clinear_right: "bounded_clinear (\b. prod a b)" apply standard apply (auto simp add: scaleC_right add_right) by (metis ab_semigroup_mult_class.mult_ac(1) ordered_field_class.sign_simps(34) pos_bounded) lemma (in bounded_sesquilinear) comp1: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x. prod (g x))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (g (a + a')) b = prod (g a) b + prod (g a') b" by (simp add: add add_left) show "prod (g a) (b + b') = prod (g a) b + prod (g a) b'" by (simp add: add add_right) show "prod (g (r *\<^sub>C a)) b = cnj r *\<^sub>C prod (g a) b" by (simp add: scaleC scaleC_left) show "prod (g a) (r *\<^sub>C b) = r *\<^sub>C prod (g a) b" by (simp add: scaleC_right) interpret bounded_bilinear \(\x. prod (g x))\ by (simp add: bounded_linear comp1) show "\K. \a b. norm (prod (g a) b) \ norm a * norm b * K" using bounded by blast qed lemma (in bounded_sesquilinear) comp2: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x y. prod x (g y))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (a + a') (g b) = prod a (g b) + prod a' (g b)" by (simp add: add add_left) show "prod a (g (b + b')) = prod a (g b) + prod a (g b')" by (simp add: add add_right) show "prod (r *\<^sub>C a) (g b) = cnj r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_left) show "prod a (g (r *\<^sub>C b)) = r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_right) interpret bounded_bilinear \(\x y. prod x (g y))\ apply (rule bounded_bilinear.flip) using _ bounded_linear apply (rule bounded_bilinear.comp1) using bounded_bilinear by (rule bounded_bilinear.flip) show "\K. \a b. norm (prod a (g b)) \ norm a * norm b * K" using bounded by blast qed -lemma (in bounded_sesquilinear) comp: "bounded_clinear f \ bounded_clinear g \ bounded_sesquilinear (\x y. prod (f x) (g y))" +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) + 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\ + fixes A :: \'b::complex_normed_vector \ 'c::complex_normed_vector\ and B :: \'a::complex_normed_vector \ 'b\ assumes \bounded_clinear A\ and \bounded_clinear B\ shows \bounded_clinear (A \ B)\ by (metis clinear_compose assms(1) assms(2) bounded_clinear_axioms_def bounded_clinear_compose bounded_clinear_def o_def) lemmas isCont_scaleC [simp] = bounded_bilinear.isCont [OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] subsection \Misc 2\ lemmas sums_of_complex = bounded_linear.sums [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas summable_of_complex = bounded_linear.summable [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas suminf_of_complex = bounded_linear.suminf [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_left = bounded_linear.sums[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_left = bounded_linear.summable[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_left = bounded_linear.suminf[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_right = bounded_linear.sums[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_right = bounded_linear.summable[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_right = bounded_linear.suminf[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] -lemma closed_scaleC: +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 + 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: +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) + 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 + 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) + by (simp add: bounded_clinear.bounded_linear le_onorm) hence t2: \bdd_above {(norm (f x)) / norm x | x. True}\ - by fastforce + by fastforce have \continuous_on UNIV ( (*) w ) \ for w::real by simp hence \isCont ( ((*) (cmod r)) ) x\ for x - by simp + 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 + by blast moreover have \mono ((*) (cmod r))\ - by (simp add: monoI ordered_comm_semiring_class.comm_mult_left_mono) + 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) + 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) + 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) + 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 + 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) + 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) + 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) + 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) + 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]: + 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" + "cspan (set canonical_basis) = UNIV" setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a::complex_vector set \ 'a set\)\ lemma cdim_UNIV_basis_enum[simp]: \cdim (UNIV::'a::basis_enum set) = length (canonical_basis::'a list)\ apply (subst is_generator_set[symmetric]) apply (subst complex_vector.dim_span_eq_card_independent) apply (rule is_cindependent_set) using distinct_canonical_basis distinct_card by blast lemma finite_basis: "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = UNIV" proof - from cfinitely_spanned obtain S :: \'a set\ where \finite S\ and \cspan S = UNIV\ by auto from complex_vector.maximal_independent_subset obtain B :: \'a set\ where \B \ S\ and \cindependent B\ and \S \ cspan B\ by metis moreover have \finite B\ using \B \ S\ \finite S\ - by (meson finite_subset) + 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 + \finite_span_complete\ below (see there for an intuition about the mathematical content of the lemmas). However, there is one difference: Here we additionally assume here that there is a bijection rep/abs between a finite type \<^typ>\'basis\ and the set $B$. This is needed to be able to use results about euclidean spaces that are formulated w.r.t. the type class \<^class>\finite\ Since we anyway assume that $B$ is finite, this added assumption does not make the lemma weaker. However, we cannot derive the existence of \<^typ>\'basis\ inside the proof (HOL does not support such reasoning). Therefore we have the type \<^typ>\'basis\ as an explicit assumption and remove it using @{attribute internalize_sort} after the proof.\ lemma finite_span_complete_aux: fixes b :: "'b::real_normed_vector" and B :: "'b set" and rep :: "'basis::finite \ 'b" and abs :: "'b \ 'basis" assumes t: "type_definition rep abs B" and t1: "finite B" and t2: "b\B" and t3: "independent B" shows "\D>0. \\. norm (representation B \ b) \ norm \ * D" and "complete (span B)" proof - define repr where "repr = real_vector.representation B" define repr' where "repr' \ = Abs_euclidean_space (repr \ o rep)" for \ define comb where "comb l = (\b\B. l b *\<^sub>R b)" for l define comb' where "comb' l = comb (Rep_euclidean_space l o abs)" for l have comb_cong: "comb x = comb y" if "\z. z\B \ x z = y z" for x y unfolding comb_def using that by auto have comb_repr[simp]: "comb (repr \) = \" if "\ \ real_vector.span B" for \ - using \comb \ \l. \b\B. l b *\<^sub>R b\ local.repr_def real_vector.sum_representation_eq t1 t3 that - by fastforce + 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) + 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) + 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 + by meson show "finite {b. (if b \ B then x b else 0) \ 0}" - using t1 by auto + 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' \) = \" + 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 + (\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 + 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 + using compact_imp_closed compact_norm_comb' by blast moreover have "0 \ norm ` comb' ` sphere 0 1" - by (simp add: not0) + 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" + 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" + 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 + 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 + 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) + 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 + by simp show "bounded_linear comb'" - by (simp add: blin_comb') + by (simp add: blin_comb') show "\x\UNIV. d * norm x \ norm (comb' x)" - by (simp add: norm_comb') + by (simp add: norm_comb') show "complete (UNIV::'basis euclidean_space set)" - by (simp add: \complete UNIV\) + 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)\ + by (metis (no_types, lifting) \\ \ span B\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' f' fun.map_cong repr_comb t type_definition.Rep_range x_def) thus "\x. \ = comb' x" by auto qed from range_comb' complete_comb' show "complete (real_vector.span B)" by simp qed lemma finite_span_complete[simp]: fixes A :: "'a::real_normed_vector set" assumes "finite A" shows "complete (span A)" text \The span of a finite set is complete.\ proof (cases "A \ {} \ A \ {0}") case True obtain B where BT: "real_vector.span B = real_vector.span A" - and "independent B" + and "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, hide_lams) ex_new_if_finite finite_imageI image_eqI type_definition_def) + by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) note finite_span_complete_aux(2)[internalize_sort "'basis::finite"] note this[OF basisT_finite t] } note this[cancel_type_definition, OF \B\{}\ \finite B\ _ \independent B\] hence "complete (real_vector.span B)" using \B\{}\ by auto thus "complete (real_vector.span A)" unfolding BT by simp next case False thus ?thesis using complete_singleton by auto qed lemma finite_span_representation_bounded: fixes B :: "'a::real_normed_vector set" assumes "finite B" and "independent B" shows "\D>0. \\ b. abs (representation B \ b) \ norm \ * D" text \ Assume $B$ is a finite linear independent set of vectors (in a real normed vector space). Let $\alpha^\psi_b$ be the coefficients of $\psi$ expressed as a linear combination over $B$. Then $\alpha$ is is uniformly cblinfun (i.e., $\lvert\alpha^\psi_b \leq D \lVert\psi\rVert\psi$ for some $D$ independent of $\psi,b$). (This also holds when $b$ is not in the span of $B$ because of the way \real_vector.representation\ is defined in this corner case.)\ proof (cases "B\{}") case True (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) define repr where "repr = real_vector.representation B" { (* Step 1: Create a fake type definition by introducing a new type variable 'basis and then assuming the existence of the morphisms Rep/Abs to B This is then roughly equivalent to "typedef 'basis = B" *) (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux (I.e., we cannot call it 'basis) *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto (* Step 2: We show that our fake typedef 'basisT could be instantiated as type class finite *) have basisT_finite: "class.finite TYPE('basisT)" - apply intro_classes + apply intro_classes using \finite B\ t - by (metis (mono_tags, hide_lams) ex_new_if_finite finite_imageI image_eqI type_definition_def) + 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" + 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) + 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 + 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]: +lemma finite_cspan_complete[simp]: fixes B :: "'a::complex_normed_vector set" assumes "finite B" shows "complete (cspan B)" by (simp add: assms cspan_as_span) lemma finite_span_closed[simp]: fixes B :: "'a::real_normed_vector set" assumes "finite B" shows "closed (real_vector.span B)" by (simp add: assms complete_imp_closed) lemma finite_cspan_closed[simp]: fixes S::\'a::complex_normed_vector set\ assumes a1: \finite S\ - shows \closed (cspan S)\ + shows \closed (cspan S)\ by (simp add: assms complete_imp_closed) lemma closure_finite_cspan: fixes T::\'a::complex_normed_vector set\ assumes \finite T\ shows \closure (cspan T) = cspan T\ by (simp add: assms) lemma finite_cspan_crepresentation_bounded: fixes B :: "'a::complex_normed_vector set" assumes a1: "finite B" and a2: "cindependent B" shows "\D>0. \\ b. norm (crepresentation B \ b) \ norm \ * D" proof - define B' where "B' = (B \ scaleC \ ` B)" have independent_B': "independent B'" using B'_def \cindependent B\ by (simp add: real_independent_from_complex_independent a1) have "finite B'" unfolding B'_def using \finite B\ by simp - obtain D' where "D' > 0" and D': "norm (real_vector.representation B' \ b) \ norm \ * D'" + 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'" + 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 + 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) + 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 + 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" + 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" + 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 + 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 + 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) + 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\)" + 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 + by linarith finally show ?thesis . qed thus ?thesis using assms bounded_clinear_def bounded_clinear_axioms_def by blast qed subsection \Closed subspaces\ lemma csubspace_INF[simp]: "(\x. x \ A \ csubspace x) \ csubspace (\A)" by (simp add: complex_vector.subspace_Inter) locale closed_csubspace = fixes A::"('a::{complex_vector,topological_space}) set" assumes subspace: "csubspace A" assumes closed: "closed A" declare closed_csubspace.subspace[simp] lemma closure_is_csubspace[simp]: fixes A::"('a::complex_normed_vector) set" assumes \csubspace A\ shows \csubspace (closure A)\ proof- have "x \ closure A \ y \ closure A \ x+y \ closure A" for x y proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast assume \y\(closure A)\ then obtain yy where \\ n::nat. yy n \ A\ and \yy \ y\ using closure_sequential by blast - have \\ n::nat. (xx n) + (yy n) \ A\ + 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: 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\ + 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\ + 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\ + 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) + by (metis in_mono) ultimately show ?thesis - by (simp add: complex_vector.subspaceI) + 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) + 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 + by blast moreover have "closed ({0} :: 'a set)" - by simp - ultimately show ?thesis + 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 + by simp moreover have \closed UNIV\ by simp - ultimately show ?thesis + 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" + moreover have "0 \ C" using \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def by fastforce - ultimately show ?thesis + 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 + ultimately show ?thesis by (simp add: closed_csubspace.intro) qed -typedef (overloaded) ('a::"{complex_vector,topological_space}") +typedef (overloaded) ('a::"{complex_vector,topological_space}") ccsubspace = \{S::'a set. closed_csubspace S}\ morphisms space_as_set Abs_clinear_space using Complex_Vector_Spaces.closed_csubspace_UNIV by blast setup_lifting type_definition_ccsubspace lemma csubspace_space_as_set[simp]: \csubspace (space_as_set S)\ by (metis closed_csubspace_def mem_Collect_eq space_as_set) instantiation ccsubspace :: (complex_normed_vector) scaleC begin lift_definition scaleC_ccsubspace :: "complex \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>C) c ` S" proof show "csubspace ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that - by (simp add: closed_csubspace.subspace complex_vector.linear_subspace_image) + by (simp add: closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that - by (simp add: closed_scaleC closed_csubspace.closed) + by (simp add: closed_scaleC closed_csubspace.closed) qed lift_definition scaleR_ccsubspace :: "real \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>R) c ` S" proof show "csubspace ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that using bounded_clinear_def bounded_clinear_scaleC_right scaleR_scaleC by (simp add: scaleR_scaleC closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" - using that + using that by (simp add: closed_scaling closed_csubspace.closed) qed -instance +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) + 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 + 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 + using that by (auto intro !: exI [of _ 0]) qed -lemma csubspace_scaleC_invariant: +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 + 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 + qed ultimately show ?thesis by blast qed lemma ccsubspace_scaleC_invariant[simp]: "a \ 0 \ a *\<^sub>C S = S" for S :: "_ ccsubspace" apply transfer by (simp add: closed_csubspace.subspace csubspace_scaleC_invariant) instantiation ccsubspace :: ("{complex_vector,topological_space}") "top" begin lift_definition top_ccsubspace :: \'a ccsubspace\ is \UNIV\ by simp instance .. end -lemma ccsubspace_top_not_bot[simp]: +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) + by (simp add: closed closed_csubspace.subspace) show "closed (\ S::'a set)" - by (simp add: closed closed_csubspace.closed) + 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) + by (simp add: closure_is_csubspace) show "closed (closure (cspan S))" by simp qed lemma ccspan_canonical_basis[simp]: "ccspan (set canonical_basis) = top" using ccspan.rep_eq space_as_set_inject top_ccsubspace.rep_eq closure_UNIV is_generator_set by metis lemma ccspan_Inf_def: \ccspan A = Inf {S. A \ space_as_set S}\ for A::\('a::cbanach) set\ proof- - have \x \ space_as_set (ccspan A) + 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 + 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 + 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})\ + 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 + 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) + 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 + 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 + by (smt complex_vector.dependent_single complex_vector.independent_insert + complex_vector.scale_eq_0_iff complex_vector.span_base complex_vector.span_redundant + complex_vector.span_scale doubleton_eq_iff insert_absorb insert_absorb2 insert_commute singletonI) lemma closure_is_closed_csubspace[simp]: fixes S::\'a::complex_normed_vector set\ assumes \csubspace S\ shows \closed_csubspace (closure S)\ proof- fix x y :: 'a and c :: complex have "x + y \ closure S" if "x \ closure S" and "y \ closure S" proof- have \\ r. (\ n::nat. r n \ S) \ r \ x\ using closure_sequential that(1) by auto then obtain r where \\ n::nat. r n \ S\ and \r \ x\ by blast have \\ s. (\ n::nat. s n \ S) \ s \ y\ using closure_sequential that(2) by auto then obtain s where \\ n::nat. s n \ S\ and \s \ y\ by blast have \\ n::nat. r n + s n \ S\ - using \\n. r n \ S\ \\n. s n \ S\ assms complex_vector.subspace_add by blast + using \\n. r n \ S\ \\n. s n \ S\ assms complex_vector.subspace_add by blast moreover have \(\ n. r n + s n) \ x + y\ by (simp add: \r \ x\ \s \ y\ tendsto_add) ultimately show ?thesis using assms that(1) that(2) - by (simp add: complex_vector.subspace_add) + by (simp add: complex_vector.subspace_add) qed moreover have "c *\<^sub>C x \ closure S" if "x \ closure S" proof- have \\ y. (\ n::nat. y n \ S) \ y \ x\ using Elementary_Topology.closure_sequential that by auto then obtain y where \\ n::nat. y n \ S\ and \y \ x\ by blast have \isCont (scaleC c) x\ by simp hence \(\ n. scaleC c (y n)) \ scaleC c x\ using \y \ x\ - by (simp add: isCont_tendsto_compose) + by (simp add: isCont_tendsto_compose) from \\ n::nat. y n \ S\ have \\ n::nat. scaleC c (y n) \ S\ using assms complex_vector.subspace_scale by auto thus ?thesis using assms that - by (simp add: complex_vector.subspace_scale) + by (simp add: complex_vector.subspace_scale) qed moreover have "0 \ closure S" by (simp add: assms complex_vector.subspace_0) moreover have "closed (closure S)" by auto ultimately show ?thesis by (simp add: assms closed_csubspace_def) qed lemma ccspan_singleton_scaleC[simp]: "(a::complex)\0 \ ccspan {a *\<^sub>C \} = ccspan {\}" apply transfer by simp lemma clinear_continuous_at: - assumes \bounded_clinear f\ + 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\ + 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\ + 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\ + 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) + 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) + 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 + 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) + using that by (simp add: space_as_set_inject less_eq_ccsubspace.rep_eq) qed end lemma ccspan_leqI: assumes \M \ space_as_set S\ shows \ccspan M \ S\ using assms apply transfer by (simp add: closed_csubspace.closed closure_minimal complex_vector.span_minimal) lemma ccspan_mono: assumes \A \ B\ shows \ccspan A \ ccspan B\ apply (transfer fixing: A B) by (simp add: assms closure_mono complex_vector.span_mono) lemma bounded_sesquilinear_add: \bounded_sesquilinear (\ x y. A x y + B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "A (a + a') b + B (a + a') b = (A a b + B a b) + (A a' b + B a' b)" by (simp add: bounded_sesquilinear.add_left that(1) that(2)) show \A a (b + b') + B a (b + b') = (A a b + B a b) + (A a b' + B a b')\ by (simp add: bounded_sesquilinear.add_right that(1) that(2)) show \A (r *\<^sub>C a) b + B (r *\<^sub>C a) b = cnj r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_left scaleC_add_right that(1) that(2)) show \A a (r *\<^sub>C b) + B a (r *\<^sub>C b) = r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_right scaleC_add_right that(1) that(2)) show \\K. \a b. norm (A a b + B a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \\ KB. \ a b. norm (B a b) \ norm a * norm b * KB\ by (simp add: bounded_sesquilinear.bounded that(2)) then obtain KB where \\ a b. norm (B a b) \ norm a * norm b * KB\ by blast have \norm (A a b + B a b) \ norm a * norm b * (KA + KB)\ for a b proof- have \norm (A a b + B a b) \ norm (A a b) + norm (B a b)\ using norm_triangle_ineq by blast also have \\ \ norm a * norm b * KA + norm a * norm b * KB\ using \\ a b. norm (A a b) \ norm a * norm b * KA\ \\ a b. norm (B a b) \ norm a * norm b * KB\ using add_mono by blast also have \\= norm a * norm b * (KA + KB)\ by (simp add: mult.commute ring_class.ring_distribs(2)) - finally show ?thesis + finally show ?thesis by blast qed thus ?thesis by blast qed qed lemma bounded_sesquilinear_uminus: \bounded_sesquilinear (\ x y. - A x y)\ if \bounded_sesquilinear A\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "- A (a + a') b = (- A a b) + (- A a' b)" by (simp add: bounded_sesquilinear.add_left that) show \- A a (b + b') = (- A a b) + (- A a b')\ by (simp add: bounded_sesquilinear.add_right that) show \- A (r *\<^sub>C a) b = cnj r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_left that) show \- A a (r *\<^sub>C b) = r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_right that) show \\K. \a b. norm (- A a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \norm (- A a b) \ norm a * norm b * KA\ for a b by (simp add: \\a b. norm (A a b) \ norm a * norm b * KA\) thus ?thesis by blast qed qed lemma bounded_sesquilinear_diff: \bounded_sesquilinear (\ x y. A x y - B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof - have \bounded_sesquilinear (\ x y. - B x y)\ using that(2) by (rule bounded_sesquilinear_uminus) then have \bounded_sesquilinear (\ x y. A x y + (- B x y))\ using that(1) by (rule bounded_sesquilinear_add[rotated]) then show ?thesis by auto qed lemma ccsubspace_leI: assumes t1: "space_as_set A \ space_as_set B" shows "A \ B" using t1 apply transfer by - lemma ccspan_of_empty[simp]: "ccspan {} = bot" proof transfer show "closure (cspan {}) = {0::'a}" by simp qed -instantiation ccsubspace :: ("{complex_vector,topological_space}") inf begin -lift_definition inf_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" +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 +instance proof show "a \ \" for a :: "'a ccsubspace" apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") order_bot begin -instance +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 +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 +qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") zero begin definition zero_ccsubspace :: "'a ccsubspace" where [simp]: "zero_ccsubspace = bot" lemma zero_ccsubspace_transfer[transfer_rule]: \pcr_ccsubspace (=) {0} 0\ unfolding zero_ccsubspace_def by transfer_prover instance .. end subsection \Closed sums\ definition closed_sum:: \'a::{semigroup_add,topological_space} set \ 'a set \ 'a set\ where \closed_sum A B = closure (A + B)\ notation closed_sum (infixl "+\<^sub>M" 65) lemma closed_sum_comm: \A +\<^sub>M B = B +\<^sub>M A\ for A B :: "_::ab_semigroup_add" by (simp add: add.commute closed_sum_def) lemma closed_sum_left_subset: \0 \ B \ A \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.right_neutral closed_sum_def closure_subset in_mono set_plus_intro subsetI) lemma closed_sum_right_subset: \0 \ A \ B \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.left_neutral closed_sum_def closure_subset set_plus_intro subset_iff) lemma finite_cspan_closed_csubspace: assumes "finite (S::'a::complex_normed_vector set)" shows "closed_csubspace (cspan S)" by (simp add: assms closed_csubspace.intro) lemma closed_sum_is_sup: fixes A B C:: \('a::{complex_vector,topological_space}) set\ assumes \closed_csubspace C\ assumes \A \ C\ and \B \ C\ shows \(A +\<^sub>M B) \ C\ proof - have \A + B \ C\ using assms unfolding set_plus_def using closed_csubspace.subspace complex_vector.subspace_add by blast then show \(A +\<^sub>M B) \ C\ unfolding closed_sum_def using \closed_csubspace C\ by (simp add: closed_csubspace.closed closure_minimal) qed lemma closed_subspace_closed_sum: fixes A B::"('a::complex_normed_vector) set" assumes a1: \csubspace A\ and a2: \csubspace B\ shows \closed_csubspace (A +\<^sub>M B)\ - using a1 a2 closed_sum_def + 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 + 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" +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 .. + 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: +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) + 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 + 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 + 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 + 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 + 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 + using that apply transfer - by (meson Union_upper closure_subset complex_vector.span_superset dual_order.trans) + 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 + 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 +instance proof fix a b c :: \'a ccsubspace\ show "a + b + c = a + (b + c)" - using sup.assoc by auto + using sup.assoc by auto show "a + b = b + a" - by (simp add: sup.commute) + 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" +lemma ccsubspace_plus_sup: "y \ x \ z \ x \ y + z \ x" for x y z :: "'a::complex_normed_vector ccsubspace" unfolding plus_ccsubspace_def by auto lemma ccsubspace_Sup_empty: "Sup {} = (0::_ ccsubspace)" unfolding zero_ccsubspace_def by auto lemma ccsubspace_add_right_incr[simp]: "a \ a + c" for a::"_ ccsubspace" by (simp add: add_increasing2) lemma ccsubspace_add_left_incr[simp]: "a \ c + a" for a::"_ ccsubspace" by (simp add: add_increasing) subsection \Conjugate space\ typedef 'a conjugate_space = "UNIV :: 'a set" morphisms from_conjugate_space to_conjugate_space .. setup_lifting type_definition_conjugate_space instantiation conjugate_space :: (complex_vector) complex_vector begin lift_definition scaleC_conjugate_space :: \complex \ 'a conjugate_space \ 'a conjugate_space\ is \\c x. cnj c *\<^sub>C x\. lift_definition scaleR_conjugate_space :: \real \ 'a conjugate_space \ 'a conjugate_space\ is \\r x. r *\<^sub>R x\. lift_definition plus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(+)". lift_definition uminus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is \\x. -x\. lift_definition zero_conjugate_space :: "'a conjugate_space" is 0. lift_definition minus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(-)". instance apply (intro_classes; transfer) by (simp_all add: scaleR_scaleC scaleC_add_right scaleC_left.add) end instantiation conjugate_space :: (complex_normed_vector) complex_normed_vector begin lift_definition sgn_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is "sgn". lift_definition norm_conjugate_space :: "'a conjugate_space \ real" is norm. lift_definition dist_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ real" is dist. lift_definition uniformity_conjugate_space :: "('a conjugate_space \ 'a conjugate_space) filter" is uniformity. lift_definition open_conjugate_space :: "'a conjugate_space set \ bool" is "open". -instance +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 +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, hide_lams) complex_cnj_cnj) + by (metis (no_types, opaque_lifting) complex_cnj_cnj) lemma surj_to_conjugate_space[simp]: "surj to_conjugate_space" by (meson surj_def to_conjugate_space_cases) lemmas has_derivative_scaleC[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] lemma norm_to_conjugate_space[simp]: \norm (to_conjugate_space x) = norm x\ by (fact norm_conjugate_space.abs_eq) lemma norm_from_conjugate_space[simp]: \norm (from_conjugate_space x) = norm x\ by (simp add: norm_conjugate_space.rep_eq) lemma closure_to_conjugate_space: \closure (to_conjugate_space ` X) = to_conjugate_space ` closure X\ proof - have 1: \to_conjugate_space ` closure X \ closure (to_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = to_conjugate_space ` from_conjugate_space ` closure (to_conjugate_space ` X)\ by (simp add: from_conjugate_space_inverse image_image) also have \\ \ to_conjugate_space ` closure (from_conjugate_space ` to_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = to_conjugate_space ` closure X\ by (simp add: to_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma closure_from_conjugate_space: \closure (from_conjugate_space ` X) = from_conjugate_space ` closure X\ proof - have 1: \from_conjugate_space ` closure X \ closure (from_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = from_conjugate_space ` to_conjugate_space ` closure (from_conjugate_space ` X)\ by (simp add: to_conjugate_space_inverse image_image) also have \\ \ from_conjugate_space ` closure (to_conjugate_space ` from_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = from_conjugate_space ` closure X\ by (simp add: from_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma bounded_antilinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_antilinear A\ and \bounded_antilinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - let ?A = \\x. A (from_conjugate_space x)\ and ?B = \\x. B (from_conjugate_space x)\ and ?G = \to_conjugate_space ` G\ and ?t = \to_conjugate_space t\ have \bounded_clinear ?A\ and \bounded_clinear ?B\ by (auto intro!: bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear A\] bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear B\]) moreover from eq have \\x. x \ ?G \ ?A x = ?B x\ by (metis image_iff iso_tuple_UNIV_I to_conjugate_space_inverse) moreover from t have \?t \ closure (cspan ?G)\ by (metis bounded_antilinear.bounded_linear bounded_antilinear_to_conjugate_space closure_bounded_linear_image_subset cspan_to_conjugate_space imageI subsetD) ultimately have \?A ?t = ?B ?t\ by (rule bounded_clinear_eq_on) then show \A t = B t\ by (simp add: to_conjugate_space_inverse) qed instantiation complex :: basis_enum begin definition "canonical_basis = [1::complex]" instance proof show "distinct (canonical_basis::complex list)" - by (simp add: canonical_basis_complex_def) + 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 + unfolding canonical_basis_complex_def apply (auto simp add: cspan_raw_def vector_space_over_itself.span_Basis) by (metis complex_scaleC_def complex_vector.span_base complex_vector.span_scale cspan_raw_def insertI1 mult.right_neutral) qed end lemma csubspace_is_convex[simp]: assumes a1: "csubspace M" shows "convex M" proof- have \\x\M. \y\ M. \u. \v. u *\<^sub>C x + v *\<^sub>C y \ M\ using a1 by (simp add: complex_vector.subspace_def) hence \\x\M. \y\M. \u::real. \v::real. u *\<^sub>R x + v *\<^sub>R y \ M\ by (simp add: scaleR_scaleC) hence \\x\M. \y\M. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \M\ by blast thus ?thesis using convex_def by blast qed lemma kernel_is_csubspace[simp]: assumes a1: "clinear f" shows "csubspace (f -` {0})" proof- - have w3: \t *\<^sub>C x \ {x. f x = 0}\ + have w3: \t *\<^sub>C x \ {x. f x = 0}\ if b1: "x \ {x. f x = 0}" for x t by (metis assms complex_vector.linear_subspace_kernel complex_vector.subspace_def that) have \f 0 = 0\ by (simp add: assms complex_vector.linear_0) hence s2: \0 \ {x. f x = 0}\ by blast have w4: "x + y \ {x. f x = 0}" if c1: "x \ {x. f x = 0}" and c2: "y \ {x. f x = 0}" for x y using assms c1 c2 complex_vector.linear_add by fastforce - have s4: \c *\<^sub>C t \ {x. f x = 0}\ + have s4: \c *\<^sub>C t \ {x. f x = 0}\ if "t \ {x. f x = 0}" for t c using that w3 by auto have s5: "u + v \ {x. f x = 0}" if "u \ {x. f x = 0}" and "v \ {x. f x = 0}" for u v - using w4 that(1) that(2) by auto + using w4 that(1) that(2) by auto have f3: "f -` {b. b = 0 \ b \ {}} = {a. f a = 0}" by blast have "csubspace {a. f a = 0}" by (metis complex_vector.subspace_def s2 s4 s5) thus ?thesis using f3 by auto qed lemma kernel_is_closed_csubspace[simp]: assumes a1: "bounded_clinear f" shows "closed_csubspace (f -` {0})" proof- have \csubspace (f -` {0})\ using assms bounded_clinear.clinear complex_vector.linear_subspace_vimage complex_vector.subspace_single_0 by blast have "L \ {x. f x = 0}" if "r \ L" and "\ n. r n \ {x. f x = 0}" - for r and L + 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) + using assms clinear_continuous_at continuous_within_tendsto_compose' that(1) by fastforce hence \(\ n. 0) \ f L\ - using d1 by simp + using d1 by simp hence \f L = 0\ using limI by fastforce thus ?thesis by blast qed then have s3: \closed (f -` {0})\ using closed_sequential_limits by force with \csubspace (f -` {0})\ show ?thesis using closed_csubspace.intro by blast qed lemma range_is_clinear[simp]: assumes a1: "clinear f" shows "csubspace (range f)" using assms complex_vector.linear_subspace_image complex_vector.subspace_UNIV by blast lemma ccspan_superset: - \A \ space_as_set (ccspan A)\ + \A \ space_as_set (ccspan A)\ for A :: \'a::complex_normed_vector set\ apply transfer by (meson closure_subset complex_vector.span_superset subset_trans) subsection \Product is a Complex Vector Space\ (* Follows closely Product_Vector.thy *) instantiation prod :: (complex_vector, complex_vector) complex_vector begin definition scaleC_prod_def: "scaleC r A = (scaleC r (fst A), scaleC r (snd A))" lemma fst_scaleC [simp]: "fst (scaleC r A) = scaleC r (fst A)" unfolding scaleC_prod_def by simp lemma snd_scaleC [simp]: "snd (scaleC r A) = scaleC r (snd A)" unfolding scaleC_prod_def by simp proposition scaleC_Pair [simp]: "scaleC r (a, b) = (scaleC r a, scaleC r b)" unfolding scaleC_prod_def by simp instance proof fix a b :: complex and x y :: "'a \ 'b" show "scaleC a (x + y) = scaleC a x + scaleC a y" by (simp add: scaleC_add_right scaleC_prod_def) show "scaleC (a + b) x = scaleC a x + scaleC b x" by (simp add: Complex_Vector_Spaces.scaleC_prod_def scaleC_left.add) show "scaleC a (scaleC b x) = scaleC (a * b) x" by (simp add: prod_eq_iff) show "scaleC 1 x = x" by (simp add: prod_eq_iff) show \(scaleR :: _ \ _ \ 'a*'b) r = (*\<^sub>C) (complex_of_real r)\ for r by (auto intro!: ext simp: scaleR_scaleC scaleC_prod_def scaleR_prod_def) qed end lemma module_prod_scale_eq_scaleC: "module_prod.scale (*\<^sub>C) (*\<^sub>C) = scaleC" apply (rule ext) apply (rule ext) apply (subst module_prod.scale_def) subgoal by unfold_locales by (simp add: scaleC_prod_def) interpretation complex_vector?: vector_space_prod "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" rewrites "scale = ((*\<^sub>C)::_\_\('a \ 'b))" and "module.dependent (*\<^sub>C) = cdependent" and "module.representation (*\<^sub>C) = crepresentation" and "module.subspace (*\<^sub>C) = csubspace" and "module.span (*\<^sub>C) = cspan" and "vector_space.extend_basis (*\<^sub>C) = cextend_basis" and "vector_space.dim (*\<^sub>C) = cdim" and "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" subgoal by unfold_locales subgoal by (fact module_prod_scale_eq_scaleC) unfolding cdependent_raw_def crepresentation_raw_def csubspace_raw_def cspan_raw_def cextend_basis_raw_def cdim_raw_def clinear_def by (rule refl)+ subsection \Copying existing theorems into sublocales\ context bounded_clinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = continuous lemmas uniform_limit = uniform_limit lemmas Cauchy = Cauchy end context bounded_antilinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = continuous lemmas uniform_limit = uniform_limit end context bounded_cbilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = tendsto lemmas isCont = isCont end context bounded_sesquilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = tendsto lemmas isCont = isCont end lemmas tendsto_scaleC [tendsto_intros] = bounded_cbilinear.tendsto [OF bounded_cbilinear_scaleC] end diff --git a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces0.thy b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces0.thy --- a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces0.thy +++ b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces0.thy @@ -1,1573 +1,1573 @@ (* Based on HOL/Real_Vector_Spaces.thy by Brian Huffman, Johannes Hölzl Adapted to the complex case by Dominique Unruh *) section \\Complex_Vector_Spaces0\ -- Vector Spaces and Algebras over the Complex Numbers\ theory Complex_Vector_Spaces0 imports HOL.Real_Vector_Spaces HOL.Topological_Spaces HOL.Vector_Spaces Complex_Main Jordan_Normal_Form.Conjugate -begin +begin (* Jordan_Normal_Form.Conjugate declares these as simps. Seems too aggressive to me. *) declare less_complex_def[simp del] declare less_eq_complex_def[simp del] subsection \Complex vector spaces\ class scaleC = scaleR + fixes scaleC :: "complex \ 'a \ 'a" (infixr "*\<^sub>C" 75) assumes scaleR_scaleC: "scaleR r = scaleC (complex_of_real r)" begin abbreviation divideC :: "'a \ complex \ 'a" (infixl "'/\<^sub>C" 70) where "x /\<^sub>C c \ inverse c *\<^sub>C x" end class complex_vector = scaleC + ab_group_add + assumes scaleC_add_right: "a *\<^sub>C (x + y) = (a *\<^sub>C x) + (a *\<^sub>C y)" and scaleC_add_left: "(a + b) *\<^sub>C x = (a *\<^sub>C x) + (b *\<^sub>C x)" and scaleC_scaleC[simp]: "a *\<^sub>C (b *\<^sub>C x) = (a * b) *\<^sub>C x" and scaleC_one[simp]: "1 *\<^sub>C x = x" (* Not present in Real_Vector_Spaces *) subclass (in complex_vector) real_vector by (standard, simp_all add: scaleR_scaleC scaleC_add_right scaleC_add_left) class complex_algebra = complex_vector + ring + assumes mult_scaleC_left [simp]: "a *\<^sub>C x * y = a *\<^sub>C (x * y)" and mult_scaleC_right [simp]: "x * a *\<^sub>C y = a *\<^sub>C (x * y)" (* Not present in Real_Vector_Spaces *) subclass (in complex_algebra) real_algebra by (standard, simp_all add: scaleR_scaleC) class complex_algebra_1 = complex_algebra + ring_1 (* Not present in Real_Vector_Spaces *) subclass (in complex_algebra_1) real_algebra_1 .. class complex_div_algebra = complex_algebra_1 + division_ring (* Not present in Real_Vector_Spaces *) subclass (in complex_div_algebra) real_div_algebra .. class complex_field = complex_div_algebra + field (* Not present in Real_Vector_Spaces *) subclass (in complex_field) real_field .. instantiation complex :: complex_field begin definition complex_scaleC_def [simp]: "scaleC a x = a * x" instance proof intro_classes fix r :: real and a b x y :: complex show "((*\<^sub>R) r::complex \ _) = (*\<^sub>C) (complex_of_real r)" by (auto simp add: scaleR_conv_of_real) show "a *\<^sub>C (x + y) = a *\<^sub>C x + a *\<^sub>C y" by (simp add: ring_class.ring_distribs(1)) show "(a + b) *\<^sub>C x = a *\<^sub>C x + b *\<^sub>C x" by (simp add: algebra_simps) show "a *\<^sub>C b *\<^sub>C x = (a * b) *\<^sub>C x" by simp show "1 *\<^sub>C x = x" by simp show "a *\<^sub>C (x::complex) * y = a *\<^sub>C (x * y)" by simp show "(x::complex) * a *\<^sub>C y = a *\<^sub>C (x * y)" by simp qed end locale clinear = Vector_Spaces.linear "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" begin lemmas scaleC = scale end global_interpretation complex_vector: vector_space "scaleC :: complex \ 'a \ 'a :: complex_vector" rewrites "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" and "Vector_Spaces.linear (*) (*\<^sub>C) = clinear" defines cdependent_raw_def: cdependent = complex_vector.dependent and crepresentation_raw_def: crepresentation = complex_vector.representation and csubspace_raw_def: csubspace = complex_vector.subspace and cspan_raw_def: cspan = complex_vector.span and cextend_basis_raw_def: cextend_basis = complex_vector.extend_basis and cdim_raw_def: cdim = complex_vector.dim proof unfold_locales show "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" "Vector_Spaces.linear (*) (*\<^sub>C) = clinear" by (force simp: clinear_def complex_scaleC_def[abs_def])+ qed (use scaleC_add_right scaleC_add_left in auto) (* Not needed since we did the global_interpretation with mandatory complex_vector-prefix: hide_const (open)\ \locale constants\ complex_vector.dependent complex_vector.independent complex_vector.representation complex_vector.subspace complex_vector.span complex_vector.extend_basis complex_vector.dim *) abbreviation "cindependent x \ \ cdependent x" global_interpretation complex_vector: vector_space_pair "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" rewrites "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" and "Vector_Spaces.linear (*) (*\<^sub>C) = clinear" defines cconstruct_raw_def: cconstruct = complex_vector.construct proof unfold_locales show "Vector_Spaces.linear (*) (*\<^sub>C) = clinear" unfolding clinear_def complex_scaleC_def by auto qed (auto simp: clinear_def) (* Not needed since we did the global_interpretation with mandatory complex_vector-prefix: hide_const (open)\ \locale constants\ complex_vector.construct *) lemma clinear_compose: "clinear f \ clinear g \ clinear (g \ f)" unfolding clinear_def by (rule Vector_Spaces.linear_compose) text \Recover original theorem names\ lemmas scaleC_left_commute = complex_vector.scale_left_commute lemmas scaleC_zero_left = complex_vector.scale_zero_left lemmas scaleC_minus_left = complex_vector.scale_minus_left lemmas scaleC_diff_left = complex_vector.scale_left_diff_distrib lemmas scaleC_sum_left = complex_vector.scale_sum_left lemmas scaleC_zero_right = complex_vector.scale_zero_right lemmas scaleC_minus_right = complex_vector.scale_minus_right lemmas scaleC_diff_right = complex_vector.scale_right_diff_distrib lemmas scaleC_sum_right = complex_vector.scale_sum_right lemmas scaleC_eq_0_iff = complex_vector.scale_eq_0_iff lemmas scaleC_left_imp_eq = complex_vector.scale_left_imp_eq lemmas scaleC_right_imp_eq = complex_vector.scale_right_imp_eq lemmas scaleC_cancel_left = complex_vector.scale_cancel_left lemmas scaleC_cancel_right = complex_vector.scale_cancel_right lemma divideC_field_simps[field_simps]: (* In Real_Vector_Spaces, these lemmas are unnamed *) "c \ 0 \ a = b /\<^sub>C c \ c *\<^sub>C a = b" "c \ 0 \ b /\<^sub>C c = a \ b = c *\<^sub>C a" "c \ 0 \ a + b /\<^sub>C c = (c *\<^sub>C a + b) /\<^sub>C c" "c \ 0 \ a /\<^sub>C c + b = (a + c *\<^sub>C b) /\<^sub>C c" "c \ 0 \ a - b /\<^sub>C c = (c *\<^sub>C a - b) /\<^sub>C c" "c \ 0 \ a /\<^sub>C c - b = (a - c *\<^sub>C b) /\<^sub>C c" "c \ 0 \ - (a /\<^sub>C c) + b = (- a + c *\<^sub>C b) /\<^sub>C c" "c \ 0 \ - (a /\<^sub>C c) - b = (- a - c *\<^sub>C b) /\<^sub>C c" for a b :: "'a :: complex_vector" by (auto simp add: scaleC_add_right scaleC_add_left scaleC_diff_right scaleC_diff_left) text \Legacy names -- omitted\ (* lemmas scaleC_left_distrib = scaleC_add_left lemmas scaleC_right_distrib = scaleC_add_right lemmas scaleC_left_diff_distrib = scaleC_diff_left lemmas scaleC_right_diff_distrib = scaleC_diff_right *) lemmas clinear_injective_0 = linear_inj_iff_eq_0 and clinear_injective_on_subspace_0 = linear_inj_on_iff_eq_0 and clinear_cmul = linear_scale and clinear_scaleC = linear_scale_self and csubspace_mul = subspace_scale and cspan_linear_image = linear_span_image and cspan_0 = span_zero and cspan_mul = span_scale and injective_scaleC = injective_scale lemma scaleC_minus1_left [simp]: "scaleC (-1) x = - x" for x :: "'a::complex_vector" using scaleC_minus_left [of 1 x] by simp lemma scaleC_2: fixes x :: "'a::complex_vector" shows "scaleC 2 x = x + x" unfolding one_add_one [symmetric] scaleC_add_left by simp lemma scaleC_half_double [simp]: fixes a :: "'a::complex_vector" shows "(1 / 2) *\<^sub>C (a + a) = a" proof - have "\r. r *\<^sub>C (a + a) = (r * 2) *\<^sub>C a" by (metis scaleC_2 scaleC_scaleC) thus ?thesis by simp qed lemma clinear_scale_complex: fixes c::complex shows "clinear f \ f (c * b) = c * f b" using complex_vector.linear_scale by fastforce interpretation scaleC_left: additive "(\a. scaleC a x :: 'a::complex_vector)" by standard (rule scaleC_add_left) interpretation scaleC_right: additive "(\x. scaleC a x :: 'a::complex_vector)" by standard (rule scaleC_add_right) lemma nonzero_inverse_scaleC_distrib: "a \ 0 \ x \ 0 \ inverse (scaleC a x) = scaleC (inverse a) (inverse x)" for x :: "'a::complex_div_algebra" by (rule inverse_unique) simp lemma inverse_scaleC_distrib: "inverse (scaleC a x) = scaleC (inverse a) (inverse x)" for x :: "'a::{complex_div_algebra,division_ring}" by (metis inverse_zero nonzero_inverse_scaleC_distrib complex_vector.scale_eq_0_iff) (* lemmas sum_constant_scaleC = real_vector.sum_constant_scale\ \legacy name\ *) (* Defined in Real_Vector_Spaces: named_theorems vector_add_divide_simps "to simplify sums of scaled vectors" *) lemma complex_add_divide_simps[vector_add_divide_simps]: (* In Real_Vector_Spaces, these lemmas are unnamed *) "v + (b / z) *\<^sub>C w = (if z = 0 then v else (z *\<^sub>C v + b *\<^sub>C w) /\<^sub>C z)" "a *\<^sub>C v + (b / z) *\<^sub>C w = (if z = 0 then a *\<^sub>C v else ((a * z) *\<^sub>C v + b *\<^sub>C w) /\<^sub>C z)" "(a / z) *\<^sub>C v + w = (if z = 0 then w else (a *\<^sub>C v + z *\<^sub>C w) /\<^sub>C z)" "(a / z) *\<^sub>C v + b *\<^sub>C w = (if z = 0 then b *\<^sub>C w else (a *\<^sub>C v + (b * z) *\<^sub>C w) /\<^sub>C z)" "v - (b / z) *\<^sub>C w = (if z = 0 then v else (z *\<^sub>C v - b *\<^sub>C w) /\<^sub>C z)" "a *\<^sub>C v - (b / z) *\<^sub>C w = (if z = 0 then a *\<^sub>C v else ((a * z) *\<^sub>C v - b *\<^sub>C w) /\<^sub>C z)" "(a / z) *\<^sub>C v - w = (if z = 0 then -w else (a *\<^sub>C v - z *\<^sub>C w) /\<^sub>C z)" "(a / z) *\<^sub>C v - b *\<^sub>C w = (if z = 0 then -b *\<^sub>C w else (a *\<^sub>C v - (b * z) *\<^sub>C w) /\<^sub>C z)" for v :: "'a :: complex_vector" by (simp_all add: divide_inverse_commute scaleC_add_right scaleC_diff_right) lemma ceq_vector_fraction_iff [vector_add_divide_simps]: fixes x :: "'a :: complex_vector" shows "(x = (u / v) *\<^sub>C a) \ (if v=0 then x = 0 else v *\<^sub>C x = u *\<^sub>C a)" by auto (metis (no_types) divide_eq_1_iff divide_inverse_commute scaleC_one scaleC_scaleC) lemma cvector_fraction_eq_iff [vector_add_divide_simps]: fixes x :: "'a :: complex_vector" shows "((u / v) *\<^sub>C a = x) \ (if v=0 then x = 0 else u *\<^sub>C a = v *\<^sub>C x)" by (metis ceq_vector_fraction_iff) lemma complex_vector_affinity_eq: fixes x :: "'a :: complex_vector" assumes m0: "m \ 0" shows "m *\<^sub>C x + c = y \ x = inverse m *\<^sub>C y - (inverse m *\<^sub>C c)" (is "?lhs \ ?rhs") proof assume ?lhs hence "m *\<^sub>C x = y - c" by (simp add: field_simps) hence "inverse m *\<^sub>C (m *\<^sub>C x) = inverse m *\<^sub>C (y - c)" by simp thus "x = inverse m *\<^sub>C y - (inverse m *\<^sub>C c)" using m0 by (simp add: complex_vector.scale_right_diff_distrib) next assume ?rhs with m0 show "m *\<^sub>C x + c = y" by (simp add: complex_vector.scale_right_diff_distrib) qed lemma complex_vector_eq_affinity: "m \ 0 \ y = m *\<^sub>C x + c \ inverse m *\<^sub>C y - (inverse m *\<^sub>C c) = x" for x :: "'a::complex_vector" using complex_vector_affinity_eq[where m=m and x=x and y=y and c=c] by metis lemma scaleC_eq_iff [simp]: "b + u *\<^sub>C a = a + u *\<^sub>C b \ a = b \ u = 1" for a :: "'a::complex_vector" proof (cases "u = 1") case True thus ?thesis by auto next case False have "a = b" if "b + u *\<^sub>C a = a + u *\<^sub>C b" proof - from that have "(u - 1) *\<^sub>C a = (u - 1) *\<^sub>C b" by (simp add: algebra_simps) with False show ?thesis by auto qed thus ?thesis by auto qed lemma scaleC_collapse [simp]: "(1 - u) *\<^sub>C a + u *\<^sub>C a = a" for a :: "'a::complex_vector" by (simp add: algebra_simps) subsection \Embedding of the Complex Numbers into any \complex_algebra_1\: \of_complex\\ definition of_complex :: "complex \ 'a::complex_algebra_1" where "of_complex c = scaleC c 1" lemma scaleC_conv_of_complex: "scaleC r x = of_complex r * x" by (simp add: of_complex_def) lemma of_complex_0 [simp]: "of_complex 0 = 0" by (simp add: of_complex_def) lemma of_complex_1 [simp]: "of_complex 1 = 1" by (simp add: of_complex_def) lemma of_complex_add [simp]: "of_complex (x + y) = of_complex x + of_complex y" by (simp add: of_complex_def scaleC_add_left) lemma of_complex_minus [simp]: "of_complex (- x) = - of_complex x" by (simp add: of_complex_def) lemma of_complex_diff [simp]: "of_complex (x - y) = of_complex x - of_complex y" by (simp add: of_complex_def scaleC_diff_left) lemma of_complex_mult [simp]: "of_complex (x * y) = of_complex x * of_complex y" by (simp add: of_complex_def mult.commute) lemma of_complex_sum[simp]: "of_complex (sum f s) = (\x\s. of_complex (f x))" by (induct s rule: infinite_finite_induct) auto lemma of_complex_prod[simp]: "of_complex (prod f s) = (\x\s. of_complex (f x))" by (induct s rule: infinite_finite_induct) auto lemma nonzero_of_complex_inverse: "x \ 0 \ of_complex (inverse x) = inverse (of_complex x :: 'a::complex_div_algebra)" by (simp add: of_complex_def nonzero_inverse_scaleC_distrib) lemma of_complex_inverse [simp]: "of_complex (inverse x) = inverse (of_complex x :: 'a::{complex_div_algebra,division_ring})" by (simp add: of_complex_def inverse_scaleC_distrib) lemma nonzero_of_complex_divide: "y \ 0 \ of_complex (x / y) = (of_complex x / of_complex y :: 'a::complex_field)" by (simp add: divide_inverse nonzero_of_complex_inverse) lemma of_complex_divide [simp]: "of_complex (x / y) = (of_complex x / of_complex y :: 'a::complex_div_algebra)" by (simp add: divide_inverse) lemma of_complex_power [simp]: "of_complex (x ^ n) = (of_complex x :: 'a::{complex_algebra_1}) ^ n" by (induct n) simp_all lemma of_complex_power_int [simp]: "of_complex (power_int x n) = power_int (of_complex x :: 'a :: {complex_div_algebra,division_ring}) n" by (auto simp: power_int_def) lemma of_complex_eq_iff [simp]: "of_complex x = of_complex y \ x = y" by (simp add: of_complex_def) lemma inj_of_complex: "inj of_complex" by (auto intro: injI) lemmas of_complex_eq_0_iff [simp] = of_complex_eq_iff [of _ 0, simplified] lemmas of_complex_eq_1_iff [simp] = of_complex_eq_iff [of _ 1, simplified] lemma minus_of_complex_eq_of_complex_iff [simp]: "-of_complex x = of_complex y \ -x = y" using of_complex_eq_iff[of "-x" y] by (simp only: of_complex_minus) lemma of_complex_eq_minus_of_complex_iff [simp]: "of_complex x = -of_complex y \ x = -y" using of_complex_eq_iff[of x "-y"] by (simp only: of_complex_minus) lemma of_complex_eq_id [simp]: "of_complex = (id :: complex \ complex)" by (rule ext) (simp add: of_complex_def) text \Collapse nested embeddings.\ lemma of_complex_of_nat_eq [simp]: "of_complex (of_nat n) = of_nat n" by (induct n) auto lemma of_complex_of_int_eq [simp]: "of_complex (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma of_complex_numeral [simp]: "of_complex (numeral w) = numeral w" using of_complex_of_int_eq [of "numeral w"] by simp lemma of_complex_neg_numeral [simp]: "of_complex (- numeral w) = - numeral w" using of_complex_of_int_eq [of "- numeral w"] by simp lemma numeral_power_int_eq_of_complex_cancel_iff [simp]: "power_int (numeral x) n = (of_complex y :: 'a :: {complex_div_algebra, division_ring}) \ power_int (numeral x) n = y" proof - have "power_int (numeral x) n = (of_complex (power_int (numeral x) n) :: 'a)" by simp also have "\ = of_complex y \ power_int (numeral x) n = y" by (subst of_complex_eq_iff) auto finally show ?thesis . qed lemma of_complex_eq_numeral_power_int_cancel_iff [simp]: "(of_complex y :: 'a :: {complex_div_algebra, division_ring}) = power_int (numeral x) n \ y = power_int (numeral x) n" by (subst (1 2) eq_commute) simp lemma of_complex_eq_of_complex_power_int_cancel_iff [simp]: "power_int (of_complex b :: 'a :: {complex_div_algebra, division_ring}) w = of_complex x \ power_int b w = x" by (metis of_complex_power_int of_complex_eq_iff) lemma of_complex_in_Ints_iff [simp]: "of_complex x \ \ \ x \ \" proof safe fix x assume "(of_complex x :: 'a) \ \" then obtain n where "(of_complex x :: 'a) = of_int n" by (auto simp: Ints_def) also have "of_int n = of_complex (of_int n)" by simp finally have "x = of_int n" by (subst (asm) of_complex_eq_iff) thus "x \ \" by auto qed (auto simp: Ints_def) lemma Ints_of_complex [intro]: "x \ \ \ of_complex x \ \" by simp text \Every complex algebra has characteristic zero.\ (* Inherited from real_algebra_1 *) (* instance complex_algebra_1 < ring_char_0 .. *) lemma fraction_scaleC_times [simp]: fixes a :: "'a::complex_algebra_1" shows "(numeral u / numeral v) *\<^sub>C (numeral w * a) = (numeral u * numeral w / numeral v) *\<^sub>C a" by (metis (no_types, lifting) of_complex_numeral scaleC_conv_of_complex scaleC_scaleC times_divide_eq_left) lemma inverse_scaleC_times [simp]: fixes a :: "'a::complex_algebra_1" shows "(1 / numeral v) *\<^sub>C (numeral w * a) = (numeral w / numeral v) *\<^sub>C a" by (metis divide_inverse_commute inverse_eq_divide of_complex_numeral scaleC_conv_of_complex scaleC_scaleC) lemma scaleC_times [simp]: fixes a :: "'a::complex_algebra_1" shows "(numeral u) *\<^sub>C (numeral w * a) = (numeral u * numeral w) *\<^sub>C a" by (simp add: scaleC_conv_of_complex) (* Inherited from real_field *) (* instance complex_field < field_char_0 .. *) subsection \The Set of Real Numbers\ definition Complexs :: "'a::complex_algebra_1 set" ("\") where "\ = range of_complex" lemma Complexs_of_complex [simp]: "of_complex r \ \" by (simp add: Complexs_def) lemma Complexs_of_int [simp]: "of_int z \ \" by (subst of_complex_of_int_eq [symmetric], rule Complexs_of_complex) lemma Complexs_of_nat [simp]: "of_nat n \ \" by (subst of_complex_of_nat_eq [symmetric], rule Complexs_of_complex) lemma Complexs_numeral [simp]: "numeral w \ \" by (subst of_complex_numeral [symmetric], rule Complexs_of_complex) lemma Complexs_0 [simp]: "0 \ \" and Complexs_1 [simp]: "1 \ \" by (simp_all add: Complexs_def) lemma Complexs_add [simp]: "a \ \ \ b \ \ \ a + b \ \" apply (auto simp add: Complexs_def) - by (metis of_complex_add range_eqI) + by (metis of_complex_add range_eqI) lemma Complexs_minus [simp]: "a \ \ \ - a \ \" by (auto simp: Complexs_def) lemma Complexs_minus_iff [simp]: "- a \ \ \ a \ \" using Complexs_minus by fastforce lemma Complexs_diff [simp]: "a \ \ \ b \ \ \ a - b \ \" by (metis Complexs_add Complexs_minus_iff add_uminus_conv_diff) lemma Complexs_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" apply (auto simp add: Complexs_def) by (metis of_complex_mult rangeI) lemma nonzero_Complexs_inverse: "a \ \ \ a \ 0 \ inverse a \ \" for a :: "'a::complex_div_algebra" apply (auto simp add: Complexs_def) - by (metis of_complex_inverse range_eqI) + by (metis of_complex_inverse range_eqI) lemma Complexs_inverse: "a \ \ \ inverse a \ \" for a :: "'a::{complex_div_algebra,division_ring}" using nonzero_Complexs_inverse by fastforce lemma Complexs_inverse_iff [simp]: "inverse x \ \ \ x \ \" for x :: "'a::{complex_div_algebra,division_ring}" by (metis Complexs_inverse inverse_inverse_eq) lemma nonzero_Complexs_divide: "a \ \ \ b \ \ \ b \ 0 \ a / b \ \" for a b :: "'a::complex_field" by (simp add: divide_inverse) lemma Complexs_divide [simp]: "a \ \ \ b \ \ \ a / b \ \" for a b :: "'a::{complex_field,field}" using nonzero_Complexs_divide by fastforce lemma Complexs_power [simp]: "a \ \ \ a ^ n \ \" for a :: "'a::complex_algebra_1" apply (auto simp add: Complexs_def) by (metis range_eqI of_complex_power[symmetric]) lemma Complexs_cases [cases set: Complexs]: assumes "q \ \" obtains (of_complex) c where "q = of_complex c" unfolding Complexs_def proof - from \q \ \\ have "q \ range of_complex" unfolding Complexs_def . then obtain c where "q = of_complex c" .. then show thesis .. qed lemma sum_in_Complexs [intro,simp]: "(\i. i \ s \ f i \ \) \ sum f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Complexs_0 sum.infinite) qed simp_all lemma prod_in_Complexs [intro,simp]: "(\i. i \ s \ f i \ \) \ prod f s \ \" proof (induct s rule: infinite_finite_induct) case infinite then show ?case by (metis Complexs_1 prod.infinite) qed simp_all lemma Complexs_induct [case_names of_complex, induct set: Complexs]: "q \ \ \ (\r. P (of_complex r)) \ P q" by (rule Complexs_cases) auto subsection \Ordered complex vector spaces\ class ordered_complex_vector = complex_vector + ordered_ab_group_add + assumes scaleC_left_mono: "x \ y \ 0 \ a \ a *\<^sub>C x \ a *\<^sub>C y" and scaleC_right_mono: "a \ b \ 0 \ x \ a *\<^sub>C x \ b *\<^sub>C x" begin subclass (in ordered_complex_vector) ordered_real_vector apply standard by (auto simp add: less_eq_complex_def scaleC_left_mono scaleC_right_mono scaleR_scaleC) lemma scaleC_mono: "a \ b \ x \ y \ 0 \ b \ 0 \ x \ a *\<^sub>C x \ b *\<^sub>C y" by (meson order_trans scaleC_left_mono scaleC_right_mono) lemma scaleC_mono': "a \ b \ c \ d \ 0 \ a \ 0 \ c \ a *\<^sub>C c \ b *\<^sub>C d" by (rule scaleC_mono) (auto intro: order.trans) lemma pos_le_divideC_eq [field_simps]: "a \ b /\<^sub>C c \ c *\<^sub>C a \ b" (is "?P \ ?Q") if "0 < c" proof assume ?P with scaleC_left_mono that have "c *\<^sub>C a \ c *\<^sub>C (b /\<^sub>C c)" using preorder_class.less_imp_le by blast with that show ?Q by auto next assume ?Q with scaleC_left_mono that have "c *\<^sub>C a /\<^sub>C c \ b /\<^sub>C c" using less_complex_def less_eq_complex_def by fastforce with that show ?P by auto qed lemma pos_less_divideC_eq [field_simps]: "a < b /\<^sub>C c \ c *\<^sub>C a < b" if "c > 0" using that pos_le_divideC_eq [of c a b] by (auto simp add: le_less) lemma pos_divideC_le_eq [field_simps]: "b /\<^sub>C c \ a \ b \ c *\<^sub>C a" if "c > 0" using that pos_le_divideC_eq [of "inverse c" b a] less_complex_def by auto lemma pos_divideC_less_eq [field_simps]: "b /\<^sub>C c < a \ b < c *\<^sub>C a" if "c > 0" using that pos_less_divideC_eq [of "inverse c" b a] by (simp add: local.less_le_not_le local.pos_divideC_le_eq local.pos_le_divideC_eq) lemma pos_le_minus_divideC_eq [field_simps]: "a \ - (b /\<^sub>C c) \ c *\<^sub>C a \ - b" if "c > 0" using that by (metis local.ab_left_minus local.add.inverse_unique local.add.right_inverse local.add_minus_cancel local.le_minus_iff local.pos_divideC_le_eq local.scaleC_add_right local.scaleC_one local.scaleC_scaleC) lemma pos_less_minus_divideC_eq [field_simps]: "a < - (b /\<^sub>C c) \ c *\<^sub>C a < - b" if "c > 0" using that by (metis le_less less_le_not_le pos_divideC_le_eq pos_divideC_less_eq pos_le_minus_divideC_eq) lemma pos_minus_divideC_le_eq [field_simps]: "- (b /\<^sub>C c) \ a \ - b \ c *\<^sub>C a" if "c > 0" using that by (metis local.add_minus_cancel local.left_minus local.pos_divideC_le_eq local.scaleC_add_right) lemma pos_minus_divideC_less_eq [field_simps]: "- (b /\<^sub>C c) < a \ - b < c *\<^sub>C a" if "c > 0" - using that by (simp add: less_le_not_le pos_le_minus_divideC_eq pos_minus_divideC_le_eq) + using that by (simp add: less_le_not_le pos_le_minus_divideC_eq pos_minus_divideC_le_eq) lemma scaleC_image_atLeastAtMost: "c > 0 \ scaleC c ` {x..y} = {c *\<^sub>C x..c *\<^sub>C y}" apply (auto intro!: scaleC_left_mono simp: image_iff Bex_def) - by (meson local.eq_iff pos_divideC_le_eq pos_le_divideC_eq) + by (meson order.eq_iff pos_divideC_le_eq pos_le_divideC_eq) end (* class ordered_complex_vector *) lemma neg_le_divideC_eq [field_simps]: "a \ b /\<^sub>C c \ b \ c *\<^sub>C a" (is "?P \ ?Q") if "c < 0" for a b :: "'a :: ordered_complex_vector" using that pos_le_divideC_eq [of "- c" a "- b"] by (simp add: less_complex_def) lemma neg_less_divideC_eq [field_simps]: "a < b /\<^sub>C c \ b < c *\<^sub>C a" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that neg_le_divideC_eq [of c a b] by (smt (verit, ccfv_SIG) neg_le_divideC_eq antisym_conv2 complex_vector.scale_minus_right dual_order.strict_implies_order le_less_trans neg_le_iff_le scaleC_scaleC) lemma neg_divideC_le_eq [field_simps]: "b /\<^sub>C c \ a \ c *\<^sub>C a \ b" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that pos_divideC_le_eq [of "- c" "- b" a] by (simp add: less_complex_def) lemma neg_divideC_less_eq [field_simps]: "b /\<^sub>C c < a \ c *\<^sub>C a < b" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that neg_divideC_le_eq [of c b a] by (meson neg_le_divideC_eq less_le_not_le) lemma neg_le_minus_divideC_eq [field_simps]: "a \ - (b /\<^sub>C c) \ - b \ c *\<^sub>C a" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that pos_le_minus_divideC_eq [of "- c" a "- b"] by (metis neg_le_divideC_eq complex_vector.scale_minus_right) lemma neg_less_minus_divideC_eq [field_simps]: "a < - (b /\<^sub>C c) \ - b < c *\<^sub>C a" if "c < 0" for a b :: "'a :: ordered_complex_vector" proof - have *: "- b = c *\<^sub>C a \ b = - (c *\<^sub>C a)" by (metis add.inverse_inverse) from that neg_le_minus_divideC_eq [of c a b] show ?thesis by (auto simp add: le_less *) qed lemma neg_minus_divideC_le_eq [field_simps]: "- (b /\<^sub>C c) \ a \ c *\<^sub>C a \ - b" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that pos_minus_divideC_le_eq [of "- c" "- b" a] by (metis Complex_Vector_Spaces0.neg_divideC_le_eq complex_vector.scale_minus_right) lemma neg_minus_divideC_less_eq [field_simps]: "- (b /\<^sub>C c) < a \ c *\<^sub>C a < - b" if "c < 0" for a b :: "'a :: ordered_complex_vector" using that by (simp add: less_le_not_le neg_le_minus_divideC_eq neg_minus_divideC_le_eq) lemma divideC_field_splits_simps_1 [field_split_simps]: (* In Real_Vector_Spaces, these lemmas are unnamed *) "a = b /\<^sub>C c \ (if c = 0 then a = 0 else c *\<^sub>C a = b)" "b /\<^sub>C c = a \ (if c = 0 then a = 0 else b = c *\<^sub>C a)" "a + b /\<^sub>C c = (if c = 0 then a else (c *\<^sub>C a + b) /\<^sub>C c)" "a /\<^sub>C c + b = (if c = 0 then b else (a + c *\<^sub>C b) /\<^sub>C c)" "a - b /\<^sub>C c = (if c = 0 then a else (c *\<^sub>C a - b) /\<^sub>C c)" "a /\<^sub>C c - b = (if c = 0 then - b else (a - c *\<^sub>C b) /\<^sub>C c)" "- (a /\<^sub>C c) + b = (if c = 0 then b else (- a + c *\<^sub>C b) /\<^sub>C c)" "- (a /\<^sub>C c) - b = (if c = 0 then - b else (- a - c *\<^sub>C b) /\<^sub>C c)" for a b :: "'a :: complex_vector" by (auto simp add: field_simps) lemma divideC_field_splits_simps_2 [field_split_simps]: (* In Real_Vector_Spaces, these lemmas are unnamed *) "0 < c \ a \ b /\<^sub>C c \ (if c > 0 then c *\<^sub>C a \ b else if c < 0 then b \ c *\<^sub>C a else a \ 0)" "0 < c \ a < b /\<^sub>C c \ (if c > 0 then c *\<^sub>C a < b else if c < 0 then b < c *\<^sub>C a else a < 0)" "0 < c \ b /\<^sub>C c \ a \ (if c > 0 then b \ c *\<^sub>C a else if c < 0 then c *\<^sub>C a \ b else a \ 0)" "0 < c \ b /\<^sub>C c < a \ (if c > 0 then b < c *\<^sub>C a else if c < 0 then c *\<^sub>C a < b else a > 0)" "0 < c \ a \ - (b /\<^sub>C c) \ (if c > 0 then c *\<^sub>C a \ - b else if c < 0 then - b \ c *\<^sub>C a else a \ 0)" "0 < c \ a < - (b /\<^sub>C c) \ (if c > 0 then c *\<^sub>C a < - b else if c < 0 then - b < c *\<^sub>C a else a < 0)" "0 < c \ - (b /\<^sub>C c) \ a \ (if c > 0 then - b \ c *\<^sub>C a else if c < 0 then c *\<^sub>C a \ - b else a \ 0)" "0 < c \ - (b /\<^sub>C c) < a \ (if c > 0 then - b < c *\<^sub>C a else if c < 0 then c *\<^sub>C a < - b else a > 0)" for a b :: "'a :: ordered_complex_vector" by (clarsimp intro!: field_simps)+ lemma scaleC_nonneg_nonneg: "0 \ a \ 0 \ x \ 0 \ a *\<^sub>C x" for x :: "'a::ordered_complex_vector" using scaleC_left_mono [of 0 x a] by simp lemma scaleC_nonneg_nonpos: "0 \ a \ x \ 0 \ a *\<^sub>C x \ 0" for x :: "'a::ordered_complex_vector" using scaleC_left_mono [of x 0 a] by simp lemma scaleC_nonpos_nonneg: "a \ 0 \ 0 \ x \ a *\<^sub>C x \ 0" for x :: "'a::ordered_complex_vector" using scaleC_right_mono [of a 0 x] by simp lemma split_scaleC_neg_le: "(0 \ a \ x \ 0) \ (a \ 0 \ 0 \ x) \ a *\<^sub>C x \ 0" for x :: "'a::ordered_complex_vector" by (auto simp: scaleC_nonneg_nonpos scaleC_nonpos_nonneg) lemma cle_add_iff1: "a *\<^sub>C e + c \ b *\<^sub>C e + d \ (a - b) *\<^sub>C e + c \ d" for c d e :: "'a::ordered_complex_vector" by (simp add: algebra_simps) lemma cle_add_iff2: "a *\<^sub>C e + c \ b *\<^sub>C e + d \ c \ (b - a) *\<^sub>C e + d" for c d e :: "'a::ordered_complex_vector" by (simp add: algebra_simps) lemma scaleC_left_mono_neg: "b \ a \ c \ 0 \ c *\<^sub>C a \ c *\<^sub>C b" for a b :: "'a::ordered_complex_vector" by (drule scaleC_left_mono [of _ _ "- c"], simp_all add: less_eq_complex_def) lemma scaleC_right_mono_neg: "b \ a \ c \ 0 \ a *\<^sub>C c \ b *\<^sub>C c" for c :: "'a::ordered_complex_vector" by (drule scaleC_right_mono [of _ _ "- c"], simp_all) lemma scaleC_nonpos_nonpos: "a \ 0 \ b \ 0 \ 0 \ a *\<^sub>C b" for b :: "'a::ordered_complex_vector" using scaleC_right_mono_neg [of a 0 b] by simp lemma split_scaleC_pos_le: "(0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0) \ 0 \ a *\<^sub>C b" for b :: "'a::ordered_complex_vector" by (auto simp: scaleC_nonneg_nonneg scaleC_nonpos_nonpos) lemma zero_le_scaleC_iff: fixes b :: "'a::ordered_complex_vector" assumes "a \ \" (* Not present in Real_Vector_Spaces.thy *) shows "0 \ a *\<^sub>C b \ 0 < a \ 0 \ b \ a < 0 \ b \ 0 \ a = 0" (is "?lhs = ?rhs") proof (cases "a = 0") case True then show ?thesis by simp next case False show ?thesis proof assume ?lhs from \a \ 0\ consider "a > 0" | "a < 0" by (metis assms complex_is_Real_iff less_complex_def less_eq_complex_def not_le order.not_eq_order_implies_strict that(1) zero_complex.sel(2)) then show ?rhs proof cases case 1 with \?lhs\ have "inverse a *\<^sub>C 0 \ inverse a *\<^sub>C (a *\<^sub>C b)" by (metis complex_vector.scale_zero_right ordered_complex_vector_class.pos_le_divideC_eq) with 1 show ?thesis by simp next case 2 with \?lhs\ have "- inverse a *\<^sub>C 0 \ - inverse a *\<^sub>C (a *\<^sub>C b)" by (metis Complex_Vector_Spaces0.neg_le_minus_divideC_eq complex_vector.scale_zero_right neg_le_0_iff_le scaleC_left.minus) with 2 show ?thesis by simp qed next assume ?rhs then show ?lhs using less_imp_le split_scaleC_pos_le by auto qed qed lemma scaleC_le_0_iff: "a *\<^sub>C b \ 0 \ 0 < a \ b \ 0 \ a < 0 \ 0 \ b \ a = 0" if "a \ \" (* Not present in Real_Vector_Spaces *) for b::"'a::ordered_complex_vector" apply (insert zero_le_scaleC_iff [of "-a" b]) using less_complex_def that by force lemma scaleC_le_cancel_left: "c *\<^sub>C a \ c *\<^sub>C b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" if "c \ \" (* Not present in Real_Vector_Spaces *) for b :: "'a::ordered_complex_vector" by (smt (verit, ccfv_threshold) Complex_Vector_Spaces0.neg_divideC_le_eq complex_vector.scale_cancel_left complex_vector.scale_zero_right dual_order.eq_iff dual_order.trans ordered_complex_vector_class.pos_le_divideC_eq that zero_le_scaleC_iff) lemma scaleC_le_cancel_left_pos: "0 < c \ c *\<^sub>C a \ c *\<^sub>C b \ a \ b" for b :: "'a::ordered_complex_vector" by (simp add: complex_is_Real_iff less_complex_def scaleC_le_cancel_left) lemma scaleC_le_cancel_left_neg: "c < 0 \ c *\<^sub>C a \ c *\<^sub>C b \ b \ a" for b :: "'a::ordered_complex_vector" by (simp add: complex_is_Real_iff less_complex_def scaleC_le_cancel_left) lemma scaleC_left_le_one_le: "0 \ x \ a \ 1 \ a *\<^sub>C x \ x" for x :: "'a::ordered_complex_vector" and a :: complex using scaleC_right_mono[of a 1 x] by simp subsection \Complex normed vector spaces\ (* Classes dist, norm, sgn_div_norm, dist_norm, uniformity_dist defined in Real_Vector_Spaces are unchanged in the complex setting. No need to define them here. *) class complex_normed_vector = complex_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + real_normed_vector + (* Not present in Real_Normed_Vector *) assumes norm_scaleC [simp]: "norm (scaleC a x) = cmod a * norm x" begin (* lemma norm_ge_zero [simp]: "0 \ norm x" *) (* Not needed, included from real_normed_vector *) end class complex_normed_algebra = complex_algebra + complex_normed_vector + real_normed_algebra (* Not present in Real_Normed_Vector *) (* assumes norm_mult_ineq: "norm (x * y) \ norm x * norm y" *) (* Not needed, included from real_normed_algebra *) class complex_normed_algebra_1 = complex_algebra_1 + complex_normed_algebra + real_normed_algebra_1 (* Not present in Real_Normed_Vector *) (* assumes norm_one [simp]: "norm 1 = 1" *) (* Not needed, included from real_normed_algebra_1 *) lemma (in complex_normed_algebra_1) scaleC_power [simp]: "(scaleC x y) ^ n = scaleC (x^n) (y^n)" by (induct n) (simp_all add: mult_ac) class complex_normed_div_algebra = complex_div_algebra + complex_normed_vector + real_normed_div_algebra (* Not present in Real_Normed_Vector *) (* assumes norm_mult: "norm (x * y) = norm x * norm y" *) (* Not needed, included from real_normed_div_algebra *) class complex_normed_field = complex_field + complex_normed_div_algebra subclass (in complex_normed_field) real_normed_field .. instance complex_normed_div_algebra < complex_normed_algebra_1 .. context complex_normed_vector begin (* Inherited from real_normed_vector: lemma norm_zero [simp]: "norm (0::'a) = 0" lemma zero_less_norm_iff [simp]: "norm x > 0 \ x \ 0" lemma norm_not_less_zero [simp]: "\ norm x < 0" lemma norm_le_zero_iff [simp]: "norm x \ 0 \ x = 0" lemma norm_minus_cancel [simp]: "norm (- x) = norm x" lemma norm_minus_commute: "norm (a - b) = norm (b - a)" lemma dist_add_cancel [simp]: "dist (a + b) (a + c) = dist b c" lemma dist_add_cancel2 [simp]: "dist (b + a) (c + a) = dist b c" lemma norm_uminus_minus: "norm (- x - y) = norm (x + y)" lemma norm_triangle_ineq2: "norm a - norm b \ norm (a - b)" lemma norm_triangle_ineq3: "\norm a - norm b\ \ norm (a - b)" lemma norm_triangle_ineq4: "norm (a - b) \ norm a + norm b" lemma norm_triangle_le_diff: "norm x + norm y \ e \ norm (x - y) \ e" lemma norm_diff_ineq: "norm a - norm b \ norm (a + b)" lemma norm_triangle_sub: "norm x \ norm y + norm (x - y)" lemma norm_triangle_le: "norm x + norm y \ e \ norm (x + y) \ e" lemma norm_triangle_lt: "norm x + norm y < e \ norm (x + y) < e" lemma norm_add_leD: "norm (a + b) \ c \ norm b \ norm a + c" lemma norm_diff_triangle_ineq: "norm ((a + b) - (c + d)) \ norm (a - c) + norm (b - d)" lemma norm_diff_triangle_le: "norm (x - z) \ e1 + e2" if "norm (x - y) \ e1" "norm (y - z) \ e2" lemma norm_diff_triangle_less: "norm (x - z) < e1 + e2" if "norm (x - y) < e1" "norm (y - z) < e2" lemma norm_triangle_mono: "norm a \ r \ norm b \ s \ norm (a + b) \ r + s" lemma norm_sum: "norm (sum f A) \ (\i\A. norm (f i))" for f::"'b \ 'a" lemma sum_norm_le: "norm (sum f S) \ sum g S" if "\x. x \ S \ norm (f x) \ g x" for f::"'b \ 'a" lemma abs_norm_cancel [simp]: "\norm a\ = norm a" lemma sum_norm_bound: "norm (sum f S) \ of_nat (card S)*K" if "\x. x \ S \ norm (f x) \ K" for f :: "'b \ 'a" lemma norm_add_less: "norm x < r \ norm y < s \ norm (x + y) < r + s" *) end lemma dist_scaleC [simp]: "dist (x *\<^sub>C a) (y *\<^sub>C a) = \x - y\ * norm a" for a :: "'a::complex_normed_vector" by (metis dist_scaleR scaleR_scaleC) (* Inherited from real_normed_vector *) (* lemma norm_mult_less: "norm x < r \ norm y < s \ norm (x * y) < r * s" for x y :: "'a::complex_normed_algebra" *) lemma norm_of_complex [simp]: "norm (of_complex c :: 'a::complex_normed_algebra_1) = cmod c" by (simp add: of_complex_def) (* Inherited from real_normed_vector: lemma norm_numeral [simp]: "norm (numeral w::'a::complex_normed_algebra_1) = numeral w" lemma norm_neg_numeral [simp]: "norm (- numeral w::'a::complex_normed_algebra_1) = numeral w" lemma norm_of_complex_add1 [simp]: "norm (of_real x + 1 :: 'a :: complex_normed_div_algebra) = \x + 1\" lemma norm_of_complex_addn [simp]: "norm (of_real x + numeral b :: 'a :: complex_normed_div_algebra) = \x + numeral b\" lemma norm_of_int [simp]: "norm (of_int z::'a::complex_normed_algebra_1) = \of_int z\" lemma norm_of_nat [simp]: "norm (of_nat n::'a::complex_normed_algebra_1) = of_nat n" lemma nonzero_norm_inverse: "a \ 0 \ norm (inverse a) = inverse (norm a)" for a :: "'a::complex_normed_div_algebra" lemma norm_inverse: "norm (inverse a) = inverse (norm a)" for a :: "'a::{complex_normed_div_algebra,division_ring}" lemma nonzero_norm_divide: "b \ 0 \ norm (a / b) = norm a / norm b" for a b :: "'a::complex_normed_field" lemma norm_divide: "norm (a / b) = norm a / norm b" for a b :: "'a::{complex_normed_field,field}" lemma norm_inverse_le_norm: fixes x :: "'a::complex_normed_div_algebra" shows "r \ norm x \ 0 < r \ norm (inverse x) \ inverse r" lemma norm_power_ineq: "norm (x ^ n) \ norm x ^ n" for x :: "'a::complex_normed_algebra_1" lemma norm_power: "norm (x ^ n) = norm x ^ n" for x :: "'a::complex_normed_div_algebra" lemma norm_power_int: "norm (power_int x n) = power_int (norm x) n" for x :: "'a::complex_normed_div_algebra" lemma power_eq_imp_eq_norm: fixes w :: "'a::complex_normed_div_algebra" assumes eq: "w ^ n = z ^ n" and "n > 0" shows "norm w = norm z" lemma power_eq_1_iff: fixes w :: "'a::complex_normed_div_algebra" shows "w ^ n = 1 \ norm w = 1 \ n = 0" lemma norm_mult_numeral1 [simp]: "norm (numeral w * a) = numeral w * norm a" for a b :: "'a::{complex_normed_field,field}" lemma norm_mult_numeral2 [simp]: "norm (a * numeral w) = norm a * numeral w" for a b :: "'a::{complex_normed_field,field}" lemma norm_divide_numeral [simp]: "norm (a / numeral w) = norm a / numeral w" for a b :: "'a::{complex_normed_field,field}" lemma square_norm_one: fixes x :: "'a::complex_normed_div_algebra" assumes "x\<^sup>2 = 1" shows "norm x = 1" lemma norm_less_p1: "norm x < norm (of_real (norm x) + 1 :: 'a)" for x :: "'a::complex_normed_algebra_1" lemma prod_norm: "prod (\x. norm (f x)) A = norm (prod f A)" for f :: "'a \ 'b::{comm_semiring_1,complex_normed_div_algebra}" lemma norm_prod_le: "norm (prod f A) \ (\a\A. norm (f a :: 'a :: {complex_normed_algebra_1,comm_monoid_mult}))" lemma norm_prod_diff: fixes z w :: "'i \ 'a::{complex_normed_algebra_1, comm_monoid_mult}" shows "(\i. i \ I \ norm (z i) \ 1) \ (\i. i \ I \ norm (w i) \ 1) \ norm ((\i\I. z i) - (\i\I. w i)) \ (\i\I. norm (z i - w i))" lemma norm_power_diff: fixes z w :: "'a::{complex_normed_algebra_1, comm_monoid_mult}" assumes "norm z \ 1" "norm w \ 1" shows "norm (z^m - w^m) \ m * norm (z - w)" *) lemma norm_of_complex_add1 [simp]: "norm (of_complex x + 1 :: 'a :: complex_normed_div_algebra) = cmod (x + 1)" by (metis norm_of_complex of_complex_1 of_complex_add) lemma norm_of_complex_addn [simp]: "norm (of_complex x + numeral b :: 'a :: complex_normed_div_algebra) = cmod (x + numeral b)" by (metis norm_of_complex of_complex_add of_complex_numeral) lemma norm_of_complex_diff [simp]: "norm (of_complex b - of_complex a :: 'a::complex_normed_algebra_1) \ cmod (b - a)" by (metis norm_of_complex of_complex_diff order_refl) subsection \Metric spaces\ (* Class metric_space is already defined in Real_Vector_Spaces and does not need changing here *) text \Every normed vector space is a metric space.\ (* Already follows from complex_normed_vector < real_normed_vector < metric_space *) (* instance complex_normed_vector < metric_space *) subsection \Class instances for complex numbers\ instantiation complex :: complex_normed_field begin instance apply intro_classes by (simp add: norm_mult) end declare uniformity_Abort[where 'a=complex, code] lemma dist_of_complex [simp]: "dist (of_complex x :: 'a) (of_complex y) = dist x y" for a :: "'a::complex_normed_div_algebra" by (metis dist_norm norm_of_complex of_complex_diff) declare [[code abort: "open :: complex set \ bool"]] (* As far as I can tell, there is no analogue to this for complex: instance real :: order_topology instance real :: linear_continuum_topology .. lemmas open_complex_greaterThan = open_greaterThan[where 'a=complex] lemmas open_complex_lessThan = open_lessThan[where 'a=complex] lemmas open_complex_greaterThanLessThan = open_greaterThanLessThan[where 'a=complex] *) lemma closed_complex_atMost: \closed {..a::complex}\ proof - have \{..a} = Im -` {Im a} \ Re -` {..Re a}\ by (auto simp: less_eq_complex_def) also have \closed \\ by (auto intro!: closed_Int closed_vimage continuous_on_Im continuous_on_Re) finally show ?thesis by - qed lemma closed_complex_atLeast: \closed {a::complex..}\ proof - have \{a..} = Im -` {Im a} \ Re -` {Re a..}\ by (auto simp: less_eq_complex_def) also have \closed \\ by (auto intro!: closed_Int closed_vimage continuous_on_Im continuous_on_Re) finally show ?thesis by - qed lemma closed_complex_atLeastAtMost: \closed {a::complex .. b}\ proof (cases \Im a = Im b\) case True have \{a..b} = Im -` {Im a} \ Re -` {Re a..Re b}\ by (auto simp add: less_eq_complex_def intro!: True) also have \closed \\ by (auto intro!: closed_Int closed_vimage continuous_on_Im continuous_on_Re) finally show ?thesis by - next case False then have *: \{a..b} = {}\ using less_eq_complex_def by auto show ?thesis - by (simp add: *) + by (simp add: *) qed (* As far as I can tell, there is no analogue to this for complex: instance real :: ordered_real_vector by standard (auto intro: mult_left_mono mult_right_mono) *) (* subsection \Extra type constraints\ *) (* Everything is commented out, so we comment out the heading, too. *) (* These are already configured in Real_Vector_Spaces: text \Only allow \<^term>\open\ in class \topological_space\.\ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ text \Only allow \<^term>\uniformity\ in class \uniform_space\.\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ text \Only allow \<^term>\dist\ in class \metric_space\.\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ text \Only allow \<^term>\norm\ in class \complex_normed_vector\.\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::complex_normed_vector \ real\)\ *) subsection \Sign function\ -(* Inherited from real_normed_vector: +(* Inherited from real_normed_vector: lemma norm_sgn: "norm (sgn x) = (if x = 0 then 0 else 1)" for x :: "'a::complex_normed_vector" lemma sgn_zero [simp]: "sgn (0::'a::complex_normed_vector) = 0" lemma sgn_zero_iff: "sgn x = 0 \ x = 0" for x :: "'a::complex_normed_vector" lemma sgn_minus: "sgn (- x) = - sgn x" for x :: "'a::complex_normed_vector" lemma sgn_one [simp]: "sgn (1::'a::complex_normed_algebra_1) = 1" lemma sgn_mult: "sgn (x * y) = sgn x * sgn y" for x y :: "'a::complex_normed_div_algebra" hide_fact (open) sgn_mult lemma norm_conv_dist: "norm x = dist x 0" declare norm_conv_dist [symmetric, simp] lemma dist_0_norm [simp]: "dist 0 x = norm x" for x :: "'a::complex_normed_vector" lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" lemma dist_of_int: "dist (of_int m) (of_int n :: 'a :: complex_normed_algebra_1) = of_int \m - n\" lemma dist_of_nat: "dist (of_nat m) (of_nat n :: 'a :: complex_normed_algebra_1) = of_int \int m - int n\" *) lemma sgn_scaleC: "sgn (scaleC r x) = scaleC (sgn r) (sgn x)" for x :: "'a::complex_normed_vector" by (simp add: scaleR_scaleC sgn_div_norm ac_simps) lemma sgn_of_complex: "sgn (of_complex r :: 'a::complex_normed_algebra_1) = of_complex (sgn r)" unfolding of_complex_def by (simp only: sgn_scaleC sgn_one) lemma complex_sgn_eq: "sgn x = x / \x\" for x :: complex by (simp add: abs_complex_def scaleR_scaleC sgn_div_norm divide_inverse) lemma czero_le_sgn_iff [simp]: "0 \ sgn x \ 0 \ x" for x :: complex using cmod_eq_Re divide_eq_0_iff less_eq_complex_def by auto lemma csgn_le_0_iff [simp]: "sgn x \ 0 \ x \ 0" for x :: complex by (smt (verit, best) czero_le_sgn_iff Im_sgn Re_sgn divide_eq_0_iff dual_order.eq_iff less_eq_complex_def sgn_zero_iff zero_complex.sel(1) zero_complex.sel(2)) subsection \Bounded Linear and Bilinear Operators\ lemma clinearI: "clinear f" if "\b1 b2. f (b1 + b2) = f b1 + f b2" "\r b. f (r *\<^sub>C b) = r *\<^sub>C f b" using that by unfold_locales (auto simp: algebra_simps) lemma clinear_iff: "clinear f \ (\x y. f (x + y) = f x + f y) \ (\c x. f (c *\<^sub>C x) = c *\<^sub>C f x)" (is "clinear f \ ?rhs") proof assume "clinear f" then interpret f: clinear f . show "?rhs" by (simp add: f.add f.scale complex_vector.linear_scale f.clinear_axioms) next assume "?rhs" then show "clinear f" by (intro clinearI) auto qed lemmas clinear_scaleC_left = complex_vector.linear_scale_left lemmas clinear_imp_scaleC = complex_vector.linear_imp_scale corollary complex_clinearD: fixes f :: "complex \ complex" assumes "clinear f" obtains c where "f = (*) c" by (rule clinear_imp_scaleC [OF assms]) (force simp: scaleC_conv_of_complex) lemma clinear_times_of_complex: "clinear (\x. a * of_complex x)" by (auto intro!: clinearI simp: distrib_left) (metis mult_scaleC_right scaleC_conv_of_complex) locale bounded_clinear = clinear f for f :: "'a::complex_normed_vector \ 'b::complex_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" begin (* Not present in Real_Vector_Spaces *) lemma bounded_linear: "bounded_linear f" apply standard by (simp_all add: add scaleC scaleR_scaleC bounded) lemma pos_bounded: "\K>0. \x. norm (f x) \ norm x * K" proof - obtain K where K: "\x. norm (f x) \ norm x * K" using bounded by blast show ?thesis proof (intro exI impI conjI allI) show "0 < max 1 K" by (rule order_less_le_trans [OF zero_less_one max.cobounded1]) next fix x have "norm (f x) \ norm x * K" using K . also have "\ \ norm x * max 1 K" by (rule mult_left_mono [OF max.cobounded2 norm_ge_zero]) finally show "norm (f x) \ norm x * max 1 K" . qed qed (* Inherited from bounded_linear *) lemma nonneg_bounded: "\K\0. \x. norm (f x) \ norm x * K" by (meson less_imp_le pos_bounded) lemma clinear: "clinear f" by (fact local.clinear_axioms) end lemma bounded_clinear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleC r x) = scaleC r (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_clinear f" by standard (blast intro: assms)+ locale bounded_cbilinear = 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 (scaleC r a) b = scaleC r (prod a b)" and scaleC_right: "prod a (scaleC r b) = scaleC r (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" begin (* Not present in Real_Vector_Spaces *) lemma bounded_bilinear[simp]: "bounded_bilinear prod" apply standard by (auto simp add: add_left add_right scaleR_scaleC scaleC_left scaleC_right bounded) (* Not present in Real_Vector_Spaces. Has only temporary effect (until "end") *) interpretation bounded_bilinear prod by simp lemmas pos_bounded = pos_bounded (* "\K>0. \a b. norm (a ** b) \ norm a * norm b * K" *) lemmas nonneg_bounded = nonneg_bounded (* "\K\0. \a b. norm (a ** b) \ norm a * norm b * K" *) lemmas additive_right = additive_right (* "additive (\b. prod a b)" *) lemmas additive_left = additive_left (* "additive (\a. prod a b)" *) lemmas zero_left = zero_left (* "prod 0 b = 0" *) lemmas zero_right = zero_right (* "prod a 0 = 0" *) lemmas minus_left = minus_left (* "prod (- a) b = - prod a b" *) lemmas minus_right = minus_right (* "prod a (- b) = - prod a b" *) lemmas diff_left = diff_left (* "prod (a - a') b = prod a b - prod a' b" *) lemmas diff_right = diff_right (* "prod a (b - b') = prod a b - prod a b'" *) lemmas sum_left = sum_left (* "prod (sum g S) x = sum ((\i. prod (g i) x)) S" *) lemmas sum_right = sum_right (* "prod x (sum g S) = sum ((\i. (prod x (g i)))) S" *) lemmas prod_diff_prod = prod_diff_prod (* "(x ** y - a ** b) = (x - a) ** (y - b) + (x - a) ** b + a ** (y - b)" *) lemma bounded_clinear_left: "bounded_clinear (\a. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm b * K" in bounded_clinear_intro) (auto simp: algebra_simps scaleC_left add_left) qed lemma bounded_clinear_right: "bounded_clinear (\b. a ** b)" proof - obtain K where "\a b. norm (a ** b) \ norm a * norm b * K" using pos_bounded by blast then show ?thesis by (rule_tac K="norm a * K" in bounded_clinear_intro) (auto simp: algebra_simps scaleC_right add_right) qed lemma flip: "bounded_cbilinear (\x y. y ** x)" proof show "\K. \a b. norm (b ** a) \ norm a * norm b * K" by (metis bounded mult.commute) qed (simp_all add: add_right add_left scaleC_right scaleC_left) lemma comp1: assumes "bounded_clinear g" shows "bounded_cbilinear (\x. (**) (g x))" proof interpret g: bounded_clinear g by fact show "\a a' b. g (a + a') ** b = g a ** b + g a' ** b" "\a b b'. g a ** (b + b') = g a ** b + g a ** b'" "\r a b. g (r *\<^sub>C a) ** b = r *\<^sub>C (g a ** b)" "\a r b. g a ** (r *\<^sub>C b) = r *\<^sub>C (g a ** b)" by (auto simp: g.add add_left add_right g.scaleC scaleC_left scaleC_right) have "bounded_bilinear (\a b. g a ** b)" using g.bounded_linear by (rule comp1) then show "\K. \a b. norm (g a ** b) \ norm a * norm b * K" by (rule bounded_bilinear.bounded) qed lemma comp: "bounded_clinear f \ bounded_clinear g \ bounded_cbilinear (\x y. f x ** g y)" by (rule bounded_cbilinear.flip[OF bounded_cbilinear.comp1[OF bounded_cbilinear.flip[OF comp1]]]) end (* locale bounded_cbilinear *) lemma bounded_clinear_ident[simp]: "bounded_clinear (\x. x)" by standard (auto intro!: exI[of _ 1]) lemma bounded_clinear_zero[simp]: "bounded_clinear (\x. 0)" by standard (auto intro!: exI[of _ 1]) lemma bounded_clinear_add: assumes "bounded_clinear f" and "bounded_clinear g" shows "bounded_clinear (\x. f x + g x)" proof - interpret f: bounded_clinear f by fact interpret g: bounded_clinear g by fact show ?thesis proof from f.bounded obtain Kf where Kf: "norm (f x) \ norm x * Kf" for x by blast from g.bounded obtain Kg where Kg: "norm (g x) \ norm x * Kg" for x by blast show "\K. \x. norm (f x + g x) \ norm x * K" using add_mono[OF Kf Kg] by (intro exI[of _ "Kf + Kg"]) (auto simp: field_simps intro: norm_triangle_ineq order_trans) qed (simp_all add: f.add g.add f.scaleC g.scaleC scaleC_add_right) qed lemma bounded_clinear_minus: assumes "bounded_clinear f" shows "bounded_clinear (\x. - f x)" proof - interpret f: bounded_clinear f by fact show ?thesis by unfold_locales (simp_all add: f.add f.scaleC f.bounded) qed lemma bounded_clinear_sub: "bounded_clinear f \ bounded_clinear g \ bounded_clinear (\x. f x - g x)" using bounded_clinear_add[of f "\x. - g x"] bounded_clinear_minus[of g] by (auto simp: algebra_simps) lemma bounded_clinear_sum: fixes f :: "'i \ 'a::complex_normed_vector \ 'b::complex_normed_vector" shows "(\i. i \ I \ bounded_clinear (f i)) \ bounded_clinear (\x. \i\I. f i x)" by (induct I rule: infinite_finite_induct) (auto intro!: bounded_clinear_add) lemma bounded_clinear_compose: assumes "bounded_clinear f" and "bounded_clinear g" shows "bounded_clinear (\x. f (g x))" proof interpret f: bounded_clinear 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 r (f (g x))" for r x by (simp only: f.scaleC g.scaleC) from f.pos_bounded obtain Kf where f: "\x. norm (f x) \ norm x * Kf" and Kf: "0 < Kf" by blast from g.pos_bounded obtain Kg where g: "\x. norm (g x) \ norm x * Kg" by blast show "\K. \x. norm (f (g x)) \ norm x * K" proof (intro exI allI) fix x have "norm (f (g x)) \ norm (g x) * Kf" using f . also have "\ \ (norm x * Kg) * Kf" using g Kf [THEN order_less_imp_le] by (rule mult_right_mono) also have "(norm x * Kg) * Kf = norm x * (Kg * Kf)" by (rule mult.assoc) finally show "norm (f (g x)) \ norm x * (Kg * Kf)" . qed qed lemma bounded_cbilinear_mult: "bounded_cbilinear ((*) :: 'a \ 'a \ 'a::complex_normed_algebra)" proof (rule bounded_cbilinear.intro) show "\K. \a b::'a. norm (a * b) \ norm a * norm b * K" by (rule_tac x=1 in exI) (simp add: norm_mult_ineq) qed (auto simp: algebra_simps) lemma bounded_clinear_mult_left: "bounded_clinear (\x::'a::complex_normed_algebra. x * y)" using bounded_cbilinear_mult by (rule bounded_cbilinear.bounded_clinear_left) lemma bounded_clinear_mult_right: "bounded_clinear (\y::'a::complex_normed_algebra. x * y)" using bounded_cbilinear_mult by (rule bounded_cbilinear.bounded_clinear_right) lemmas bounded_clinear_mult_const = bounded_clinear_mult_left [THEN bounded_clinear_compose] lemmas bounded_clinear_const_mult = bounded_clinear_mult_right [THEN bounded_clinear_compose] lemma bounded_clinear_divide: "bounded_clinear (\x. x / y)" for y :: "'a::complex_normed_field" unfolding divide_inverse by (rule bounded_clinear_mult_left) lemma bounded_cbilinear_scaleC: "bounded_cbilinear scaleC" proof (rule bounded_cbilinear.intro) obtain K where K: \\a (b::'a). norm b \ norm b * K\ using less_eq_real_def by auto show "\K. \a (b::'a). norm (a *\<^sub>C b) \ norm a * norm b * K" apply (rule exI[where x=K]) using K by (metis norm_scaleC) qed (auto simp: algebra_simps) lemma bounded_clinear_scaleC_left: "bounded_clinear (\c. scaleC c x)" using bounded_cbilinear_scaleC by (rule bounded_cbilinear.bounded_clinear_left) lemma bounded_clinear_scaleC_right: "bounded_clinear (\x. scaleC c x)" using bounded_cbilinear_scaleC by (rule bounded_cbilinear.bounded_clinear_right) lemmas bounded_clinear_scaleC_const = bounded_clinear_scaleC_left[THEN bounded_clinear_compose] lemmas bounded_clinear_const_scaleC = bounded_clinear_scaleC_right[THEN bounded_clinear_compose] lemma bounded_clinear_of_complex: "bounded_clinear (\r. of_complex r)" unfolding of_complex_def by (rule bounded_clinear_scaleC_left) lemma complex_bounded_clinear: "bounded_clinear f \ (\c::complex. f = (\x. x * c))" for f :: "complex \ complex" proof - { fix x assume "bounded_clinear f" then interpret bounded_clinear f . from scaleC[of x 1] have "f x = x * f 1" by simp } then show ?thesis by (auto intro: exI[of _ "f 1"] bounded_clinear_mult_left) qed (* Inherited from real_normed_algebra_1 *) (* instance complex_normed_algebra_1 \ perfect_space *) (* subsection \Filters and Limits on Metric Space\ *) (* Everything is commented out, so we comment out the heading, too. *) (* Not specific to real/complex *) (* lemma (in metric_space) nhds_metric: "nhds x = (INF e\{0 <..}. principal {y. dist y x < e})" *) (* lemma (in metric_space) tendsto_iff: "(f \ l) F \ (\e>0. eventually (\x. dist (f x) l < e) F)" *) (* lemma tendsto_dist_iff: "((f \ l) F) \ (((\x. dist (f x) l) \ 0) F)" *) (* lemma (in metric_space) tendstoI [intro?]: "(\e. 0 < e \ eventually (\x. dist (f x) l < e) F) \ (f \ l) F" *) (* lemma (in metric_space) tendstoD: "(f \ l) F \ 0 < e \ eventually (\x. dist (f x) l < e) F" *) (* lemma (in metric_space) eventually_nhds_metric: "eventually P (nhds a) \ (\d>0. \x. dist x a < d \ P x)" *) (* lemma eventually_at: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" *) (* lemma frequently_at: "frequently P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a < d \ P x)" for a :: "'a :: metric_space" *) (* lemma eventually_at_le: "eventually P (at a within S) \ (\d>0. \x\S. x \ a \ dist x a \ d \ P x)" for a :: "'a::metric_space" *) (* Does not work in complex case because it needs complex :: order_toplogy *) (* lemma eventually_at_left_real: "a > (b :: real) \ eventually (\x. x \ {b<.. eventually (\x. x \ {a<.. a) F" and le: "eventually (\x. dist (g x) b \ dist (f x) a) F" shows "(g \ b) F" *) (* Not sure if this makes sense in the complex case *) (* lemma filterlim_complex_sequentially: "LIM x sequentially. (of_nat x :: complex) :> at_top" *) (* Not specific to real/complex *) (* lemma filterlim_nat_sequentially: "filterlim nat sequentially at_top" *) (* lemma filterlim_floor_sequentially: "filterlim floor at_top at_top" *) (* Not sure if this makes sense in the complex case *) (* lemma filterlim_sequentially_iff_filterlim_real: "filterlim f sequentially F \ filterlim (\x. real (f x)) at_top F" (is "?lhs = ?rhs") *) subsubsection \Limits of Sequences\ (* Not specific to real/complex *) (* lemma lim_sequentially: "X \ L \ (\r>0. \no. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" *) (* lemmas LIMSEQ_def = lim_sequentially (*legacy binding*) *) (* lemma LIMSEQ_iff_nz: "X \ L \ (\r>0. \no>0. \n\no. dist (X n) L < r)" for L :: "'a::metric_space" *) (* lemma metric_LIMSEQ_I: "(\r. 0 < r \ \no. \n\no. dist (X n) L < r) \ X \ L" for L :: "'a::metric_space" *) (* lemma metric_LIMSEQ_D: "X \ L \ 0 < r \ \no. \n\no. dist (X n) L < r" for L :: "'a::metric_space" *) (* lemma LIMSEQ_norm_0: assumes "\n::nat. norm (f n) < 1 / real (Suc n)" shows "f \ 0" *) (* subsubsection \Limits of Functions\ *) (* Everything is commented out, so we comment out the heading, too. *) (* Not specific to real/complex *) (* lemma LIM_def: "f \a\ L \ (\r > 0. \s > 0. \x. x \ a \ dist x a < s \ dist (f x) L < r)" for a :: "'a::metric_space" and L :: "'b::metric_space" *) (* lemma metric_LIM_I: "(\r. 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r) \ f \a\ L" for a :: "'a::metric_space" and L :: "'b::metric_space" *) (* lemma metric_LIM_D: "f \a\ L \ 0 < r \ \s>0. \x. x \ a \ dist x a < s \ dist (f x) L < r" for a :: "'a::metric_space" and L :: "'b::metric_space" *) (* lemma metric_LIM_imp_LIM: fixes l :: "'a::metric_space" and m :: "'b::metric_space" assumes f: "f \a\ l" and le: "\x. x \ a \ dist (g x) m \ dist (f x) l" shows "g \a\ m" *) (* lemma metric_LIM_equal2: fixes a :: "'a::metric_space" assumes "g \a\ l" "0 < R" and "\x. x \ a \ dist x a < R \ f x = g x" shows "f \a\ l" *) (* lemma metric_LIM_compose2: fixes a :: "'a::metric_space" assumes f: "f \a\ b" and g: "g \b\ c" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ b" shows "(\x. g (f x)) \a\ c" *) (* lemma metric_isCont_LIM_compose2: fixes f :: "'a :: metric_space \ _" assumes f [unfolded isCont_def]: "isCont f a" and g: "g \f a\ l" and inj: "\d>0. \x. x \ a \ dist x a < d \ f x \ f a" shows "(\x. g (f x)) \a\ l" *) (* subsection \Complete metric spaces\ *) (* Everything is commented out, so we comment out the heading, too. *) subsection \Cauchy sequences\ (* Not specific to real/complex *) (* lemma (in metric_space) Cauchy_def: "Cauchy X = (\e>0. \M. \m\M. \n\M. dist (X m) (X n) < e)" *) (* lemma (in metric_space) Cauchy_altdef: "Cauchy f \ (\e>0. \M. \m\M. \n>m. dist (f m) (f n) < e)" *) (* lemma (in metric_space) Cauchy_altdef2: "Cauchy s \ (\e>0. \N::nat. \n\N. dist(s n)(s N) < e)" (is "?lhs = ?rhs") *) (* lemma (in metric_space) metric_CauchyI: "(\e. 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e) \ Cauchy X" *) (* lemma (in metric_space) CauchyI': "(\e. 0 < e \ \M. \m\M. \n>m. dist (X m) (X n) < e) \ Cauchy X" *) (* lemma (in metric_space) metric_CauchyD: "Cauchy X \ 0 < e \ \M. \m\M. \n\M. dist (X m) (X n) < e" *) (* lemma (in metric_space) metric_Cauchy_iff2: "Cauchy X = (\j. (\M. \m \ M. \n \ M. dist (X m) (X n) < inverse(real (Suc j))))" *) lemma cCauchy_iff2: "Cauchy X \ (\j. (\M. \m \ M. \n \ M. cmod (X m - X n) < inverse (real (Suc j))))" by (simp only: metric_Cauchy_iff2 dist_complex_def) (* Not specific to real/complex *) (* lemma lim_1_over_n [tendsto_intros]: "((\n. 1 / of_nat n) \ (0::'a::complex_normed_field)) sequentially" *) (* lemma (in metric_space) complete_def: shows "complete S = (\f. (\n. f n \ S) \ Cauchy f \ (\l\S. f \ l))" *) (* lemma (in metric_space) totally_bounded_metric: "totally_bounded S \ (\e>0. \k. finite k \ S \ (\x\k. {y. dist x y < e}))" *) (* subsubsection \Cauchy Sequences are Convergent\ *) (* Everything is commented out, so we comment out the heading, too. *) (* Not specific to real/complex *) (* class complete_space *) (* lemma Cauchy_convergent_iff: "Cauchy X \ convergent X" for X :: "nat \ 'a::complete_space" *) (* text \To prove that a Cauchy sequence converges, it suffices to show that a subsequence converges.\ *) (* Not specific to real/complex *) (* lemma Cauchy_converges_subseq: fixes u::"nat \ 'a::metric_space" assumes "Cauchy u" "strict_mono r" "(u \ r) \ l" shows "u \ l" *) subsection \The set of real numbers is a complete metric space\ text \ Proof that Cauchy sequences converge based on the one from \<^url>\http://pirate.shu.edu/~wachsmut/ira/numseq/proofs/cauconv.html\ \ text \ If sequence \<^term>\X\ is Cauchy, then its limit is the lub of \<^term>\{r::real. \N. \n\N. r < X n}\ \ lemma complex_increasing_LIMSEQ: fixes f :: "nat \ complex" assumes inc: "\n. f n \ f (Suc n)" and bdd: "\n. f n \ l" and en: "\e. 0 < e \ \n. l \ f n + e" shows "f \ l" proof - have \(\n. Re (f n)) \ Re l\ apply (rule increasing_LIMSEQ) using assms apply (auto simp: less_eq_complex_def less_complex_def) by (metis Im_complex_of_real Re_complex_of_real) moreover have \Im (f n) = Im l\ for n using bdd by (auto simp: less_eq_complex_def) then have \(\n. Im (f n)) \ Im l\ by auto ultimately show \f \ l\ by (simp add: tendsto_complex_iff) qed lemma complex_Cauchy_convergent: fixes X :: "nat \ complex" assumes X: "Cauchy X" shows "convergent X" using assms by (rule Cauchy_convergent) instance complex :: complete_space by intro_classes (rule complex_Cauchy_convergent) class cbanach = complex_normed_vector + complete_space (* Not present in Real_Vector_Spaces *) subclass (in cbanach) banach .. instance complex :: banach .. (* Don't know if this holds in the complex case *) (* lemma tendsto_at_topI_sequentially: fixes f :: "complex \ 'b::first_countable_topology" assumes *: "\X. filterlim X at_top sequentially \ (\n. f (X n)) \ y" shows "(f \ y) at_top" *) (* lemma tendsto_at_topI_sequentially_real: fixes f :: "real \ real" assumes mono: "mono f" and limseq: "(\n. f (real n)) \ y" shows "(f \ y) at_top" *) end diff --git a/thys/Complex_Bounded_Operators/ROOT b/thys/Complex_Bounded_Operators/ROOT --- a/thys/Complex_Bounded_Operators/ROOT +++ b/thys/Complex_Bounded_Operators/ROOT @@ -1,18 +1,19 @@ chapter AFP session Complex_Bounded_Operators(AFP) = "HOL-Analysis" + options [timeout = 1200] - sessions - "HOL-Types_To_Sets" - Jordan_Normal_Form + sessions + "HOL-Examples" + "HOL-Types_To_Sets" + Jordan_Normal_Form Banach_Steinhaus Real_Impl - directories + directories extra theories Complex_L2 Cblinfun_Code Cblinfun_Code_Examples - document_files + document_files root.tex root.bib diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Infinite_Set_Sum.thy b/thys/Complex_Bounded_Operators/extra/Extra_Infinite_Set_Sum.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Infinite_Set_Sum.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Infinite_Set_Sum.thy @@ -1,2284 +1,2283 @@ theory Extra_Infinite_Set_Sum imports "HOL-Analysis.Infinite_Set_Sum" Jordan_Normal_Form.Conjugate \ \\<^theory>\Jordan_Normal_Form.Conjugate\ contains the instantiation \complex :: ord\. If we define our own instantiation, it would be impossible to load both \<^session>\Jordan_Normal_Form\ and this theory.\ Extra_General begin subsection\Infinite Set Sum Missing\ definition infsetsum'_converges :: "('a \ 'b::{comm_monoid_add,t2_space}) \ 'a set \ bool" where "infsetsum'_converges f A = (\x. (sum f \ x) (finite_subsets_at_top A))" definition infsetsum' :: "('a \ 'b::{comm_monoid_add,t2_space}) \ 'a set \ 'b" where "infsetsum' f A = (if infsetsum'_converges f A then Lim (finite_subsets_at_top A) (sum f) else 0)" -lemma infsetsum'_converges_cong: +lemma infsetsum'_converges_cong: assumes t1: "\x. x\A \ f x = g x" shows "infsetsum'_converges f A = infsetsum'_converges g A" proof- have "sum f X = sum g X" if "finite X" and "X \ A" for X - by (meson Finite_Cartesian_Product.sum_cong_aux subsetD t1 that(2)) + by (meson Finite_Cartesian_Product.sum_cong_aux subsetD t1 that(2)) hence "\\<^sub>F x in finite_subsets_at_top A. sum f x = sum g x" by (simp add: eventually_finite_subsets_at_top_weakI) hence "(sum f \ x) (finite_subsets_at_top A) = (sum g \ x) (finite_subsets_at_top A)" for x by (simp add: filterlim_cong) thus ?thesis by (simp add: infsetsum'_converges_def) qed lemma infsetsum'_cong: assumes "\x. x\A \ f x = g x" shows "infsetsum' f A = infsetsum' g A" proof- have "sum f X = sum g X" if "finite X" and "X \ A" for X - by (meson Finite_Cartesian_Product.sum_cong_aux assms in_mono that(2)) + by (meson Finite_Cartesian_Product.sum_cong_aux assms in_mono that(2)) hence "\\<^sub>F x in finite_subsets_at_top A. sum f x = sum g x" by (rule eventually_finite_subsets_at_top_weakI) - hence "(sum f \ x) (finite_subsets_at_top A) \ (sum g \ x) (finite_subsets_at_top A)" + hence "(sum f \ x) (finite_subsets_at_top A) \ (sum g \ x) (finite_subsets_at_top A)" for x by (rule tendsto_cong) hence "Lim (finite_subsets_at_top A) (sum f) = Lim (finite_subsets_at_top A) (sum g)" unfolding Topological_Spaces.Lim_def[abs_def] by auto thus ?thesis unfolding infsetsum'_def using assms infsetsum'_converges_cong by auto qed lemma abs_summable_finiteI0: assumes "\F. finite F \ F\S \ sum (\x. norm (f x)) F \ B" and "B \ 0" shows "f abs_summable_on S" and "infsetsum (\x. norm (f x)) S \ B" proof- have t1: "f abs_summable_on S \ infsetsum (\x. norm (f x)) S \ B" proof(cases "S = {}") case True thus ?thesis - by (simp add: assms(2)) + by (simp add: assms(2)) next case False define M normf where "M = count_space S" and "normf x = ennreal (norm (f x))" for x have "sum normf F \ ennreal B" if "finite F" and "F \ S" and "\F. finite F \ F \ S \ (\i\F. ennreal (norm (f i))) \ ennreal B" and "ennreal 0 \ ennreal B" for F - using that unfolding normf_def[symmetric] by simp + using that unfolding normf_def[symmetric] by simp hence normf_B: "finite F \ F\S \ sum normf F \ ennreal B" for F - using assms[THEN ennreal_leI] + using assms[THEN ennreal_leI] by auto - have "integral\<^sup>S M g \ B" if "simple_function M g" and "g \ normf" for g + have "integral\<^sup>S M g \ B" if "simple_function M g" and "g \ normf" for g proof - define gS where "gS = g ` S" have "finite gS" using that unfolding gS_def M_def simple_function_count_space by simp have "gS \ {}" unfolding gS_def False - by (simp add: False) + by (simp add: False) define part where "part r = g -` {r} \ S" for r - have r_finite: "r < \" if "r : gS" for r + have r_finite: "r < \" if "r : gS" for r using \g \ normf\ that unfolding gS_def le_fun_def normf_def apply auto using ennreal_less_top neq_top_trans top.not_eq_extremum by blast define B' where "B' r = (SUP F\{F. finite F \ F\part r}. sum normf F)" for r have B'fin: "B' r < \" for r proof - have "B' r \ (SUP F\{F. finite F \ F\part r}. sum normf F)" unfolding B'_def by (metis (mono_tags, lifting) SUP_least SUP_upper) also have "\ \ B" using normf_B unfolding part_def by (metis (no_types, lifting) Int_subset_iff SUP_least mem_Collect_eq) also have "\ < \" by simp finally show ?thesis by simp qed have sumB': "sum B' gS \ ennreal B + \" if "\>0" for \ proof - define N \N where "N = card gS" and "\N = \ / N" - have "N > 0" + have "N > 0" unfolding N_def using \gS\{}\ \finite gS\ by (simp add: card_gt_0_iff) from \N_def that have "\N > 0" by (simp add: ennreal_of_nat_eq_real_of_nat ennreal_zero_less_divide) have c1: "\y. B' r \ sum normf y + \N \ finite y \ y \ part r" if "B' r = 0" for r using that by auto have c2: "\y. B' r \ sum normf y + \N \ finite y \ y \ part r" if "B' r \ 0" for r proof- have "B' r - \N < B' r" using B'fin \0 < \N\ ennreal_between that by fastforce have "B' r - \N < Sup (sum normf ` {F. finite F \ F \ part r}) \ \F. B' r - \N \ sum normf F \ finite F \ F \ part r" by (metis (no_types, lifting) leD le_cases less_SUP_iff mem_Collect_eq) hence "B' r - \N < B' r \ \F. B' r - \N \ sum normf F \ finite F \ F \ part r" by (subst (asm) (2) B'_def) then obtain F where "B' r - \N \ sum normf F" and "finite F" and "F \ part r" - using \B' r - \N < B' r\ by auto + using \B' r - \N < B' r\ by auto thus "\F. B' r \ sum normf F + \N \ finite F \ F \ part r" by (metis add.commute ennreal_minus_le_iff) qed have "\x. \y. B' x \ sum normf y + \N \ finite y \ y \ part x" using c1 c2 - by blast + by blast hence "\F. \x. B' x \ sum normf (F x) + \N \ finite (F x) \ F x \ part x" - by metis + by metis then obtain F where F: "sum normf (F r) + \N \ B' r" and Ffin: "finite (F r)" and Fpartr: "F r \ part r" for r using atomize_elim by auto have w1: "finite gS" - by (simp add: \finite gS\) + by (simp add: \finite gS\) have w2: "\i\gS. finite (F i)" - by (simp add: Ffin) + by (simp add: Ffin) have False if "\r. F r \ g -` {r} \ F r \ S" and "i \ gS" and "j \ gS" and "i \ j" and "x \ F i" and "x \ F j" for i j x - by (metis subsetD that(1) that(4) that(5) that(6) vimage_singleton_eq) + by (metis subsetD that(1) that(4) that(5) that(6) vimage_singleton_eq) hence w3: "\i\gS. \j\gS. i \ j \ F i \ F j = {}" - using Fpartr[unfolded part_def] by auto + using Fpartr[unfolded part_def] by auto have w4: "sum normf (\ (F ` gS)) + \ = sum normf (\ (F ` gS)) + \" by simp have "sum B' gS \ (\r\gS. sum normf (F r) + \N)" using F by (simp add: sum_mono) also have "\ = (\r\gS. sum normf (F r)) + (\r\gS. \N)" by (simp add: sum.distrib) also have "\ = (\r\gS. sum normf (F r)) + (card gS * \N)" by auto also have "\ = (\r\gS. sum normf (F r)) + \" - unfolding \N_def N_def[symmetric] using \N>0\ + unfolding \N_def N_def[symmetric] using \N>0\ by (simp add: ennreal_times_divide mult.commute mult_divide_eq_ennreal) - also have "\ = sum normf (\ (F ` gS)) + \" + also have "\ = sum normf (\ (F ` gS)) + \" using w1 w2 w3 w4 by (subst sum.UNION_disjoint[symmetric]) also have "\ \ B + \" using \finite gS\ normf_B add_right_mono Ffin Fpartr unfolding part_def - by (simp add: \gS \ {}\ cSUP_least) + by (simp add: \gS \ {}\ cSUP_least) finally show ?thesis by auto qed hence sumB': "sum B' gS \ B" using ennreal_le_epsilon ennreal_less_zero_iff by blast have "\r. \y. r \ gS \ B' r = ennreal y" using B'fin less_top_ennreal by auto hence "\B''. \r. r \ gS \ B' r = ennreal (B'' r)" - by (rule_tac choice) + by (rule_tac choice) then obtain B'' where B'': "B' r = ennreal (B'' r)" if "r \ gS" for r - by atomize_elim + by atomize_elim have cases[case_names zero finite infinite]: "P" if "r=0 \ P" and "finite (part r) \ P" and "infinite (part r) \ r\0 \ P" for P r using that by metis have emeasure_B': "r * emeasure M (part r) \ B' r" if "r : gS" for r proof (cases rule:cases[of r]) case zero thus ?thesis by simp next case finite have s1: "sum g F \ sum normf F" if "F \ {F. finite F \ F \ part r}" for F - using \g \ normf\ + using \g \ normf\ by (simp add: le_fun_def sum_mono) have "r * of_nat (card (part r)) = r * (\x\part r. 1)" by simp also have "\ = (\x\part r. r)" using mult.commute by auto also have "\ = (\x\part r. g x)" unfolding part_def by auto also have "\ \ (SUP F\{F. finite F \ F\part r}. sum g F)" using finite by (simp add: Sup_upper) - also have "\ \ B' r" + also have "\ \ B' r" unfolding B'_def using s1 SUP_subset_mono by blast finally have "r * of_nat (card (part r)) \ B' r" by assumption thus ?thesis unfolding M_def using part_def finite by auto next case infinite from r_finite[OF \r : gS\] obtain r' where r': "r = ennreal r'" using ennreal_cases by auto with infinite have "r' > 0" using ennreal_less_zero_iff not_gr_zero by blast obtain N::nat where N:"N > B / r'" and "real N > 0" apply atomize_elim using reals_Archimedean2 by (metis less_trans linorder_neqE_linordered_idom) obtain F where "finite F" and "card F = N" and "F \ part r" using infinite(1) infinite_arbitrarily_large by blast from \F \ part r\ have "F \ S" unfolding part_def by simp have "B < r * N" unfolding r' ennreal_of_nat_eq_real_of_nat using N \0 < r'\ assms(2) r' by (metis enn2real_ennreal enn2real_less_iff ennreal_less_top ennreal_mult' less_le mult_less_cancel_left_pos nonzero_mult_div_cancel_left times_divide_eq_right) also have "r * N = (\x\F. r)" using \card F = N\ by (simp add: mult.commute) also have "(\x\F. r) = (\x\F. g x)" using \F \ part r\ part_def sum.cong subsetD by fastforce also have "(\x\F. g x) \ (\x\F. ennreal (norm (f x)))" - by (metis (mono_tags, lifting) \g \ normf\ \normf \ \x. ennreal (norm (f x))\ le_fun_def + by (metis (mono_tags, lifting) \g \ normf\ \normf \ \x. ennreal (norm (f x))\ le_fun_def sum_mono) also have "(\x\F. ennreal (norm (f x))) \ B" - using \F \ S\ \finite F\ \normf \ \x. ennreal (norm (f x))\ normf_B by blast + using \F \ S\ \finite F\ \normf \ \x. ennreal (norm (f x))\ normf_B by blast finally have "B < B" by auto thus ?thesis by simp qed have "integral\<^sup>S M g = (\r \ gS. r * emeasure M (part r))" unfolding simple_integral_def gS_def M_def part_def by simp also have "\ \ (\r \ gS. B' r)" by (simp add: emeasure_B' sum_mono) also have "\ \ B" using sumB' by blast finally show ?thesis by assumption qed hence int_leq_B: "integral\<^sup>N M normf \ B" unfolding nn_integral_def by (metis (no_types, lifting) SUP_least mem_Collect_eq) hence "integral\<^sup>N M normf < \" using le_less_trans by fastforce hence "integrable M f" unfolding M_def normf_def by (rule integrableI_bounded[rotated], simp) hence v1: "f abs_summable_on S" - unfolding abs_summable_on_def M_def by simp + unfolding abs_summable_on_def M_def by simp have "(\x. norm (f x)) abs_summable_on S" using v1 Infinite_Set_Sum.abs_summable_on_norm_iff[where A = S and f = f] by auto moreover have "0 \ norm (f x)" if "x \ S" for x - by simp + by simp moreover have "(\\<^sup>+ x. ennreal (norm (f x)) \count_space S) \ ennreal B" - using M_def \normf \ \x. ennreal (norm (f x))\ int_leq_B by auto + using M_def \normf \ \x. ennreal (norm (f x))\ int_leq_B by auto ultimately have "ennreal (\\<^sub>ax\S. norm (f x)) \ ennreal B" - by (simp add: nn_integral_conv_infsetsum) + by (simp add: nn_integral_conv_infsetsum) hence v2: "(\\<^sub>ax\S. norm (f x)) \ B" by (subst ennreal_le_iff[symmetric], simp add: assms) show ?thesis using v1 v2 by auto qed show "f abs_summable_on S" using t1 by blast show "(\\<^sub>ax\S. norm (f x)) \ B" using t1 by blast qed lemma abs_summable_finiteI: assumes "\F. finite F \ F\S \ sum (\x. norm (f x)) F \ B" shows "f abs_summable_on S" proof - from assms have "sum (\x. norm (f x)) {} \ B" by blast hence "0 \ B" by simp - thus ?thesis + thus ?thesis using assms by (rule abs_summable_finiteI0[rotated]) qed lemma infsetsum_finite_sets: assumes "\F. finite F \ F\S \ sum (\x. norm (f x)) F \ B" and "B \ 0" and "\x. f x \ 0" shows "infsetsum f S \ B" using abs_summable_finiteI0(2)[where f=f and S=S, OF assms(1-2), simplified] assms(3) by auto lemma abs_summable_finiteI_converse: assumes f_sum_S: "f abs_summable_on S" and finite_F: "finite F" and FS: "F\S" defines "B \ (infsetsum (\x. norm (f x)) S)" shows "sum (\x. norm (f x)) F \ B" proof- have a1: "(\x. norm (f x)) abs_summable_on F" - by (simp add: finite_F) + by (simp add: finite_F) have a2: "(\x. norm (f x)) abs_summable_on S" - by (simp add: f_sum_S) + by (simp add: f_sum_S) have a3: "x \ F \ norm (f x) \ norm (f x)" for x by simp have a4: "F \ S" - by (simp add: FS) + by (simp add: FS) have a5: "x \ S - F \ 0 \ norm (f x)" for x - by simp + by simp have "sum (\x. norm (f x)) F = infsetsum (\x. norm (f x)) F" - by (simp add: finite_F) + by (simp add: finite_F) also have "infsetsum (\x. norm (f x)) F \ B" - unfolding B_def - using a1 a2 a3 a4 a5 + unfolding B_def + using a1 a2 a3 a4 a5 by (simp add: infsetsum_mono_neutral_left) finally show ?thesis by assumption qed lemma abs_summable_countable: fixes \ :: "'a \ 'b::{banach,second_countable_topology}" assumes "\ abs_summable_on T" shows "countable {x\T. 0 \ \ x}" proof- define B where "B = infsetsum (\x. norm (\ x)) T" have \sum: "sum (\x. norm (\ x)) F \ B" if "finite F" and "F \ T" for F - unfolding B_def + unfolding B_def using assms that abs_summable_finiteI_converse by auto define S where "S i = {x\T. 1/real (Suc i) < norm (\ x)}" for i::nat have "\i. x \ S i" if "0 < norm (\ x)" and "x \ T" for x using that unfolding S_def - by (metis (full_types, lifting) mem_Collect_eq nat_approx_posE) + by (metis (full_types, lifting) mem_Collect_eq nat_approx_posE) hence union: "{x\T. 0 < norm (\ x)} = (\i. S i)" unfolding S_def by auto have finiteS: "finite (S i)" for i proof (rule ccontr) assume "infinite (S i)" then obtain F where F_Si: "F \ S i" and cardF: "card F > (Suc i)*B" and finiteF: "finite F" by (metis One_nat_def ex_less_of_nat_mult infinite_arbitrarily_large lessI mult.commute mult.left_neutral of_nat_0_less_iff of_nat_1) - from F_Si have F_T: "F \ T" + from F_Si have F_T: "F \ T" unfolding S_def by blast from finiteF \sum F_T have sumF: "sum (\x. norm (\ x)) F \ B" by simp have "1 / real (Suc i) \ norm (\ x)" if "x \ F" for x using that F_Si S_def by auto hence "sum (\x. norm (\ x)) F \ sum (\_. 1/real (Suc i)) F" (is "_ \ \") using sum_mono - by metis + by metis moreover have "\ = real (card F) / (Suc i)" by (subst sum_constant_scale, auto) moreover have "\ > B" using cardF by (simp add: linordered_field_class.mult_imp_less_div_pos algebra_simps) ultimately have "sum (\x. norm (\ x)) F > B" - by linarith + by linarith with sumF show False by simp qed have "countable (S i)" if "i \ UNIV" for i using finiteS by (simp add: countable_finite) hence "countable (\i. S i)" using countable_UN by simp hence "countable {x\T. 0 < norm (\ x)}" unfolding union by simp thus ?thesis by simp qed lemma infsetsum_cnj[simp]: "infsetsum (\x. cnj (f x)) M = cnj (infsetsum f M)" unfolding infsetsum_def by (rule integral_cnj) -lemma infsetsum_Re: +lemma infsetsum_Re: assumes "f abs_summable_on M" shows "infsetsum (\x. Re (f x)) M = Re (infsetsum f M)" - unfolding infsetsum_def + unfolding infsetsum_def using integral_Re assms by (simp add: abs_summable_on_def) -lemma infsetsum_Im: +lemma infsetsum_Im: assumes "f abs_summable_on M" shows "infsetsum (\x. Im (f x)) M = Im (infsetsum f M)" unfolding infsetsum_def using assms by (simp add: abs_summable_on_def) lemma infsetsum_mono_complex: fixes f g :: "'a \ complex" assumes f_sum: "f abs_summable_on A" and g_sum: "g abs_summable_on A" assumes x: "\x. x \ A \ f x \ g x" shows "infsetsum f A \ infsetsum g A" proof - have a1: "infsetsum f A = Complex (Re (infsetsum f A)) (Im (infsetsum f A))" by auto also have a2: "Re (infsetsum f A) = infsetsum (\x. Re (f x)) A" - unfolding infsetsum_def + unfolding infsetsum_def using assms by (simp add: abs_summable_on_def) also have a3: "Im (infsetsum f A) = infsetsum (\x. Im (f x)) A" using f_sum by (rule infsetsum_Im[symmetric]) finally have fsplit: "infsetsum f A = Complex (\\<^sub>ax\A. Re (f x)) (\\<^sub>ax\A. Im (f x))" by assumption have "infsetsum g A = Complex (Re (infsetsum g A)) (Im (infsetsum g A))" by auto also have b2: "Re (infsetsum g A) = infsetsum (\x. Re (g x)) A" using g_sum by (rule infsetsum_Re[symmetric]) also have b1: "Im (infsetsum g A) = infsetsum (\x. Im (g x)) A " using g_sum by (rule infsetsum_Im[symmetric]) - finally have gsplit: "infsetsum g A = Complex (\\<^sub>ax\A. Re (g x)) (\\<^sub>ax\A. Im (g x))" + finally have gsplit: "infsetsum g A = Complex (\\<^sub>ax\A. Re (g x)) (\\<^sub>ax\A. Im (g x))" by assumption have Re_leq: "Re (f x) \ Re (g x)" if "x\A" for x using that assms unfolding less_eq_complex_def by simp have Im_eq: "Im (f x) = Im (g x)" if "x\A" for x - using that assms + using that assms unfolding less_eq_complex_def by simp have Refsum: "(\x. Re (f x)) abs_summable_on A" using assms(1) abs_Re_le_cmod by (simp add: abs_summable_on_comparison_test[where g=f]) have Regsum: "(\x. Re (g x)) abs_summable_on A" - using assms(2) abs_Re_le_cmod + using assms(2) abs_Re_le_cmod by (simp add: abs_summable_on_comparison_test[where g=g]) show "infsetsum f A \ infsetsum g A" unfolding fsplit gsplit by (smt (verit, ccfv_SIG) Im_eq Re_leq Refsum Regsum a2 a3 b1 b2 fsplit gsplit infsetsum_cong infsetsum_mono less_eq_complex_def) qed lemma infsetsum_subset_complex: fixes f :: "'a \ complex" assumes "f abs_summable_on B" and "A \ B" and "\x. x\A \ f x \ 0" shows "infsetsum f A \ infsetsum f B" proof - have fBA: "f abs_summable_on B - A" by (meson Diff_subset abs_summable_on_subset assms(1)) have "0 = infsetsum (\_.0) (B-A)" by auto also have "... \ infsetsum f (B - A)" using assms fBA infsetsum_mono_complex by (metis DiffD2 abs_summable_on_0) also have "... = infsetsum f B - infsetsum f A" using assms by (simp add: infsetsum_Diff) finally show ?thesis by auto qed lemma infsetsum_subset_real: fixes f :: "'a \ real" assumes "f abs_summable_on B" and "A \ B" and "\x. x\A \ f x \ 0" shows "infsetsum f A \ infsetsum f B" proof - have fBA: "f abs_summable_on B - A" by (meson Diff_subset abs_summable_on_subset assms(1)) have "0 = infsetsum (\_.0) (B-A)" by auto also have "... \ infsetsum f (B - A)" - using assms fBA - by (metis DiffD2 calculation infsetsum_nonneg) + using assms fBA + by (metis DiffD2 calculation infsetsum_nonneg) also have "... = infsetsum f B - infsetsum f A" using assms by (simp add: infsetsum_Diff) finally show ?thesis by auto qed lemma abs_summable_product: fixes x :: "'a \ 'b::{real_normed_div_algebra,banach,second_countable_topology}" assumes x2_sum: "(\i. (x i) * (x i)) abs_summable_on A" and y2_sum: "(\i. (y i) * (y i)) abs_summable_on A" shows "(\i. x i * y i) abs_summable_on A" proof (rule abs_summable_finiteI) have aux: "a\a' \ b\b' \ a+b \ a'+b'" for a b a' b' :: real by simp fix F assume r1: "finite F" and b4: "F \ A" define B :: real where "B = (\\<^sub>ai\A. norm (x i * x i)) + (\\<^sub>ai\A. norm (y i * y i))" have a1: "(\\<^sub>ai\F. norm (x i * x i)) \ (\\<^sub>ai\A. norm (x i * x i))" proof (rule infsetsum_mono_neutral_left) show "(\i. norm (x i * x i)) abs_summable_on F" - by (simp add: r1) + by (simp add: r1) show "(\i. norm (x i * x i)) abs_summable_on A" - by (simp add: x2_sum) + by (simp add: x2_sum) show "norm (x i * x i) \ norm (x i * x i)" if "i \ F" for i :: 'a by simp show "F \ A" - by (simp add: b4) + by (simp add: b4) show "0 \ norm (x i * x i)" if "i \ A - F" for i :: 'a - by simp + by simp qed have "norm (x i * y i) \ norm (x i * x i) + norm (y i * y i)" for i unfolding norm_mult by (smt mult_left_mono mult_nonneg_nonneg mult_right_mono norm_ge_zero) hence "(\i\F. norm (x i * y i)) \ (\i\F. norm (x i * x i) + norm (y i * y i))" by (simp add: sum_mono) also have "\ = (\i\F. norm (x i * x i)) + (\i\F. norm (y i * y i))" by (simp add: sum.distrib) also have "\ = (\\<^sub>ai\F. norm (x i * x i)) + (\\<^sub>ai\F. norm (y i * y i))" by (simp add: \finite F\) - also have "\ \ (\\<^sub>ai\A. norm (x i * x i)) + (\\<^sub>ai\A. norm (y i * y i))" + also have "\ \ (\\<^sub>ai\A. norm (x i * x i)) + (\\<^sub>ai\A. norm (y i * y i))" using aux a1 by (simp add: aux \F \ A\ \finite F\ abs_summable_finiteI_converse x2_sum y2_sum) also have "\ = B" unfolding B_def by simp finally show "(\i\F. norm (x i * y i)) \ B" by assumption qed lemma abs_summable_cnj_iff[simp]: "(\i. cnj (f i)) abs_summable_on A \ f abs_summable_on A" proof show "f abs_summable_on A" if "(\i. cnj (f i)) abs_summable_on A" using that abs_summable_on_norm_iff[symmetric] - abs_summable_on_comparison_test by fastforce + abs_summable_on_comparison_test by fastforce show "(\i. cnj (f i)) abs_summable_on A" if "f abs_summable_on A" using that abs_summable_on_norm_iff[symmetric] - abs_summable_on_comparison_test by fastforce + abs_summable_on_comparison_test by fastforce qed lemma ennreal_Sup: assumes "bdd_above A" and nonempty: "A\{}" shows "ennreal (Sup A) = Sup (ennreal ` A)" proof (rule Sup_eqI[symmetric]) fix y assume "y \ ennreal ` A" thus "y \ ennreal (Sup A)" using assms cSup_upper ennreal_leI by auto next fix y assume asm: "\z. z \ ennreal ` A \ z \ y" show "ennreal (Sup A) \ y" proof (cases y) case (real r) - show ?thesis - by (metis assms(1) cSup_le_iff ennreal_leI real(1) real(2) asm Sup_least bdd_above_top + show ?thesis + by (metis assms(1) cSup_le_iff ennreal_leI real(1) real(2) asm Sup_least bdd_above_top cSUP_le_iff ennreal_le_iff nonempty) next case top thus ?thesis by auto qed qed lemma infsetsum_nonneg_is_SUPREMUM_ennreal: fixes f :: "'a \ real" assumes summable: "f abs_summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ennreal (infsetsum f A) = (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))" proof- - have sum_F_A: "sum f F \ infsetsum f A" - if "F \ {F. finite F \ F \ A}" + have sum_F_A: "sum f F \ infsetsum f A" + if "F \ {F. finite F \ F \ A}" for F proof- from that have "finite F" and "F \ A" by auto from \finite F\ have "sum f F = infsetsum f F" by auto also have "\ \ infsetsum f A" proof (rule infsetsum_mono_neutral_left) show "f abs_summable_on F" - by (simp add: \finite F\) + by (simp add: \finite F\) show "f abs_summable_on A" - by (simp add: local.summable) + by (simp add: local.summable) show "f x \ f x" if "x \ F" for x :: 'a - by simp + by simp show "F \ A" - by (simp add: \F \ A\) + by (simp add: \F \ A\) show "0 \ f x" if "x \ A - F" for x :: 'a - using that fnn by auto + using that fnn by auto qed finally show ?thesis by assumption - qed + qed hence geq: "ennreal (infsetsum f A) \ (SUP F\{G. finite G \ G \ A}. (ennreal (sum f F)))" by (meson SUP_least ennreal_leI) define fe where "fe x = ennreal (f x)" for x have sum_f_int: "infsetsum f A = \\<^sup>+ x. fe x \(count_space A)" unfolding infsetsum_def fe_def proof (rule nn_integral_eq_integral [symmetric]) show "integrable (count_space A) f" - using abs_summable_on_def local.summable by blast + using abs_summable_on_def local.summable by blast show "AE x in count_space A. 0 \ f x" - using fnn by auto + using fnn by auto qed also have "\ = (SUP g \ {g. finite (g`A) \ g \ fe}. integral\<^sup>S (count_space A) g)" unfolding nn_integral_def simple_function_count_space by simp also have "\ \ (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))" proof (rule Sup_least) fix x assume "x \ integral\<^sup>S (count_space A) ` {g. finite (g ` A) \ g \ fe}" - then obtain g where xg: "x = integral\<^sup>S (count_space A) g" and fin_gA: "finite (g`A)" + then obtain g where xg: "x = integral\<^sup>S (count_space A) g" and fin_gA: "finite (g`A)" and g_fe: "g \ fe" by auto define F where "F = {z:A. g z \ 0}" hence "F \ A" by simp have fin: "finite {z:A. g z = t}" if "t \ 0" for t proof (rule ccontr) assume inf: "infinite {z:A. g z = t}" hence tgA: "t \ g ` A" by (metis (mono_tags, lifting) image_eqI not_finite_existsD) have "x = (\x \ g ` A. x * emeasure (count_space A) (g -` {x} \ A))" unfolding xg simple_integral_def space_count_space by simp also have "\ \ (\x \ {t}. x * emeasure (count_space A) (g -` {x} \ A))" (is "_ \ \") proof (rule sum_mono2) show "finite (g ` A)" - by (simp add: fin_gA) + by (simp add: fin_gA) show "{t} \ g ` A" - by (simp add: tgA) + by (simp add: tgA) show "0 \ b * emeasure (count_space A) (g -` {b} \ A)" if "b \ g ` A - {t}" for b :: ennreal using that by simp qed also have "\ = t * emeasure (count_space A) (g -` {t} \ A)" by auto also have "\ = t * \" proof (subst emeasure_count_space_infinite) show "g -` {t} \ A \ A" - by simp + by simp have "{a \ A. g a = t} = {a \ g -` {t}. a \ A}" by auto thus "infinite (g -` {t} \ A)" - by (metis (full_types) Int_def inf) + by (metis (full_types) Int_def inf) show "t * \ = t * \" by simp qed also have "\ = \" using \t \ 0\ by (simp add: ennreal_mult_eq_top_iff) finally have x_inf: "x = \" using neq_top_trans by auto have "x = integral\<^sup>S (count_space A) g" by (fact xg) also have "\ = integral\<^sup>N (count_space A) g" by (simp add: fin_gA nn_integral_eq_simple_integral) also have "\ \ integral\<^sup>N (count_space A) fe" using g_fe by (simp add: le_funD nn_integral_mono) also have "\ < \" - by (metis sum_f_int ennreal_less_top infinity_ennreal_def) + by (metis sum_f_int ennreal_less_top infinity_ennreal_def) finally have x_fin: "x < \" by simp from x_inf x_fin show False by simp qed have F: "F = (\t\g`A-{0}. {z\A. g z = t})" unfolding F_def by auto hence "finite F" unfolding F using fin_gA fin by auto have "x = integral\<^sup>N (count_space A) g" unfolding xg by (simp add: fin_gA nn_integral_eq_simple_integral) also have "\ = set_nn_integral (count_space UNIV) A g" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) also have "\ = set_nn_integral (count_space UNIV) F g" proof - have "\a. g a * (if a \ {a \ A. g a \ 0} then 1 else 0) = g a * (if a \ A then 1 else 0)" by auto hence "(\\<^sup>+ a. g a * (if a \ A then 1 else 0) \count_space UNIV) = (\\<^sup>+ a. g a * (if a \ {a \ A. g a \ 0} then 1 else 0) \count_space UNIV)" by presburger thus ?thesis unfolding F_def indicator_def - using mult.right_neutral mult_zero_right nn_integral_cong - by blast + by (simp add: of_bool_def) qed also have "\ = integral\<^sup>N (count_space F) g" by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) - also have "\ = sum g F" + also have "\ = sum g F" using \finite F\ by (rule nn_integral_count_space_finite) also have "sum g F \ sum fe F" using g_fe unfolding le_fun_def - by (simp add: sum_mono) + by (simp add: sum_mono) also have "\ \ (SUP F \ {G. finite G \ G \ A}. (sum fe F))" using \finite F\ \F\A\ by (simp add: SUP_upper) also have "\ = (SUP F \ {F. finite F \ F \ A}. (ennreal (sum f F)))" proof (rule SUP_cong [OF refl]) have "finite x \ x \ A \ (\x\x. ennreal (f x)) = ennreal (sum f x)" for x by (metis fnn subsetCE sum_ennreal) thus "sum fe x = ennreal (sum f x)" if "x \ {G. finite G \ G \ A}" for x :: "'a set" - using that unfolding fe_def by auto - qed + using that unfolding fe_def by auto + qed finally show "x \ \" by simp qed finally have leq: "ennreal (infsetsum f A) \ (SUP F\{F. finite F \ F \ A}. (ennreal (sum f F)))" by assumption from leq geq show ?thesis by simp qed lemma infsetsum_nonneg_is_SUPREMUM_ereal: fixes f :: "'a \ real" assumes summable: "f abs_summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "ereal (infsetsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" proof - have "ereal (infsetsum f A) = enn2ereal (ennreal (infsetsum f A))" by (simp add: fnn infsetsum_nonneg) also have "\ = enn2ereal (SUP F\{F. finite F \ F \ A}. ennreal (sum f F))" proof (subst infsetsum_nonneg_is_SUPREMUM_ennreal) show "f abs_summable_on A" - by (simp add: local.summable) + by (simp add: local.summable) show "0 \ f x" if "x \ A" for x :: 'a using that - by (simp add: fnn) + by (simp add: fnn) show "enn2ereal (SUP F\{F. finite F \ F \ A}. ennreal (sum f F)) = enn2ereal (SUP F\{F. finite F \ F \ A}. ennreal (sum f F))" - by simp - qed + by simp + qed also have "\ = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" proof (simp add: image_def Sup_ennreal.rep_eq) have "0 \ Sup {y. \x. (\xa. finite xa \ xa \ A \ x = ennreal (sum f xa)) \ y = enn2ereal x}" by (metis (mono_tags, lifting) Sup_upper empty_subsetI ennreal_0 finite.emptyI mem_Collect_eq sum.empty zero_ennreal.rep_eq) moreover have "Sup {y. \x. (\y. finite y \ y \ A \ x = ennreal (sum f y)) \ y = enn2ereal x} = Sup {y. \x. finite x \ x \ A \ y = ereal (sum f x)}" using enn2ereal_ennreal fnn in_mono sum_nonneg Collect_cong by smt ultimately show "max 0 (Sup {y. \x. (\xa. finite xa \ xa \ A \ x = ennreal (sum f xa)) \ y = enn2ereal x}) = Sup {y. \x. finite x \ x \ A \ y = ereal (sum f x)}" by linarith - qed + qed finally show ?thesis by simp qed lemma infsetsum_nonneg_is_SUPREMUM: fixes f :: "'a \ real" assumes summable: "f abs_summable_on A" and fnn: "\x. x\A \ f x \ 0" shows "infsetsum f A = (SUP F\{F. finite F \ F \ A}. (sum f F))" proof - have "ereal (infsetsum f A) = (SUP F\{F. finite F \ F \ A}. (ereal (sum f F)))" using assms by (rule infsetsum_nonneg_is_SUPREMUM_ereal) also have "\ = ereal (SUP F\{F. finite F \ F \ A}. (sum f F))" proof (subst ereal_SUP) show "\SUP a\{F. finite F \ F \ A}. ereal (sum f a)\ \ \" - using calculation by fastforce + using calculation by fastforce show "(SUP F\{F. finite F \ F \ A}. ereal (sum f F)) = (SUP a\{F. finite F \ F \ A}. ereal (sum f a))" - by simp + by simp qed finally show ?thesis by simp qed lemma infsetsum_geq0_complex: fixes f :: "'a \ complex" assumes "f abs_summable_on M" and fnn: "\x. x \ M \ 0 \ f x" shows "infsetsum f M \ 0" (is "?lhs \ _") proof - have "?lhs \ infsetsum (\x. 0) M" (is "_ \ \") proof (rule infsetsum_mono_complex) show "(\x. 0::complex) abs_summable_on M" - by simp + by simp show "f abs_summable_on M" - by (simp add: assms(1)) + by (simp add: assms(1)) show "0 \ f x" if "x \ M" for x :: 'a using that using fnn by blast qed also have "\ = 0" by auto finally show ?thesis by assumption qed lemma infsetsum_cmod: assumes "f abs_summable_on M" and fnn: "\x. x \ M \ 0 \ f x" shows "infsetsum (\x. cmod (f x)) M = cmod (infsetsum f M)" proof - - have nn: "infsetsum f M \ 0" - using assms by (rule infsetsum_geq0_complex) + have nn: "infsetsum f M \ 0" + using assms by (rule infsetsum_geq0_complex) have "infsetsum (\x. cmod (f x)) M = infsetsum (\x. Re (f x)) M" using fnn cmod_eq_Re less_eq_complex_def by auto also have "\ = Re (infsetsum f M)" using assms(1) infsetsum_Re by blast also have "\ = cmod (infsetsum f M)" using nn cmod_eq_Re less_eq_complex_def by auto finally show ?thesis by assumption qed lemma infsetsum_Sigma: fixes A :: "'a set" and B :: "'a \ 'b set" assumes summable: "f abs_summable_on (Sigma A B)" shows "infsetsum f (Sigma A B) = infsetsum (\x. infsetsum (\y. f (x, y)) (B x)) A" proof- from summable have countable_Sigma: "countable {x \ Sigma A B. 0 \ f x}" by (rule abs_summable_countable) define A' where "A' = fst ` {x \ Sigma A B. 0 \ f x}" have countA': "countable A'" using countable_Sigma unfolding A'_def by auto define B' where "B' a = snd ` ({x \ Sigma A B. 0 \ f x} \ {(a,b) | b. True})" for a have countB': "countable (B' a)" for a using countable_Sigma unfolding B'_def by auto have Sigma_eq: "x \ Sigma A B \ x \ Sigma A' B'" if "f x \ 0" for x - unfolding A'_def B'_def Sigma_def + unfolding A'_def B'_def Sigma_def using that by force have Sigma'_smaller: "Sigma A' B' \ Sigma A B" unfolding A'_def B'_def by auto with summable have summable': "f abs_summable_on Sigma A' B'" using abs_summable_on_subset by blast have A'_smaller: "A' \ A" unfolding A'_def by auto have B'_smaller: "B' a \ B a" for a unfolding B'_def by auto have "infsetsum f (Sigma A B) = infsetsum f (Sigma A' B')" proof (rule infsetsum_cong_neutral) show "f x = 0" if "x \ Sigma A B - Sigma A' B'" for x :: "'a \ 'b" using that - by (meson DiffD1 DiffD2 Sigma_eq) + by (meson DiffD1 DiffD2 Sigma_eq) show "f x = 0" if "x \ Sigma A' B' - Sigma A B" for x :: "'a \ 'b" - using that Sigma'_smaller by auto + using that Sigma'_smaller by auto show "f x = f x" if "x \ Sigma A B \ Sigma A' B'" for x :: "'a \ 'b" using that - by simp - qed + by simp + qed also from countA' countB' summable' have "\ = (\\<^sub>aa\A'. \\<^sub>ab\B' a. f (a, b))" by (rule infsetsum_Sigma) also have "\ = (\\<^sub>aa\A. \\<^sub>ab\B' a. f (a, b))" (is "_ = (\\<^sub>aa\A. ?inner a)" is "_ = ?rhs") proof (rule infsetsum_cong_neutral) show "(\\<^sub>ab\B' x. f (x, b)) = 0" if "x \ A' - A" for x :: 'a - using that A'_smaller by blast + using that A'_smaller by blast show "(\\<^sub>ab\B' x. f (x, b)) = 0" if "x \ A - A'" for x :: 'a proof - have f1: "x \ A" by (metis DiffD1 that) obtain bb :: "('b \ 'c) \ 'b set \ 'b" where "\x0 x1. (\v2. v2 \ x1 \ x0 v2 \ 0) = (bb x0 x1 \ x1 \ x0 (bb x0 x1) \ 0)" by moura hence f2: "\B f. bb f B \ B \ f (bb f B) \ 0 \ infsetsum f B = 0" by (meson infsetsum_all_0) have "(x, bb (\b. f (x, b)) (B' x)) \ Sigma A' B'" by (meson DiffD2 SigmaE2 that) thus ?thesis using f2 f1 by (meson B'_smaller SigmaI Sigma_eq in_mono) - qed + qed show "(\\<^sub>ab\B' x. f (x, b)) = (\\<^sub>ab\B' x. f (x, b))" if "x \ A' \ A" for x :: 'a using that - by simp + by simp qed also have "?inner a = (\\<^sub>ab\B a. f (a, b))" if "a\A" for a proof (rule infsetsum_cong_neutral) show "f (a, x) = 0" if "x \ B' a - B a" for x :: 'b - using that B'_smaller by blast + using that B'_smaller by blast show "f (a, x) = 0" if "x \ B a - B' a" for x :: 'b - using that Sigma_eq \a \ A\ by fastforce + using that Sigma_eq \a \ A\ by fastforce show "f (a, x) = f (a, x)" if "x \ B' a \ B a" for x :: 'b using that - by simp + by simp qed hence "?rhs = (\\<^sub>aa\A. \\<^sub>ab\B a. f (a, b))" by (rule infsetsum_cong, auto) - finally show ?thesis + finally show ?thesis by simp qed lemma infsetsum_Sigma': fixes A :: "'a set" and B :: "'a \ 'b set" assumes summable: "(\(x,y). f x y) abs_summable_on (Sigma A B)" shows "infsetsum (\x. infsetsum (\y. f x y) (B x)) A = infsetsum (\(x,y). f x y) (Sigma A B)" using assms by (subst infsetsum_Sigma) auto lemma infsetsum_Times: fixes A :: "'a set" and B :: "'b set" assumes summable: "f abs_summable_on (A \ B)" shows "infsetsum f (A \ B) = infsetsum (\x. infsetsum (\y. f (x, y)) B) A" using assms by (subst infsetsum_Sigma) auto lemma infsetsum_Times': fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c :: {banach, second_countable_topology}" assumes summable: "(\(x,y). f x y) abs_summable_on (A \ B)" shows "infsetsum (\x. infsetsum (\y. f x y) B) A = infsetsum (\(x,y). f x y) (A \ B)" using assms by (subst infsetsum_Times) auto lemma infsetsum_swap: fixes A :: "'a set" and B :: "'b set" fixes f :: "'a \ 'b \ 'c :: {banach, second_countable_topology}" assumes summable: "(\(x,y). f x y) abs_summable_on A \ B" shows "infsetsum (\x. infsetsum (\y. f x y) B) A = infsetsum (\y. infsetsum (\x. f x y) A) B" proof- from summable have summable': "(\(x,y). f y x) abs_summable_on B \ A" by (subst abs_summable_on_Times_swap) auto have bij: "bij_betw (\(x, y). (y, x)) (B \ A) (A \ B)" by (auto simp: bij_betw_def inj_on_def) have "infsetsum (\x. infsetsum (\y. f x y) B) A = infsetsum (\(x,y). f x y) (A \ B)" using summable by (subst infsetsum_Times) auto also have "\ = infsetsum (\(x,y). f y x) (B \ A)" by (subst infsetsum_reindex_bij_betw[OF bij, of "\(x,y). f x y", symmetric]) (simp_all add: case_prod_unfold) also have "\ = infsetsum (\y. infsetsum (\x. f x y) A) B" using summable' by (subst infsetsum_Times) auto finally show ?thesis. qed lemma abs_summable_infsetsum'_converges: fixes f :: "'a\'b::{second_countable_topology,banach}" and A :: "'a set" assumes "f abs_summable_on A" shows "infsetsum'_converges f A" proof- define F where "F = filtermap (sum f) (finite_subsets_at_top A)" have F_not_bot: "F \ bot" unfolding F_def filtermap_bot_iff by simp have "\P. eventually P (finite_subsets_at_top A) \ (\x y. P x \ P y \ dist (sum f x) (sum f y) < e)" if "0 < e" for e :: real proof- have is_SUP: "ereal (\\<^sub>ax\A. norm (f x)) = (SUP F\{F. finite F \ F \ A}. ereal (\x\F. norm (f x)))" proof (rule infsetsum_nonneg_is_SUPREMUM_ereal) show "(\x. norm (f x)) abs_summable_on A" by (simp add: assms) show "0 \ norm (f x)" if "x \ A" for x :: 'a using that - by simp - qed + by simp + qed have "\F0\{F. finite F \ F \ A}. (SUP F\{F. finite F \ F \ A}. ereal (\x\F. norm (f x))) - ereal e < ereal (\x\F0. norm (f x))" unfolding is_SUP Bex_def[symmetric] - by (smt less_SUP_iff[symmetric] \0 < e\ ereal_diff_le_self ereal_less_eq(5) ereal_minus(1) + by (smt less_SUP_iff[symmetric] \0 < e\ ereal_diff_le_self ereal_less_eq(5) ereal_minus(1) is_SUP less_eq_ereal_def) then obtain F0 where "F0\{F. finite F \ F \ A}" and "ereal (\x\F0. norm (f x)) > ereal (\\<^sub>ax\A. norm (f x)) - ereal e" - by (simp add: atomize_elim is_SUP) - hence [simp]: "finite F0" and [simp]: "F0 \ A" + by (simp add: atomize_elim is_SUP) + hence [simp]: "finite F0" and [simp]: "F0 \ A" and sum_diff: "sum (\x. norm (f x)) F0 > infsetsum (\x. norm (f x)) A - e" by auto define P where "P F \ finite F \ F \ F0 \ F \ A" for F have "dist (sum f F1) (sum f F2) < e" if "P F1" and "P F2" for F1 F2 proof - from that(1) have "finite F1" and "F1 \ F0" and "F1 \ A" unfolding P_def by auto from that(2) have "finite F2" and "F2 \ F0" and "F2 \ A" unfolding P_def by auto have "dist (sum f F1) (sum f F2) = norm (sum f (F1-F2) - sum f (F2-F1))" unfolding dist_norm - by (smt \finite F1\ \finite F2\ add_diff_cancel_left' add_diff_cancel_right' algebra_simps sum.Int_Diff sum.union_diff2 sum.union_inter) + by (smt \finite F1\ \finite F2\ add_diff_cancel_left' add_diff_cancel_right' algebra_simps sum.Int_Diff sum.union_diff2 sum.union_inter) also have "\ \ sum (\x. norm (f x)) (F1-F2) + sum (\x. norm (f x)) (F2-F1)" by (smt norm_triangle_ineq4 sum_norm_le) also have "\ = infsetsum (\x. norm (f x)) (F1-F2) + infsetsum (\x. norm (f x)) (F2-F1)" by (simp add: \finite F1\ \finite F2\) also have "\ = infsetsum (\x. norm (f x)) ((F1-F2)\(F2-F1))" proof (rule infsetsum_Un_disjoint [symmetric]) show "(\x. norm (f x)) abs_summable_on F1 - F2" - by (simp add: \finite F1\) + by (simp add: \finite F1\) show "(\x. norm (f x)) abs_summable_on F2 - F1" - by (simp add: \finite F2\) + by (simp add: \finite F2\) show "(F1 - F2) \ (F2 - F1) = {}" - by (simp add: Diff_Int_distrib2) + by (simp add: Diff_Int_distrib2) qed also have "\ \ infsetsum (\x. norm (f x)) (A-F0)" proof (rule infsetsum_mono_neutral_left) show "(\x. norm (f x)) abs_summable_on F1 - F2 \ (F2 - F1)" - by (simp add: \finite F1\ \finite F2\) + by (simp add: \finite F1\ \finite F2\) show "(\x. norm (f x)) abs_summable_on A - F0" - using abs_summable_on_subset assms by fastforce + using abs_summable_on_subset assms by fastforce show "norm (f x) \ norm (f x)" if "x \ F1 - F2 \ (F2 - F1)" for x :: 'a using that - by simp + by simp show "F1 - F2 \ (F2 - F1) \ A - F0" - by (simp add: Diff_mono \F0 \ F1\ \F0 \ F2\ \F1 \ A\ \F2 \ A\) + by (simp add: Diff_mono \F0 \ F1\ \F0 \ F2\ \F1 \ A\ \F2 \ A\) show "0 \ norm (f x)" if "x \ A - F0 - (F1 - F2 \ (F2 - F1))" for x :: 'a - by simp + by simp qed also have "\ = infsetsum (\x. norm (f x)) A - infsetsum (\x. norm (f x)) F0" by (simp add: assms infsetsum_Diff) also have "\ < e" using local.sum_diff by auto finally show "dist (sum f F1) (sum f F2) < e" by assumption qed moreover have "eventually P (finite_subsets_at_top A)" unfolding P_def eventually_finite_subsets_at_top - using \F0 \ A\ \finite F0\ by blast + using \F0 \ A\ \finite F0\ by blast ultimately show "\P. eventually P (finite_subsets_at_top A) \ (\F1 F2. P F1 \ P F2 \ dist (sum f F1) (sum f F2) < e)" by auto qed hence cauchy: "cauchy_filter F" unfolding F_def - by (rule cauchy_filter_metric_filtermapI) + by (rule cauchy_filter_metric_filtermapI) from complete_UNIV have "F\principal UNIV \ F \ bot \ cauchy_filter F \ (\x. F \ nhds x)" unfolding complete_uniform by auto have "(F \ principal UNIV \ F \ bot \ cauchy_filter F \ \x. F \ nhds x) \ \x. F \ nhds x" using cauchy F_not_bot by simp then obtain x where Fx: "F \ nhds x" using \\F \ principal UNIV; F \ bot; cauchy_filter F\ \ \x. F \ nhds x\ by blast hence "(sum f \ x) (finite_subsets_at_top A)" unfolding F_def by (simp add: filterlim_def) thus ?thesis unfolding infsetsum'_converges_def by auto qed lemma infsetsum'_converges_cofin_subset: fixes f :: "'a \ 'b::{topological_ab_group_add,t2_space}" assumes "infsetsum'_converges f A" and [simp]: "finite F" shows "infsetsum'_converges f (A-F)" proof- from assms(1) obtain x where lim_f: "(sum f \ x) (finite_subsets_at_top A)" unfolding infsetsum'_converges_def by auto define F' where "F' = F\A" with assms have "finite F'" and "A-F = A-F'" by auto have "filtermap ((\)F') (finite_subsets_at_top (A-F)) \ finite_subsets_at_top A" proof (rule filter_leI) fix P assume "eventually P (finite_subsets_at_top A)" - then obtain X where [simp]: "finite X" and XA: "X \ A" + then obtain X where [simp]: "finite X" and XA: "X \ A" and P: "\Y. finite Y \ X \ Y \ Y \ A \ P Y" unfolding eventually_finite_subsets_at_top by auto define X' where "X' = X-F" hence [simp]: "finite X'" and [simp]: "X' \ A-F" using XA by auto hence "finite Y \ X' \ Y \ Y \ A - F \ P (F' \ Y)" for Y using P XA unfolding X'_def using F'_def \finite F'\ by blast thus "eventually P (filtermap ((\) F') (finite_subsets_at_top (A - F)))" unfolding eventually_filtermap eventually_finite_subsets_at_top by (rule_tac x=X' in exI, simp) qed with lim_f have "(sum f \ x) (filtermap ((\)F') (finite_subsets_at_top (A-F)))" using tendsto_mono by blast have "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))" if "((sum f \ (\) F') \ x) (finite_subsets_at_top (A - F))" using that unfolding o_def by auto hence "((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A-F))" using tendsto_compose_filtermap [symmetric] - by (simp add: \(sum f \ x) (filtermap ((\) F') (finite_subsets_at_top (A - F)))\ + by (simp add: \(sum f \ x) (filtermap ((\) F') (finite_subsets_at_top (A - F)))\ tendsto_compose_filtermap) have "\Y. finite Y \ Y \ A - F \ sum f (F' \ Y) = sum f F' + sum f Y" by (metis Diff_disjoint Int_Diff \A - F = A - F'\ \finite F'\ inf.orderE sum.union_disjoint) hence "\\<^sub>F x in finite_subsets_at_top (A - F). sum f (F' \ x) = sum f F' + sum f x" unfolding eventually_finite_subsets_at_top using exI [where x = "{}"] - by (simp add: \\P. P {} \ \x. P x\) + by (simp add: \\P. P {} \ \x. P x\) hence "((\G. sum f F' + sum f G) \ x) (finite_subsets_at_top (A-F))" using tendsto_cong [THEN iffD1 , rotated] \((\G. sum f (F' \ G)) \ x) (finite_subsets_at_top (A - F))\ by fastforce hence "((\G. sum f F' + sum f G) \ sum f F' + (x-sum f F')) (finite_subsets_at_top (A-F))" by simp hence "(sum f \ x - sum f F') (finite_subsets_at_top (A-F))" - using Extra_General.tendsto_add_const_iff by blast + using Extra_General.tendsto_add_const_iff by blast thus "infsetsum'_converges f (A - F)" unfolding infsetsum'_converges_def by auto qed -lemma +lemma fixes f :: "'a \ 'b::{comm_monoid_add,t2_space}" assumes "\x. x\(A-B)\(B-A) \ f x = 0" shows infsetsum'_subset_zero: "infsetsum' f A = infsetsum' f B" and infsetsum'_converges_subset_zero: "infsetsum'_converges f A = infsetsum'_converges f B" proof - have convB: "infsetsum'_converges f B" and eq: "infsetsum' f A = infsetsum' f B" if convA: "infsetsum'_converges f A" and f0: "\x. x\(A-B)\(B-A) \ f x = 0" for A B proof - define D where "D = (A-B)" define D' where "D' = B-A" from convA obtain x where limA: "(sum f \ x) (finite_subsets_at_top A)" using infsetsum'_converges_def by blast have "sum f X = sum f (X - D)" if "finite (X::'a set)" and "X \ A" for X :: "'a set" using that f0 D_def by (simp add: subset_iff sum.mono_neutral_cong_right) hence "\\<^sub>F x in finite_subsets_at_top A. sum f x = sum f (x - D)" by (rule eventually_finite_subsets_at_top_weakI) hence "((\F. sum f (F-D)) \ x) (finite_subsets_at_top A)" using tendsto_cong [THEN iffD1, rotated] limA by fastforce hence "(sum f \ x) (filtermap (\F. F-D) (finite_subsets_at_top A))" by (simp add: filterlim_filtermap) have "D \ A" unfolding D_def by simp hence "finite_subsets_at_top (A - D) \ filtermap (\F. F - D) (finite_subsets_at_top A)" by (rule finite_subsets_at_top_minus) hence "(sum f \ x) (finite_subsets_at_top (A-D))" - using tendsto_mono [rotated] + using tendsto_mono [rotated] \(sum f \ x) (filtermap (\F. F - D) (finite_subsets_at_top A))\ by blast have "A - D \ B" unfolding D_def by auto hence "filtermap (\F. F \ (A - D)) (finite_subsets_at_top B) \ finite_subsets_at_top (A - D)" by (rule finite_subsets_at_top_inter [where B = B and A = "A-D"]) hence "(sum f \ x) (filtermap (\F. F \ (A - D)) (finite_subsets_at_top B))" using tendsto_mono [rotated] \(sum f \ x) (finite_subsets_at_top (A - D))\ by blast hence "((\F. sum f (F \ (A - D))) \ x) (finite_subsets_at_top B)" by (simp add: filterlim_filtermap) have "sum f (X \ (A - D)) = sum f X" if "finite (X::'a set)" and "X \ B" for X :: "'a set" proof (rule sum.mono_neutral_cong) show "finite X" by (simp add: that(1)) show "finite (X \ (A - D))" by (simp add: that(1)) show "f i = 0" if "i \ X - X \ (A - D)" for i :: 'a - using that D_def DiffD2 \X \ B\ f0 by auto + using that D_def DiffD2 \X \ B\ f0 by auto show "f i = 0" if "i \ X \ (A - D) - X" for i :: 'a using that - by auto + by auto show "f x = f x" if "x \ X \ (A - D) \ X" for x :: 'a by simp qed hence "\\<^sub>F x in finite_subsets_at_top B. sum f (x \ (A - D)) = sum f x" - by (rule eventually_finite_subsets_at_top_weakI) + by (rule eventually_finite_subsets_at_top_weakI) hence limB: "(sum f \ x) (finite_subsets_at_top B)" using tendsto_cong [THEN iffD1 , rotated] \((\F. sum f (F \ (A - D))) \ x) (finite_subsets_at_top B)\ by blast thus "infsetsum'_converges f B" unfolding infsetsum'_converges_def by auto have "Lim (finite_subsets_at_top A) (sum f) = Lim (finite_subsets_at_top B) (sum f)" if "infsetsum'_converges f B" using that using limA limB using finite_subsets_at_top_neq_bot tendsto_Lim by blast thus "infsetsum' f A = infsetsum' f B" - unfolding infsetsum'_def + unfolding infsetsum'_def using convA by (simp add: \infsetsum'_converges f B\) qed with assms show "infsetsum'_converges f A = infsetsum'_converges f B" by (metis Un_commute) thus "infsetsum' f A = infsetsum' f B" using assms convB eq by (metis infsetsum'_def) qed lemma fixes f :: "'a \ 'b::{topological_ab_group_add,t2_space}" assumes "infsetsum'_converges f B" and "infsetsum'_converges f A" and AB: "A \ B" shows infsetsum'_Diff: "infsetsum' f (B - A) = infsetsum' f B - infsetsum' f A" and infsetsum'_converges_Diff: "infsetsum'_converges f (B-A)" proof - define limA limB where "limA = infsetsum' f A" and "limB = infsetsum' f B" from assms(1) have limB: "(sum f \ limB) (finite_subsets_at_top B)" unfolding infsetsum'_converges_def infsetsum'_def limB_def by (auto simp: tendsto_Lim) from assms(2) have limA: "(sum f \ limA) (finite_subsets_at_top A)" unfolding infsetsum'_converges_def infsetsum'_def limA_def by (auto simp: tendsto_Lim) have "((\F. sum f (F\A)) \ limA) (finite_subsets_at_top B)" proof (subst asm_rl [of "(\F. sum f (F\A)) = sum f o (\F. F\A)"]) show "(\F. sum f (F \ A)) = sum f \ (\F. F \ A)" unfolding o_def by auto show "((sum f \ (\F. F \ A)) \ limA) (finite_subsets_at_top B)" - unfolding o_def + unfolding o_def using tendsto_compose_filtermap finite_subsets_at_top_inter[OF AB] limA tendsto_mono \(\F. sum f (F \ A)) = sum f \ (\F. F \ A)\ by fastforce qed with limB have "((\F. sum f F - sum f (F\A)) \ limB - limA) (finite_subsets_at_top B)" using tendsto_diff by blast have "sum f X - sum f (X \ A) = sum f (X - A)" if "finite (X::'a set)" and "X \ B" for X :: "'a set" using that by (metis add_diff_cancel_left' sum.Int_Diff) hence "\\<^sub>F x in finite_subsets_at_top B. sum f x - sum f (x \ A) = sum f (x - A)" - by (rule eventually_finite_subsets_at_top_weakI) + by (rule eventually_finite_subsets_at_top_weakI) hence "((\F. sum f (F-A)) \ limB - limA) (finite_subsets_at_top B)" using tendsto_cong [THEN iffD1 , rotated] \((\F. sum f F - sum f (F \ A)) \ limB - limA) (finite_subsets_at_top B)\ by fastforce hence "(sum f \ limB - limA) (filtermap (\F. F-A) (finite_subsets_at_top B))" by (subst tendsto_compose_filtermap[symmetric], simp add: o_def) hence limBA: "(sum f \ limB - limA) (finite_subsets_at_top (B-A))" using finite_subsets_at_top_minus[OF AB] by (rule tendsto_mono[rotated]) thus "infsetsum'_converges f (B-A)" unfolding infsetsum'_converges_def by auto with limBA show "infsetsum' f (B - A) = limB - limA" - unfolding infsetsum'_def by (simp add: tendsto_Lim) + unfolding infsetsum'_def by (simp add: tendsto_Lim) qed lemma infsetsum'_mono_set: fixes f :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes fx0: "\x. x\B-A \ f x \ 0" and "A \ B" and "infsetsum'_converges f A" and "infsetsum'_converges f B" shows "infsetsum' f B \ infsetsum' f A" proof - define limA limB f' where "limA = infsetsum' f A" and "limB = infsetsum' f B" and "f' x = (if x \ A then f x else 0)" for x have "infsetsum' f A = infsetsum' f' B" proof (subst infsetsum'_subset_zero [where f = f' and B = A]) show "f' x = 0" if "x \ B - A \ (A - B)" for x :: 'a - using that assms(2) f'_def by auto + using that assms(2) f'_def by auto show "infsetsum' f A = infsetsum' f' A" - by (metis f'_def infsetsum'_cong) + by (metis f'_def infsetsum'_cong) qed hence limA_def': "limA = infsetsum' f' B" unfolding limA_def by auto have convA': "infsetsum'_converges f' B" proof (rule infsetsum'_converges_subset_zero [THEN iffD1 , where A1 = A]) show "f' x = 0" if "x \ A - B \ (B - A)" for x :: 'a - using that assms(2) f'_def by auto + using that assms(2) f'_def by auto show "infsetsum'_converges f' A" - by (simp add: assms(3) f'_def infsetsum'_converges_cong) + by (simp add: assms(3) f'_def infsetsum'_converges_cong) qed - from assms have limA: "(sum f \ limA) (finite_subsets_at_top A)" + from assms have limA: "(sum f \ limA) (finite_subsets_at_top A)" and limB: "(sum f \ limB) (finite_subsets_at_top B)" by (auto simp: limA_def limB_def infsetsum'_converges_def infsetsum'_def tendsto_Lim) have limA': "(sum f' \ limA) (finite_subsets_at_top B)" using finite_subsets_at_top_neq_bot tendsto_Lim convA' unfolding limA_def' infsetsum'_def infsetsum'_converges_def - by fastforce + by fastforce have "f' i \ f i" if "i \ X" and "X \ B" for i :: 'a and X unfolding f'_def using fx0 that using \X \ B\ by auto hence "sum f' X \ sum f X" if "finite (X::'a set)" and "X \ B" for X :: "'a set" using sum_mono - by (simp add: sum_mono that(2)) + by (simp add: sum_mono that(2)) hence sumf'_leq_sumf: "\\<^sub>F x in finite_subsets_at_top B. sum f' x \ sum f x" by (rule eventually_finite_subsets_at_top_weakI) show "limA \ limB" - using finite_subsets_at_top_neq_bot limB limA' sumf'_leq_sumf + using finite_subsets_at_top_neq_bot limB limA' sumf'_leq_sumf by (rule tendsto_le) qed lemma infsetsum'_converges_finite[simp]: assumes "finite F" shows "infsetsum'_converges f F" unfolding infsetsum'_converges_def finite_subsets_at_top_finite[OF assms] - using tendsto_principal_singleton by fastforce + using tendsto_principal_singleton by fastforce lemma infsetsum'_finite[simp]: assumes "finite F" shows "infsetsum' f F = sum f F" using assms by (auto intro: tendsto_Lim simp: finite_subsets_at_top_finite infsetsum'_def principal_eq_bot_iff tendsto_principal_singleton) lemma infsetsum'_approx_sum: fixes f :: "'a \ 'b::{comm_monoid_add,metric_space}" assumes "infsetsum'_converges f A" and "\ > 0" shows "\F. finite F \ F \ A \ dist (sum f F) (infsetsum' f A) \ \" proof- have "infsetsum'_converges f A \ 0 < \ \ (sum f \ Lim (finite_subsets_at_top A) (sum f)) (finite_subsets_at_top A)" unfolding infsetsum'_converges_def using Lim_trivial_limit tendsto_Lim by blast hence "(sum f \ infsetsum' f A) (finite_subsets_at_top A)" unfolding infsetsum'_def using assms by simp hence "\\<^sub>F F in (finite_subsets_at_top A). dist (sum f F) (infsetsum' f A) < \" using assms(2) by (rule tendstoD) have "finite X \ X \ A \ \Y. finite Y \ X \ Y \ Y \ A \ dist (sum f Y) (infsetsum' f A) < \ \ \F. finite F \ F \ A \ dist (sum f F) (infsetsum' f A) \ \" for X - by fastforce + by fastforce thus ?thesis using eventually_finite_subsets_at_top by (metis (no_types, lifting) \\\<^sub>F F in finite_subsets_at_top A. dist (sum f F) (infsetsum' f A) < \\) qed lemma norm_infsetsum'_bound: fixes f :: "'b \ 'a::real_normed_vector" and A :: "'b set" assumes a1: "infsetsum'_converges (\x. norm (f x)) A" shows "norm (infsetsum' f A) \ (infsetsum' (\x. norm (f x)) A)" proof(cases "infsetsum'_converges f A") case True have "norm (infsetsum' f A) \ (infsetsum' (\x. norm (f x)) A) + \" if "\>0" for \ proof- have "\F. norm (infsetsum' f A - sum f F) \ \ \ finite F \ F \ A" using infsetsum'_approx_sum[where A=A and f=f and \="\"] a1 True \0 < \\ by (metis dist_commute dist_norm) then obtain F where "norm (infsetsum' f A - sum f F) \ \" and "finite F" and "F \ A" by (simp add: atomize_elim) hence "norm (infsetsum' f A) \ norm (sum f F) + \" by (smt norm_triangle_sub) also have "\ \ sum (\x. norm (f x)) F + \" using norm_sum by auto also have "\ \ (infsetsum' (\x. norm (f x)) A) + \" proof- have "infsetsum' (\x. norm (f x)) F \ infsetsum' (\x. norm (f x)) A" proof (rule infsetsum'_mono_set) show "0 \ norm (f x)" if "x \ A - F" for x :: 'b using that - by simp + by simp show "F \ A" - by (simp add: \F \ A\) + by (simp add: \F \ A\) show "infsetsum'_converges (\x. norm (f x)) F" - using \finite F\ by auto + using \finite F\ by auto show "infsetsum'_converges (\x. norm (f x)) A" - by (simp add: assms) + by (simp add: assms) qed thus ?thesis by (simp_all flip: infsetsum'_finite add: \finite F\) qed - finally show ?thesis + finally show ?thesis by assumption qed thus ?thesis using linordered_field_class.field_le_epsilon by blast next case False obtain t where t_def: "(sum (\x. norm (f x)) \ t) (finite_subsets_at_top A)" using a1 unfolding infsetsum'_converges_def by blast have sumpos: "sum (\x. norm (f x)) X \ 0" for X by (simp add: sum_nonneg) have tgeq0:"t \ 0" proof(rule ccontr) define S::"real set" where "S = {s. s < 0}" assume "\ 0 \ t" hence "t < 0" by simp hence "t \ S" unfolding S_def by blast moreover have "open S" proof- have "closed {s::real. s \ 0}" using Elementary_Topology.closed_sequential_limits[where S = "{s::real. s \ 0}"] by (metis Lim_bounded2 mem_Collect_eq) moreover have "{s::real. s \ 0} = UNIV - S" unfolding S_def by auto ultimately have "closed (UNIV - S)" by simp thus ?thesis - by (simp add: Compl_eq_Diff_UNIV open_closed) + by (simp add: Compl_eq_Diff_UNIV open_closed) qed ultimately have "\\<^sub>F X in finite_subsets_at_top A. (\x\X. norm (f x)) \ S" using t_def unfolding tendsto_def by blast hence "\X. (\x\X. norm (f x)) \ S" by (metis (no_types, lifting) False eventually_mono filterlim_iff infsetsum'_converges_def) then obtain X where "(\x\X. norm (f x)) \ S" by blast hence "(\x\X. norm (f x)) < 0" - unfolding S_def by auto + unfolding S_def by auto thus False using sumpos by smt qed have "\!h. (sum (\x. norm (f x)) \ h) (finite_subsets_at_top A)" using t_def finite_subsets_at_top_neq_bot tendsto_unique by blast hence "t = (Topological_Spaces.Lim (finite_subsets_at_top A) (sum (\x. norm (f x))))" using t_def unfolding Topological_Spaces.Lim_def - by (metis the_equality) + by (metis the_equality) hence "Lim (finite_subsets_at_top A) (sum (\x. norm (f x))) \ 0" using tgeq0 by blast - thus ?thesis unfolding infsetsum'_def + thus ?thesis unfolding infsetsum'_def using False by auto qed lemma infsetsum_infsetsum': assumes "f abs_summable_on A" shows "infsetsum f A = infsetsum' f A" proof- have conv_sum_norm[simp]: "infsetsum'_converges (\x. norm (f x)) A" proof (rule abs_summable_infsetsum'_converges) show "(\x. norm (f x)) abs_summable_on A" using assms by simp - qed + qed have "norm (infsetsum f A - infsetsum' f A) \ \" if "\>0" for \ proof - define \ where "\ = \/2" with that have [simp]: "\ > 0" by simp obtain F1 where F1A: "F1 \ A" and finF1: "finite F1" and leq_eps: "infsetsum (\x. norm (f x)) (A-F1) \ \" proof - have sum_SUP: "ereal (infsetsum (\x. norm (f x)) A) = (SUP F\{F. finite F \ F \ A}. ereal (sum (\x. norm (f x)) F))" (is "_ = ?SUP") proof (rule infsetsum_nonneg_is_SUPREMUM_ereal) show "(\x. norm (f x)) abs_summable_on A" - by (simp add: assms) + by (simp add: assms) show "0 \ norm (f x)" if "x \ A" for x :: 'a using that - by simp + by simp qed have "(SUP F\{F. finite F \ F \ A}. ereal (\x\F. norm (f x))) - ereal \ < (SUP i\{F. finite F \ F \ A}. ereal (\x\i. norm (f x)))" using \\>0\ by (metis diff_strict_left_mono diff_zero ereal_less_eq(3) ereal_minus(1) not_le sum_SUP) then obtain F where "F\{F. finite F \ F \ A}" and "ereal (sum (\x. norm (f x)) F) > ?SUP - ereal (\)" by (meson less_SUP_iff) hence "sum (\x. norm (f x)) F > infsetsum (\x. norm (f x)) A - (\)" unfolding sum_SUP[symmetric] by auto hence "\ > infsetsum (\x. norm (f x)) (A-F)" proof (subst infsetsum_Diff) show "(\x. norm (f x)) abs_summable_on A" if "(\\<^sub>ax\A. norm (f x)) - \ < (\x\F. norm (f x))" using that - by (simp add: assms) + by (simp add: assms) show "F \ A" if "(\\<^sub>ax\A. norm (f x)) - \ < (\x\F. norm (f x))" - using that \F \ {F. finite F \ F \ A}\ by blast + using that \F \ {F. finite F \ F \ A}\ by blast show "(\\<^sub>ax\A. norm (f x)) - (\\<^sub>ax\F. norm (f x)) < \" if "(\\<^sub>ax\A. norm (f x)) - \ < (\x\F. norm (f x))" - using that \F \ {F. finite F \ F \ A}\ by auto + using that \F \ {F. finite F \ F \ A}\ by auto qed - thus ?thesis using that + thus ?thesis using that apply atomize_elim using \F \ {F. finite F \ F \ A}\ less_imp_le by blast qed have "\F2\A. finite F2 \ dist (\x\F2. norm (f x)) (infsetsum' (\x. norm (f x)) A) \ \" using infsetsum'_approx_sum[where f="(\x. norm (f x))" and A=A and \=\] abs_summable_infsetsum'_converges assms by auto then obtain F2 where F2A: "F2 \ A" and finF2: "finite F2" and dist: "dist (sum (\x. norm (f x)) F2) (infsetsum' (\x. norm (f x)) A) \ \" - by blast + by blast have leq_eps': "infsetsum' (\x. norm (f x)) (A-F2) \ \" proof (subst infsetsum'_Diff) show "infsetsum'_converges (\x. norm (f x)) A" - by simp + by simp show "infsetsum'_converges (\x. norm (f x)) F2" - by (simp add: finF2) + by (simp add: finF2) show "F2 \ A" - by (simp add: F2A) + by (simp add: F2A) show "infsetsum' (\x. norm (f x)) A - infsetsum' (\x. norm (f x)) F2 \ \" using dist finF2 by (auto simp: dist_norm) - qed + qed define F where "F = F1 \ F2" - have FA: "F \ A" and finF: "finite F" + have FA: "F \ A" and finF: "finite F" unfolding F_def using F1A F2A finF1 finF2 by auto have "(\\<^sub>ax\A - (F1 \ F2). norm (f x)) \ (\\<^sub>ax\A - F1. norm (f x))" proof (rule infsetsum_mono_neutral_left) show "(\x. norm (f x)) abs_summable_on A - (F1 \ F2)" - using abs_summable_on_subset assms by fastforce + using abs_summable_on_subset assms by fastforce show "(\x. norm (f x)) abs_summable_on A - F1" - using abs_summable_on_subset assms by fastforce + using abs_summable_on_subset assms by fastforce show "norm (f x) \ norm (f x)" if "x \ A - (F1 \ F2)" for x :: 'a using that - by auto + by auto show "A - (F1 \ F2) \ A - F1" - by (simp add: Diff_mono) + by (simp add: Diff_mono) show "0 \ norm (f x)" if "x \ A - F1 - (A - (F1 \ F2))" for x :: 'a using that - by auto + by auto qed hence leq_eps: "infsetsum (\x. norm (f x)) (A-F) \ \" unfolding F_def using leq_eps by linarith have "infsetsum' (\x. norm (f x)) (A - (F1 \ F2)) \ infsetsum' (\x. norm (f x)) (A - F2)" proof (rule infsetsum'_mono_set) show "0 \ norm (f x)" if "x \ A - F2 - (A - (F1 \ F2))" for x :: 'a using that - by simp + by simp show "A - (F1 \ F2) \ A - F2" - by (simp add: Diff_mono) + by (simp add: Diff_mono) show "infsetsum'_converges (\x. norm (f x)) (A - (F1 \ F2))" - using F_def conv_sum_norm finF infsetsum'_converges_cofin_subset by blast + using F_def conv_sum_norm finF infsetsum'_converges_cofin_subset by blast show "infsetsum'_converges (\x. norm (f x)) (A - F2)" - by (simp add: finF2 infsetsum'_converges_cofin_subset) + by (simp add: finF2 infsetsum'_converges_cofin_subset) qed hence leq_eps': "infsetsum' (\x. norm (f x)) (A-F) \ \" - unfolding F_def + unfolding F_def by (rule order.trans[OF _ leq_eps']) have "norm (infsetsum f A - infsetsum f F) = norm (infsetsum f (A-F))" proof (subst infsetsum_Diff [symmetric]) show "f abs_summable_on A" - by (simp add: assms) + by (simp add: assms) show "F \ A" - by (simp add: FA) + by (simp add: FA) show "norm (infsetsum f (A - F)) = norm (infsetsum f (A - F))" - by simp + by simp qed also have "\ \ infsetsum (\x. norm (f x)) (A-F)" using norm_infsetsum_bound by blast also have "\ \ \" using leq_eps by simp finally have diff1: "norm (infsetsum f A - infsetsum f F) \ \" by assumption have "norm (infsetsum' f A - infsetsum' f F) = norm (infsetsum' f (A-F))" proof (subst infsetsum'_Diff [symmetric]) show "infsetsum'_converges f A" - by (simp add: abs_summable_infsetsum'_converges assms) + by (simp add: abs_summable_infsetsum'_converges assms) show "infsetsum'_converges f F" - by (simp add: finF) + by (simp add: finF) show "F \ A" - by (simp add: FA) + by (simp add: FA) show "norm (infsetsum' f (A - F)) = norm (infsetsum' f (A - F))" - by simp + by simp qed also have "\ \ infsetsum' (\x. norm (f x)) (A-F)" by (simp add: finF infsetsum'_converges_cofin_subset norm_infsetsum'_bound) also have "\ \ \" using leq_eps' by simp finally have diff2: "norm (infsetsum' f A - infsetsum' f F) \ \" by assumption have x1: "infsetsum f F = infsetsum' f F" using finF by simp have "norm (infsetsum f A - infsetsum' f A) \ norm (infsetsum f A - infsetsum f F) + norm (infsetsum' f A - infsetsum' f F)" apply (rule_tac norm_diff_triangle_le) apply auto by (simp_all add: x1 norm_minus_commute) also have "\ \ \" using diff1 diff2 \_def by linarith finally show ?thesis by assumption qed hence "norm (infsetsum f A - infsetsum' f A) = 0" by (meson antisym_conv1 dense_ge norm_not_less_zero) thus ?thesis by auto qed lemma abs_summable_partition: fixes T :: "'b set" and I :: "'a set" assumes "\i. f abs_summable_on S i" and "(\i. \\<^sub>ax\S i. norm (f x)) abs_summable_on I" and "T \ (\i\I. S i)" shows "f abs_summable_on T" proof (rule abs_summable_finiteI) fix F assume finite_F: "finite F" and FT: "F \ T" define index where "index s = (SOME i. i\I \ s\S i)" for s hence index_I: "index s \ I" and S_index: "s \ S (index s)" if "s \ (\i\I. S i)" for s proof auto show "(SOME i. i \ I \ s \ S i) \ I" if "\s. index s = (SOME i. i \ I \ s \ S i)" using that - by (metis (no_types, lifting) UN_iff \s \ \ (S ` I)\ someI_ex) + by (metis (no_types, lifting) UN_iff \s \ \ (S ` I)\ someI_ex) show "s \ S (SOME i. i \ I \ s \ S i)" if "\s. index s = (SOME i. i \ I \ s \ S i)" using that - by (metis (no_types, lifting) UN_iff \s \ \ (S ` I)\ someI_ex) + by (metis (no_types, lifting) UN_iff \s \ \ (S ` I)\ someI_ex) qed define S' where "S' i = {s\S i. i = index s}" for i have S'_S: "S' i \ S i" for i unfolding S'_def by simp hence f_sum_S': "f abs_summable_on S' i" for i by (meson abs_summable_on_subset assms(1)) with assms(1) S'_S have "(\\<^sub>ax\S' i. norm (f x)) \ (\\<^sub>ax\S i. norm (f x))" for i by (simp add: infsetsum_mono_neutral_left) with assms(2) have sum_I: "(\i. \\<^sub>ax\S' i. norm (f x)) abs_summable_on I" by (smt abs_summable_on_comparison_test' infsetsum_cong norm_ge_zero norm_infsetsum_bound real_norm_def) have "(\i\I. S i) = (\i\I. S' i)" unfolding S'_def by (auto intro!: index_I S_index) with assms(3) have T_S': "T \ (\i\I. S' i)" by simp have S'_disj: "(S' i) \ (S' j) = {}" if "i\j" for i j unfolding S'_def disjnt_def using that by auto define B where "B i = (\\<^sub>ax\S i. norm (f x))" for i have sum_FS'_B: "(\x\F\S' i. norm (f x)) \ B i" for i unfolding B_def using f_sum_S' finite_F FT - by (metis S'_S abs_summable_finiteI_converse assms(1) finite_Int le_inf_iff order_refl + by (metis S'_S abs_summable_finiteI_converse assms(1) finite_Int le_inf_iff order_refl subset_antisym) have B_pos[simp]: "B i \ 0" for i unfolding B_def by (rule infsetsum_nonneg, simp) have B_sum_I[simp]: "B abs_summable_on I" by (simp add: B_def assms(2)) define J where "J = {i\I. F\S' i \ {}}" have finite_J[simp]: "finite J" proof - define a where "a i = (SOME x. x\F\S' i)" for i hence a: "a i \ F\S' i" if "i \ J" for i unfolding J_def by (metis (mono_tags) Collect_conj_eq Int_Collect J_def some_in_eq that) have xy: "x = y" if "x \ J" and "y \ J" and "a x = a y" and "\i. i \ J \ a i \ F \ a i \ S' i" and "\i j. i \ j \ S' i \ S' j = {}" - for x y + for x y using that a S'_disj by (metis S'_disj disjoint_iff_not_equal) hence "inj_on a J" unfolding inj_on_def - using S'_disj a by auto + using S'_disj a by auto moreover have "a ` J \ F" using a by auto ultimately show "finite J" using finite_F Finite_Set.inj_on_finite by blast qed have JI[simp]: "J \ I" unfolding J_def by simp have "F = (\i\J. F\S' i)" unfolding J_def apply auto by (metis FT T_S' UN_E disjoint_iff_not_equal subsetD) hence "(\x\F. norm (f x)) = (\x\(\i\J. F\S' i). norm (f x))" by simp also have "\ = (\i\J. \x\F \ S' i. norm (f x))" proof (rule sum.UNION_disjoint) show "finite J" - by simp + by simp show "\i\J. finite (F \ S' i)" - by (simp add: finite_F) + by (simp add: finite_F) show "\i\J. \j\J. i \ j \ F \ S' i \ (F \ S' j) = {}" - using S'_disj by auto + using S'_disj by auto qed also have "\ \ (\i\J. B i)" using sum_FS'_B by (simp add: ordered_comm_monoid_add_class.sum_mono) also have "\ = (\\<^sub>ai\J. B i)" by simp also have "\ \ (\\<^sub>ai\I. B i)" proof (rule infsetsum_mono_neutral_left) show "B abs_summable_on J" - by simp + by simp show "B abs_summable_on I" by simp show "B x \ B x" if "x \ J" for x :: 'a using that - by simp + by simp show "J \ I" - by simp + by simp show "0 \ B x" if "x \ I - J" for x :: 'a using that - by simp - qed + by simp + qed finally show "(\x\F. norm(f x)) \ (\\<^sub>ai\I. B i)" by simp qed lemma abs_summable_product': fixes X :: "'a set" and Y :: "'b set" assumes "\x. (\y. f (x,y)) abs_summable_on Y" and "(\x. \\<^sub>ay\Y. norm (f (x,y))) abs_summable_on X" shows "f abs_summable_on X\Y" proof- define S where "S x = {x} \ Y" for x :: 'a have bij[simp]: "bij_betw (Pair x) Y (S x)" for x proof (rule bij_betwI [where g = snd]) show "Pair x \ Y \ S x" - by (simp add: S_def) + by (simp add: S_def) show "snd \ S x \ Y" - using Pi_I' S_def by auto + using Pi_I' S_def by auto show "snd (y, x::'b) = x" if "x \ Y" for x :: 'b and y::'a using that - by simp + by simp show "(x, snd y::'b) = y" if "y \ S x" for y :: "'a \ 'b" using that unfolding S_def by auto qed have "f abs_summable_on S x" for x proof (subst abs_summable_on_reindex_bij_betw [symmetric , where A = Y and g = "\y. (x,y)"]) show "bij_betw (Pair x) Y (S x)" - by simp + by simp show "(\xa. f (x, xa)) abs_summable_on Y" - using assms(1) by auto + using assms(1) by auto qed moreover have "bij_betw (Pair x) Y (S x)" for x unfolding S_def using bij_betw_def using S_def bij by auto hence "(\\<^sub>ay\Y. norm (f (x, y))) = (\\<^sub>ay\S x. norm (f y))" for x - by (rule infsetsum_reindex_bij_betw) + by (rule infsetsum_reindex_bij_betw) hence "(\i. \\<^sub>ax\S i. norm (f x)) abs_summable_on X" using assms(2) by simp hence "(\i. \\<^sub>ax\S i. norm (f x)) abs_summable_on X" by auto moreover have "X \ Y \ (\i\X. S i)" unfolding S_def by auto ultimately show ?thesis by (rule abs_summable_partition[where S=S and I=X]) qed lemma infsetsum_prod_PiE: fixes f :: "'a \ 'b \ 'c :: {real_normed_field,banach,second_countable_topology}" assumes finite: "finite A" and summable: "\x. x \ A \ f x abs_summable_on B x" shows "infsetsum (\g. \x\A. f x (g x)) (PiE A B) = (\x\A. infsetsum (f x) (B x))" proof- define B' where "B' x = {y\B x. 0 \ f x y}" for x have [simp]: "B' x \ B x" for x unfolding B'_def by simp have PiE_subset: "Pi\<^sub>E A B' \ Pi\<^sub>E A B" by (simp add: PiE_mono) have "f x abs_summable_on B x" if "x\A" for x using that - by (simp add: local.summable) + by (simp add: local.summable) hence countable: "countable (B' x)" if "x\A" for x unfolding B'_def using abs_summable_countable using that by blast have summable: "f x abs_summable_on B' x" if "x\A" for x using that summable[where x = x] \\x. B' x \ B x\ abs_summable_on_subset by blast have 0: "(\x\A. f x (g x)) = 0" if "g \ Pi\<^sub>E A B - Pi\<^sub>E A B'" for g proof- from that have "g \ extensional A" by (simp add: PiE_def) from that have "g \ Pi\<^sub>E A B'" by simp with \g \ extensional A\ have "g \ Pi A B'" unfolding PiE_def by simp then obtain x where "x\A" and "g x \ B' x" unfolding Pi_def by auto hence "f x (g x) = 0" unfolding B'_def using that by auto with finite show ?thesis proof (rule_tac prod_zero) show "finite A" if "finite A" and "f x (g x) = 0" using that - by simp + by simp show "\a\A. f a (g a) = 0" if "finite A" and "f x (g x) = 0" - using that \x \ A\ by blast - qed + using that \x \ A\ by blast + qed qed have d: "infsetsum (f x) (B' x) = infsetsum (f x) (B x)" if "x \ A" for x proof (rule infsetsum_cong_neutral) show "f y x = 0" if "x \ B' y - B y" for x :: 'b and y :: 'a using that - by (meson DiffD1 DiffD2 \\x. B' x \ B x\ in_mono) + by (meson DiffD1 DiffD2 \\x. B' x \ B x\ in_mono) show "f y x = 0" if "x \ B y - B' y" for x :: 'b and y - using that B'_def by auto + using that B'_def by auto show "f y x = f y x" if "x \ B' y \ B y" for x :: 'b and y using that - by simp - qed + by simp + qed have "infsetsum (\g. \x\A. f x (g x)) (PiE A B) = infsetsum (\g. \x\A. f x (g x)) (PiE A B')" proof (rule infsetsum_cong_neutral) show "(\a\A. f a (x a)) = 0" if "x \ Pi\<^sub>E A B - Pi\<^sub>E A B'" for x :: "'a \ 'b" using that - by (simp add: "0") + by (simp add: "0") show "(\a\A. f a (x a)) = 0" if "x \ Pi\<^sub>E A B' - Pi\<^sub>E A B" for x :: "'a \ 'b" - using that PiE_subset by auto + using that PiE_subset by auto show "(\a\A. f a (x a)) = (\a\A. f a (x a))" if "x \ Pi\<^sub>E A B \ Pi\<^sub>E A B'" for x :: "'a \ 'b" using that - by simp + by simp qed also have "\ = (\x\A. infsetsum (f x) (B' x))" using finite countable summable by (rule infsetsum_prod_PiE) also have "\ = (\x\A. infsetsum (f x) (B x))" using d by auto finally show ?thesis. qed lemma infsetsum_0D: fixes f :: "'a \ real" assumes "infsetsum f A = 0" and abs_sum: "f abs_summable_on A" and nneg: "\x. x \ A \ f x \ 0" and "x \ A" shows "f x = 0" proof - from abs_sum have [simp]: "f abs_summable_on (A-{x})" by (meson Diff_subset abs_summable_on_subset) from abs_sum \x\A\ have [simp]: "f abs_summable_on {x}" by auto have a: "\a. a \ A - {x} \ a \ A" - by simp + by simp from assms have "0 = infsetsum f A" by simp also have "\ = infsetsum f (A-{x}) + infsetsum f {x}" proof (subst infsetsum_Un_disjoint [symmetric]) show "f abs_summable_on A - {x}" - by simp + by simp show "f abs_summable_on {x}" - by simp + by simp show "(A - {x}) \ {x} = {}" - by simp + by simp show "infsetsum f A = infsetsum f (A - {x} \ {x})" - using assms(4) insert_Diff by fastforce + using assms(4) insert_Diff by fastforce qed also have "\ \ 0 + infsetsum f {x}" (is "_ \ \") using a - by (smt infsetsum_nonneg nneg) + by (smt infsetsum_nonneg nneg) also have "\ = f x" by simp finally have "f x \ 0". with nneg[OF \x\A\] show "f x = 0" by auto qed lemma sum_leq_infsetsum: fixes f :: "_ \ real" assumes "f abs_summable_on N" and "finite M" and "M \ N" and "\x. x\N-M \ f x \ 0" shows "sum f M \ infsetsum f N" proof - have "infsetsum f M \ infsetsum f N" proof (rule infsetsum_mono_neutral_left) show "f abs_summable_on M" - by (simp add: assms(2)) + by (simp add: assms(2)) show "f abs_summable_on N" - by (simp add: assms(1)) + by (simp add: assms(1)) show "f x \ f x" if "x \ M" for x :: 'b using that - by simp + by simp show "M \ N" - by (simp add: assms(3)) + by (simp add: assms(3)) show "0 \ f x" if "x \ N - M" for x :: 'b using that - by (simp add: assms(4)) + by (simp add: assms(4)) qed thus ?thesis using assms by auto qed lemma infsetsum_cmult_left': fixes f :: "'a \ 'b :: {banach, real_normed_algebra, second_countable_topology, division_ring}" shows "infsetsum (\x. f x * c) A = infsetsum f A * c" proof (cases "c \ 0 \ f abs_summable_on A") case True have "(\\<^sub>ax\A. f x * c) = infsetsum f A * c" if "f abs_summable_on A" using infsetsum_cmult_left that by blast thus ?thesis - using True by auto + using True by auto next case False hence "c\0" and "\ f abs_summable_on A" by auto have "\ (\x. f x * c) abs_summable_on A" proof (rule notI) assume "(\x. f x * c) abs_summable_on A" hence "(\x. (f x * c) * inverse c) abs_summable_on A" by (rule abs_summable_on_cmult_left) with \\ f abs_summable_on A\ show False by (metis (no_types, lifting) False Groups.mult_ac(1) abs_summable_on_cong mult_1_right right_inverse) qed with \\ f abs_summable_on A\ - show ?thesis + show ?thesis by (simp add: not_summable_infsetsum_eq) qed lemma abs_summable_on_zero_diff: assumes "f abs_summable_on A" and "\x. x \ B - A \ f x = 0" shows "f abs_summable_on B" proof (subst asm_rl [of "B = (B-A) \ (A\B)"]) show "B = B - A \ A \ B" by auto have "(\x. 0::real) abs_summable_on B - A" - by simp + by simp moreover have "norm (f x) \ 0" if "x \ B - A" for x :: 'a using that - by (simp add: assms(2)) + by (simp add: assms(2)) ultimately have "f abs_summable_on B - A" - by (rule abs_summable_on_comparison_test' [where g = "\x. 0"]) + by (rule abs_summable_on_comparison_test' [where g = "\x. 0"]) moreover have "f abs_summable_on A \ B" using abs_summable_on_subset assms(1) by blast ultimately show "f abs_summable_on B - A \ A \ B" - by (rule abs_summable_on_union) + by (rule abs_summable_on_union) qed lemma abs_summable_on_Sigma_iff: "f abs_summable_on Sigma A B \ (\x\A. (\y. f (x, y)) abs_summable_on B x) \ ((\x. infsetsum (\y. norm (f (x, y))) (B x)) abs_summable_on A)" proof auto assume sum_AB: "f abs_summable_on Sigma A B" define S' where "S' = {x\Sigma A B. 0 \ f x}" from sum_AB have "countable S'" unfolding S'_def by (rule abs_summable_countable) define A' B' where "A' = fst ` S'" and "B' x = B x \ snd ` S'" for x have A'A: \A' \ A\ and B'B: \B' x \ B x\ for x unfolding A'_def B'_def S'_def by auto have cntA: "countable A'" and cntB: "countable (B' x)" for x unfolding A'_def B'_def using \countable S'\ by auto have f0: "f (x,y) = 0" if "x \ A - A'" and "y \ B x" for x y proof - from that have "(x,y) \ Sigma A B" by auto moreover from that have "(x,y) \ S'" unfolding A'_def - by (metis image_eqI mem_simps(6) prod.sel(1)) + by (metis image_eqI mem_simps(6) prod.sel(1)) ultimately show "f (x,y) = 0" unfolding S'_def by auto qed have f0': "f (x,y) = 0" if "x \ A" and "y \ B x - B' x" for x y proof - from that have "(x,y) \ Sigma A B" by auto moreover from that have "(x,y) \ S'" unfolding B'_def by (auto simp add: rev_image_eqI) ultimately show "f (x,y) = 0" unfolding S'_def by auto qed have "Sigma A' B' \ Sigma A B" using A'A B'B by (rule Sigma_mono) hence sum_A'B': "f abs_summable_on Sigma A' B'" - using sum_AB abs_summable_on_subset by auto + using sum_AB abs_summable_on_subset by auto from sum_A'B' have "(\y. f (x, y)) abs_summable_on B' x" if "x \ A'" for x using abs_summable_on_Sigma_iff[OF cntA cntB, where f=f] that by auto - moreover have "(\y. f (x, y)) abs_summable_on B' x" - if t:"x \ A - A'" + moreover have "(\y. f (x, y)) abs_summable_on B' x" + if t:"x \ A - A'" for x proof (subst abs_summable_on_zero_diff [where A = "{}"]) show "(\y. f (x, y)) abs_summable_on {}" by simp have "f (x, a) = 0" if "a \ B' x" for a using t f0 that B'B by auto thus "f (x, a) = 0" if "a \ B' x - {}" for a - using that by auto + using that by auto show True by blast - qed + qed ultimately have "(\y. f (x, y)) abs_summable_on B' x" if "x \ A" for x using that by auto thus "(\y. f (x, y)) abs_summable_on B x" if "x \ A" for x apply (rule abs_summable_on_zero_diff) using that f0' by auto have Q: "\x. x \ A - A' \ (\\<^sub>ay\B' x. norm (f (x, y))) = 0" apply (subst infsetsum_cong[where g=\\x. 0\ and B="B' _"]) using f0 B'B by auto from sum_A'B' have "(\x. infsetsum (\y. norm (f (x, y))) (B' x)) abs_summable_on A'" using abs_summable_on_Sigma_iff[OF cntA cntB, where f=f] by auto hence "(\x. infsetsum (\y. norm (f (x, y))) (B' x)) abs_summable_on A" apply (rule abs_summable_on_zero_diff) using Q by auto have R: "\x. x \ A \ (\\<^sub>ay\B' x. norm (f (x, y))) = (\\<^sub>ay\B x. norm (f (x, y)))" proof (rule infsetsum_cong_neutral) show "norm (f (x, a)) = 0" if "x \ A" and "a \ B' x - B x" for x :: 'a and a :: 'b - using that B'B by blast + using that B'B by blast show "norm (f (x, a)) = 0" if "x \ A" and "a \ B x - B' x" for x :: 'a and a :: 'b using that - by (simp add: f0') + by (simp add: f0') show "norm (f (x, a)) = norm (f (x, a))" if "x \ A" and "a \ B' x \ B x" for x :: 'a and a :: 'b using that - by simp + by simp qed - thus "(\x. infsetsum (\y. norm (f (x, y))) (B x)) abs_summable_on A" - using \(\x. \\<^sub>ay\B' x. norm (f (x, y))) abs_summable_on A\ by auto + thus "(\x. infsetsum (\y. norm (f (x, y))) (B x)) abs_summable_on A" + using \(\x. \\<^sub>ay\B' x. norm (f (x, y))) abs_summable_on A\ by auto next assume sum_B: "\x\A. (\y. f (x, y)) abs_summable_on B x" assume sum_A: "(\x. \\<^sub>ay\B x. norm (f (x, y))) abs_summable_on A" define B' where "B' x = {y\B x. 0 \ f (x,y)}" for x from sum_B have cnt_B': "countable (B' x)" if "x\A" for x unfolding B'_def apply (rule_tac abs_summable_countable) using that by auto define A' where "A' = {x\A. 0 \ (\\<^sub>ay\B x. norm (f (x, y)))}" from sum_A have cnt_A': "countable A'" unfolding A'_def by (rule abs_summable_countable) have A'A: "A' \ A" and B'B: "B' x \ B x" for x unfolding A'_def B'_def by auto have f0': "f (x,y) = 0" if "y \ B x - B' x" for x y using that unfolding B'_def by auto have f0: "f (x,y) = 0" if "x \ A - A'" and "y \ B x" for x y proof - have "(\\<^sub>ay\B x. norm (f (x, y))) = 0" using that unfolding A'_def by auto hence "norm (f (x, y)) = 0" apply (rule infsetsum_0D) using sum_B that by auto thus ?thesis by auto qed from sum_B have sum_B': "(\y. f (x, y)) abs_summable_on B' x" if "x\A" for x proof (rule_tac abs_summable_on_subset [where B = "B x"]) show "(\y. f (x, y)) abs_summable_on B x" if "\x\A. (\y. f (x, y)) abs_summable_on B x" - using that \x \ A\ by blast + using that \x \ A\ by blast show "B' x \ B x" if "\x\A. (\y. f (x, y)) abs_summable_on B x" using that - by (simp add: B'B) + by (simp add: B'B) qed have *: "(\\<^sub>ay\B x. norm (f (x, y))) = (\\<^sub>ay\B' x. norm (f (x, y)))" if "x\A" for x using infsetsum_cong_neutral f0' B'B that by (metis (no_types, lifting) DiffD1 DiffD2 Int_iff inf.absorb_iff2 norm_zero) have "(\x. \\<^sub>ay\B' x. norm (f (x, y))) abs_summable_on A" using abs_summable_on_cong sum_A "*" by auto hence sum_A': "(\x. \\<^sub>ay\B' x. norm (f (x, y))) abs_summable_on A'" - using _ A'A abs_summable_on_subset by blast + using _ A'A abs_summable_on_subset by blast from sum_A' sum_B' have "f abs_summable_on Sigma A' B'" using abs_summable_on_Sigma_iff[where A=A' and B=B' and f=f, OF cnt_A' cnt_B'] A'A by auto moreover have "f x = 0" if "x \ Sigma A B - Sigma A' B'" for x - using that f0 f0' by force + using that f0 f0' by force ultimately show "f abs_summable_on Sigma A B" by (rule abs_summable_on_zero_diff) qed lemma fixes f :: "'a \ 'c :: {banach, real_normed_field, second_countable_topology}" assumes "f abs_summable_on A" and "g abs_summable_on B" shows abs_summable_on_product: "(\(x,y). f x * g y) abs_summable_on A \ B" and infsetsum_product: "infsetsum (\(x,y). f x * g y) (A \ B) = infsetsum f A * infsetsum g B" proof - from assms show "(\(x,y). f x * g y) abs_summable_on A \ B" by (subst abs_summable_on_Sigma_iff) (auto simp: norm_mult infsetsum_cmult_right) with assms show "infsetsum (\(x,y). f x * g y) (A \ B) = infsetsum f A * infsetsum g B" by (subst infsetsum_Sigma) (auto simp: infsetsum_cmult_left infsetsum_cmult_right) qed lemma infsetsum'_converges_ennreal: \infsetsum'_converges (f::_ \ ennreal) S\ proof - define B where \B = (SUP F\{F. F \ S \ finite F}. sum f F)\ have upper: \\\<^sub>F F in finite_subsets_at_top S. sum f F \ B\ apply (rule eventually_finite_subsets_at_top_weakI) unfolding B_def by (simp add: SUP_upper) have lower: \\\<^sub>F n in finite_subsets_at_top S. x < sum f n\ if \x < B\ for x proof - obtain F where Fx: \sum f F > x\ and \F \ S\ and \finite F\ using \x < B\ unfolding B_def by (metis (mono_tags, lifting) less_SUP_iff mem_Collect_eq) have geq: \sum f Y \ sum f F\ if \finite Y\ and \Y \ F\ for Y by (simp add: sum_mono2 that(1) that(2)) show ?thesis unfolding eventually_finite_subsets_at_top apply (rule exI[of _ F]) using \finite F\ \F \ S\ Fx geq by force qed show ?thesis unfolding infsetsum'_converges_def apply (rule exI[of _ B]) using upper lower by (rule increasing_tendsto) qed lemma infsetsum'_superconst_infinite: assumes geqb: \\x. x \ S \ f x \ b\ assumes b: \b > 0\ assumes \infinite S\ shows "infsetsum' f S = (\::ennreal)" proof - have \(sum f \ \) (finite_subsets_at_top S)\ proof (rule order_tendstoI[rotated], simp) fix y :: ennreal assume \y < \\ then have \y / b < \\ by (metis b ennreal_divide_eq_top_iff gr_implies_not_zero infinity_ennreal_def top.not_eq_extremum) then obtain F where \finite F\ and \F \ S\ and cardF: \card F > y / b\ using \infinite S\ by (metis ennreal_Ex_less_of_nat infinite_arbitrarily_large infinity_ennreal_def) moreover have \sum f Y > y\ if \finite Y\ and \F \ Y\ and \Y \ S\ for Y proof - have \y < b * card F\ by (metis \y < \\ b cardF divide_less_ennreal ennreal_mult_eq_top_iff gr_implies_not_zero infinity_ennreal_def mult.commute top.not_eq_extremum) also have \\ \ b * card Y\ by (meson b card_mono less_imp_le mult_left_mono of_nat_le_iff that(1) that(2)) also have \\ = sum (\_. b) Y\ by (simp add: mult.commute) also have \\ \ sum f Y\ using geqb by (meson subset_eq sum_mono that(3)) finally show ?thesis by - qed ultimately show \\\<^sub>F x in finite_subsets_at_top S. y < sum f x\ - unfolding eventually_finite_subsets_at_top + unfolding eventually_finite_subsets_at_top by auto qed then show ?thesis - unfolding infsetsum'_def + unfolding infsetsum'_def apply (simp add: infsetsum'_converges_ennreal) by (simp add: tendsto_Lim) qed lemma infsetsum'_tendsto: assumes \infsetsum'_converges f S\ shows \((\F. sum f F) \ infsetsum' f S) (finite_subsets_at_top S)\ by (metis assms finite_subsets_at_top_neq_bot infsetsum'_converges_def infsetsum'_def tendsto_Lim) lemma infsetsum'_constant[simp]: assumes \finite F\ shows \infsetsum' (\_. c) F = of_nat (card F) * c\ apply (subst infsetsum'_finite[OF assms]) by simp lemma infsetsum'_zero[simp]: shows \infsetsum' (\_. 0) S = 0\ unfolding infsetsum'_def sum.neutral_const by (simp add: tendsto_Lim) lemma fixes f g :: "'a \ 'b::{topological_monoid_add, t2_space, comm_monoid_add}" assumes \infsetsum'_converges f A\ assumes \infsetsum'_converges g A\ shows infsetsum'_add: \infsetsum' (\x. f x + g x) A = infsetsum' f A + infsetsum' g A\ and infsetsum'_converges_add: \infsetsum'_converges (\x. f x + g x) A\ proof - note lim_f = infsetsum'_tendsto[OF assms(1)] and lim_g = infsetsum'_tendsto[OF assms(2)] then have lim: \(sum (\x. f x + g x) \ infsetsum' f A + infsetsum' g A) (finite_subsets_at_top A)\ unfolding sum.distrib by (rule tendsto_add) then show conv: \infsetsum'_converges (\x. f x + g x) A\ unfolding infsetsum'_converges_def by auto show \infsetsum' (\x. f x + g x) A = infsetsum' f A + infsetsum' g A\ - unfolding infsetsum'_def + unfolding infsetsum'_def using lim_f lim_g lim by (auto simp: assms conv tendsto_Lim) qed -lemma +lemma fixes f :: "'a \ 'b::{topological_monoid_add, t2_space, comm_monoid_add}" assumes "infsetsum'_converges f A" assumes "infsetsum'_converges f B" assumes disj: "A \ B = {}" shows infsetsum'_Un_disjoint: \infsetsum' f (A \ B) = infsetsum' f A + infsetsum' f B\ and infsetsum'_converges_Un_disjoint: \infsetsum'_converges f (A \ B)\ proof - define fA fB where \fA x = (if x \ A then f x else 0)\ and \fB x = (if x \ A then f x else 0)\ for x have conv_fA: \infsetsum'_converges fA (A \ B)\ unfolding fA_def apply (subst infsetsum'_converges_subset_zero, auto) by (simp add: assms(1) infsetsum'_converges_cong) have conv_fB: \infsetsum'_converges fB (A \ B)\ unfolding fB_def apply (subst infsetsum'_converges_subset_zero, auto) by (smt (verit, ccfv_SIG) assms(2) assms(3) disjoint_iff infsetsum'_converges_cong) have fAB: \f x = fA x + fB x\ for x unfolding fA_def fB_def by simp have \infsetsum' f (A \ B) = infsetsum' fA (A \ B) + infsetsum' fB (A \ B)\ unfolding fAB using conv_fA conv_fB by (rule infsetsum'_add) also have \\ = infsetsum' fA A + infsetsum' fB B\ unfolding fA_def fB_def by (subst infsetsum'_subset_zero[where A=\A\B\], auto)+ also have \\ = infsetsum' f A + infsetsum' f B\ apply (subst infsetsum'_cong[where f=fA and g=f], simp add: fA_def) apply (subst infsetsum'_cong[where f=fB and g=f], simp add: fB_def) using disj by auto finally show \infsetsum' f (A \ B) = infsetsum' f A + infsetsum' f B\ by - from conv_fA conv_fB have \infsetsum'_converges (\x. fA x + fB x) (A \ B)\ by (rule infsetsum'_converges_add) then show \infsetsum'_converges f (A \ B)\ unfolding fAB by - qed lemma infsetsum'_converges_union_disjoint: fixes f :: "'a \ 'b::{topological_monoid_add, t2_space, comm_monoid_add}" assumes finite: \finite A\ assumes conv: \\a. a \ A \ infsetsum'_converges f (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \infsetsum'_converges f (\a\A. B a)\ using finite conv disj apply induction by (auto intro!: infsetsum'_converges_Un_disjoint) lemma sum_infsetsum': fixes f :: "'a \ 'b::{topological_monoid_add, t2_space, comm_monoid_add}" assumes finite: \finite A\ assumes conv: \\a. a \ A \ infsetsum'_converges f (B a)\ assumes disj: \\a a'. a\A \ a'\A \ a\a' \ B a \ B a' = {}\ shows \sum (\a. infsetsum' f (B a)) A = infsetsum' f (\a\A. B a)\ using assms proof (insert finite conv disj, induction) case empty - then show ?case + then show ?case by simp next case (insert x F) have \(\a\insert x F. infsetsum' f (B a)) = infsetsum' f (B x) + (\a\F. infsetsum' f (B a))\ apply (subst sum.insert) using insert by auto also have \\ = infsetsum' f (B x) + infsetsum' f (\ (B ` F))\ apply (subst insert.IH) using assms insert by auto also have \\ = infsetsum' f (B x \ \ (B ` F))\ apply (rule infsetsum'_Un_disjoint[symmetric]) using insert.prems insert.hyps by (auto simp: infsetsum'_converges_union_disjoint) also have \\ = infsetsum' f (\a\insert x F. B a)\ by auto finally show ?case by - qed theorem infsetsum'_mono: fixes f g :: "'a\'b::{ordered_comm_monoid_add,linorder_topology}" assumes "infsetsum'_converges f A" and "infsetsum'_converges g A" assumes leq: "\x. x \ A \ f x \ g x" shows "infsetsum' f A \ infsetsum' g A" proof - note limf = infsetsum'_tendsto[OF assms(1)] and limg = infsetsum'_tendsto[OF assms(2)] have sum_leq: \\F. finite F \ F \ A \ sum f F \ sum g F\ by (simp add: in_mono leq sum_mono) show ?thesis using _ limg limf apply (rule tendsto_le) by (auto intro!: sum_leq) qed end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy b/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Lattice.thy @@ -1,271 +1,271 @@ section \\Extra_Lattice\ -- Additional results about lattices\ theory Extra_Lattice imports Main begin subsection\\Lattice_Missing\ -- Miscellaneous missing facts about lattices\ text \Two bundles to activate and deactivate lattice specific notation (e.g., \\\ etc.). Activate the notation locally via "@{theory_text \includes lattice_notation\}" in a lemma statement. (Or sandwich a declaration using that notation between "@{theory_text \unbundle lattice_notation ... unbundle no_lattice_notation\}.)\ bundle lattice_notation begin notation inf (infixl "\" 70) notation sup (infixl "\" 65) notation Inf ("\") notation Sup ("\") notation bot ("\") notation top ("\") end bundle no_lattice_notation begin notation inf (infixl "\" 70) notation sup (infixl "\" 65) notation Inf ("\") notation Sup ("\") notation bot ("\") notation top ("\") end unbundle lattice_notation text \The following class \complemented_lattice\ describes complemented lattices (with - \<^const>\uminus\ for the complement). The definition follows + \<^const>\uminus\ for the complement). The definition follows \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Definition_and_basic_properties\. - Additionally, it adopts the convention from \<^class>\boolean_algebra\ of defining + Additionally, it adopts the convention from \<^class>\boolean_algebra\ of defining \<^const>\minus\ in terms of the complement.\ -class complemented_lattice = bounded_lattice + uminus + minus + +class complemented_lattice = bounded_lattice + uminus + minus + assumes inf_compl_bot[simp]: "inf x (-x) = bot" and sup_compl_top[simp]: "sup x (-x) = top" and diff_eq: "x - y = inf x (- y)" begin lemma dual_complemented_lattice: "class.complemented_lattice (\x y. x \ (- y)) uminus sup greater_eq greater inf \ \" proof (rule class.complemented_lattice.intro) show "class.bounded_lattice (\) (\x y. (y::'a) \ x) (\x y. y < x) (\) \ \" by (rule dual_bounded_lattice) show "class.complemented_lattice_axioms (\x y. (x::'a) \ - y) uminus (\) (\) \ \" by (unfold_locales, auto simp add: diff_eq) qed lemma compl_inf_bot [simp]: "inf (- x) x = bot" by (simp add: inf_commute) lemma compl_sup_top [simp]: "sup (- x) x = top" by (simp add: sup_commute) end -class complete_complemented_lattice = complemented_lattice + complete_lattice +class complete_complemented_lattice = complemented_lattice + complete_lattice text \The following class \complemented_lattice\ describes orthocomplemented lattices, following \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Orthocomplementation\.\ class orthocomplemented_lattice = complemented_lattice + assumes ortho_involution[simp]: "- (- x) = x" and ortho_antimono: "x \ y \ -x \ -y" begin lemma dual_orthocomplemented_lattice: "class.orthocomplemented_lattice (\x y. x \ - y) uminus sup greater_eq greater inf \ \" proof (rule class.orthocomplemented_lattice.intro) show "class.complemented_lattice (\x y. (x::'a) \ - y) uminus (\) (\x y. y \ x) (\x y. y < x) (\) \ \" by (rule dual_complemented_lattice) show "class.orthocomplemented_lattice_axioms uminus (\x y. (y::'a) \ x)" by (unfold_locales, auto simp add: diff_eq intro: ortho_antimono) qed lemma compl_eq_compl_iff [simp]: "- x = - y \ x = y" by (metis ortho_involution) lemma compl_bot_eq [simp]: "- bot = top" by (metis inf_compl_bot inf_top_left ortho_involution) lemma compl_top_eq [simp]: "- top = bot" using compl_bot_eq ortho_involution by blast text \De Morgan's law\ (* Proof from: https://planetmath.org/orthocomplementedlattice *) lemma compl_sup [simp]: "- (x \ y) = - x \ - y" proof - have "- (x \ y) \ - x" by (simp add: ortho_antimono) moreover have "- (x \ y) \ - y" by (simp add: ortho_antimono) ultimately have 1: "- (x \ y) \ - x \ - y" by (simp add: sup.coboundedI1) have \x \ - (-x \ -y)\ by (metis inf.cobounded1 ortho_antimono ortho_involution) moreover have \y \ - (-x \ -y)\ by (metis inf.cobounded2 ortho_antimono ortho_involution) ultimately have \x \ y \ - (-x \ -y)\ by auto hence 2: \-x \ -y \ - (x \ y)\ using ortho_antimono by fastforce from 1 2 show ?thesis - by (simp add: eq_iff) + using dual_order.antisym by blast qed text \De Morgan's law\ lemma compl_inf [simp]: "- (x \ y) = - x \ - y" using compl_sup by (metis ortho_involution) lemma compl_mono: assumes "x \ y" shows "- y \ - x" by (simp add: assms local.ortho_antimono) lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" by (auto dest: compl_mono) lemma compl_le_swap1: assumes "y \ - x" shows "x \ -y" using assms ortho_antimono by fastforce lemma compl_le_swap2: assumes "- y \ x" shows "- x \ y" using assms local.ortho_antimono by fastforce lemma compl_less_compl_iff[simp]: "- x < - y \ y < x" by (auto simp add: less_le) lemma compl_less_swap1: assumes "y < - x" shows "x < - y" using assms compl_less_compl_iff by fastforce lemma compl_less_swap2: assumes "- y < x" shows "- x < y" using assms compl_le_swap1 compl_le_swap2 less_le_not_le by auto lemma sup_cancel_left1: "sup (sup x a) (sup (- x) b) = top" by (simp add: sup_commute sup_left_commute) lemma sup_cancel_left2: "sup (sup (- x) a) (sup x b) = top" by (simp add: sup.commute sup_left_commute) lemma inf_cancel_left1: "inf (inf x a) (inf (- x) b) = bot" by (simp add: inf.left_commute inf_commute) lemma inf_cancel_left2: "inf (inf (- x) a) (inf x b) = bot" using inf.left_commute inf_commute by auto lemma sup_compl_top_left1 [simp]: "sup (- x) (sup x y) = top" by (simp add: sup_assoc[symmetric]) lemma sup_compl_top_left2 [simp]: "sup x (sup (- x) y) = top" using sup_compl_top_left1[of "- x" y] by simp lemma inf_compl_bot_left1 [simp]: "inf (- x) (inf x y) = bot" by (simp add: inf_assoc[symmetric]) lemma inf_compl_bot_left2 [simp]: "inf x (inf (- x) y) = bot" using inf_compl_bot_left1[of "- x" y] by simp lemma inf_compl_bot_right [simp]: "inf x (inf y (- x)) = bot" by (subst inf_left_commute) simp end class complete_orthocomplemented_lattice = orthocomplemented_lattice + complete_lattice instance complete_orthocomplemented_lattice \ complete_complemented_lattice by intro_classes text \The following class \orthomodular_lattice\ describes orthomodular lattices, following \<^url>\https://en.wikipedia.org/wiki/Complemented_lattice#Orthomodular_lattices\.\ class orthomodular_lattice = orthocomplemented_lattice + assumes orthomodular: "x \ y \ sup x (inf (-x) y) = y" begin lemma dual_orthomodular_lattice: "class.orthomodular_lattice (\x y. x \ - y) uminus sup greater_eq greater inf \ \" proof (rule class.orthomodular_lattice.intro) show "class.orthocomplemented_lattice (\x y. (x::'a) \ - y) uminus (\) (\x y. y \ x) (\x y. y < x) (\) \ \" by (rule dual_orthocomplemented_lattice) show "class.orthomodular_lattice_axioms uminus (\) (\x y. (y::'a) \ x) (\)" proof (unfold_locales) show "(x::'a) \ (- x \ y) = y" if "(y::'a) \ x" for x :: 'a and y :: 'a using that local.compl_eq_compl_iff local.ortho_antimono local.orthomodular by fastforce qed qed end class complete_orthomodular_lattice = orthomodular_lattice + complete_lattice begin end instance complete_orthomodular_lattice \ complete_orthocomplemented_lattice by intro_classes instance boolean_algebra \ orthomodular_lattice proof - fix x y :: 'a + fix x y :: 'a show "sup (x::'a) (inf (- x) y) = y" if "(x::'a) \ y" using that - by (simp add: sup.absorb_iff2 sup_inf_distrib1) + by (simp add: sup.absorb_iff2 sup_inf_distrib1) show "x - y = inf x (- y)" by (simp add: boolean_algebra_class.diff_eq) qed auto instance complete_boolean_algebra \ complete_orthomodular_lattice by intro_classes lemma image_of_maximum: fixes f::"'a::order \ 'b::conditionally_complete_lattice" assumes "mono f" and "\x. x:M \ x\m" and "m:M" shows "(SUP x\M. f x) = f m" by (smt (verit, ccfv_threshold) assms(1) assms(2) assms(3) cSup_eq_maximum imageE imageI monoD) lemma cSup_eq_cSup: fixes A B :: \'a::conditionally_complete_lattice set\ assumes bdd: \bdd_above A\ assumes B: \\a. a\A \ \b\B. b \ a\ assumes A: \\b. b\B \ \a\A. a \ b\ shows \Sup A = Sup B\ proof (cases \B = {}\) case True with A B have \A = {}\ by auto with True show ?thesis by simp next case False have \bdd_above B\ by (meson A bdd bdd_above_def order_trans) have \A \ {}\ using A False by blast moreover have \a \ Sup B\ if \a \ A\ for a proof - obtain b where \b \ B\ and \b \ a\ using B \a \ A\ by auto then show ?thesis apply (rule cSup_upper2) using \bdd_above B\ by simp qed moreover have \Sup B \ c\ if \\a. a \ A \ a \ c\ for c using False apply (rule cSup_least) using A that by fastforce ultimately show ?thesis by (rule cSup_eq_non_empty) qed unbundle no_lattice_notation end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Operator_Norm.thy b/thys/Complex_Bounded_Operators/extra/Extra_Operator_Norm.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Operator_Norm.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Operator_Norm.thy @@ -1,392 +1,392 @@ section \\Extra_Operator_Norm\ -- Additional facts bout the operator norm\ theory Extra_Operator_Norm imports "HOL-Analysis.Operator_Norm" Extra_General "HOL-Analysis.Bounded_Linear_Function" begin text \This theorem complements \<^theory>\HOL-Analysis.Operator_Norm\ additional useful facts about operator norms.\ -lemma ex_norm1: +lemma ex_norm1: assumes \(UNIV::'a::real_normed_vector set) \ {0}\ shows \\x::'a. norm x = 1\ proof- have \\x::'a. x \ 0\ using assms by fastforce then obtain x::'a where \x \ 0\ by blast hence \norm x \ 0\ by simp hence \(norm x) / (norm x) = 1\ by simp moreover have \(norm x) / (norm x) = norm (x /\<^sub>R (norm x))\ by simp ultimately have \norm (x /\<^sub>R (norm x)) = 1\ by simp thus ?thesis - by blast + by blast qed lemma bdd_above_norm_f: assumes "bounded_linear f" shows \bdd_above {norm (f x) |x. norm x = 1}\ proof- have \\M. \x. norm x = 1 \ norm (f x) \ M\ using assms by (metis bounded_linear.axioms(2) bounded_linear_axioms_def) thus ?thesis by auto qed lemma onorm_sphere: fixes f :: "'a::{real_normed_vector, not_singleton} \ 'b::real_normed_vector" assumes a1: "bounded_linear f" shows \onorm f = Sup {norm (f x) | x. norm x = 1}\ proof(cases \f = (\ _. 0)\) case True have \(UNIV::'a set) \ {0}\ by simp hence \\x::'a. norm x = 1\ using ex_norm1 by blast have \norm (f x) = 0\ for x - by (simp add: True) + by (simp add: True) hence \{norm (f x) | x. norm x = 1} = {0}\ using \\x. norm x = 1\ by auto hence v1: \Sup {norm (f x) | x. norm x = 1} = 0\ - by simp + by simp have \onorm f = 0\ - by (simp add: True onorm_eq_0) + by (simp add: True onorm_eq_0) thus ?thesis using v1 by simp next case False have \y \ {norm (f x) |x. norm x = 1} \ {0}\ if "y \ {norm (f x) / norm x |x. True}" for y proof(cases \y = 0\) case True thus ?thesis - by simp + by simp next case False have \\ x. y = norm (f x) / norm x\ using \y \ {norm (f x) / norm x |x. True}\ by auto then obtain x where \y = norm (f x) / norm x\ by blast hence \y = \(1/norm x)\ * norm ( f x )\ by simp hence \y = norm ( (1/norm x) *\<^sub>R f x )\ by simp hence \y = norm ( f ((1/norm x) *\<^sub>R x) )\ apply (subst linear_cmul[of f]) by (simp_all add: assms bounded_linear.linear) moreover have \norm ((1/norm x) *\<^sub>R x) = 1\ - using False \y = norm (f x) / norm x\ by auto + using False \y = norm (f x) / norm x\ by auto ultimately have \y \ {norm (f x) |x. norm x = 1}\ by blast thus ?thesis by blast qed moreover have "y \ {norm (f x) / norm x |x. True}" if \y \ {norm (f x) |x. norm x = 1} \ {0}\ for y proof(cases \y = 0\) case True thus ?thesis - by auto + by auto next case False hence \y \ {0}\ by simp hence \y \ {norm (f x) |x. norm x = 1}\ - using that by auto + using that by auto hence \\ x. norm x = 1 \ y = norm (f x)\ by auto then obtain x where \norm x = 1\ and \y = norm (f x)\ by auto have \y = norm (f x) / norm x\ using \norm x = 1\ \y = norm (f x)\ - by simp + by simp thus ?thesis - by auto + by auto qed - ultimately have \{norm (f x) / norm x |x. True} = {norm (f x) |x. norm x = 1} \ {0}\ + ultimately have \{norm (f x) / norm x |x. True} = {norm (f x) |x. norm x = 1} \ {0}\ by blast hence \Sup {norm (f x) / norm x |x. True} = Sup ({norm (f x) |x. norm x = 1} \ {0})\ by simp moreover have \Sup {norm (f x) |x. norm x = 1} \ 0\ proof- have \\ x::'a. norm x = 1\ - by (metis (mono_tags, hide_lams) False assms bounded_linear.nonneg_bounded mult_zero_left norm_le_zero_iff norm_sgn) + by (metis (full_types) False assms linear_simps(3) norm_sgn) then obtain x::'a where \norm x = 1\ by blast have \norm (f x) \ 0\ by simp hence \\ x::'a. norm x = 1 \ norm (f x) \ 0\ using \norm x = 1\ by blast hence \\ y \ {norm (f x) |x. norm x = 1}. y \ 0\ by blast - then obtain y::real where \y \ {norm (f x) |x. norm x = 1}\ + then obtain y::real where \y \ {norm (f x) |x. norm x = 1}\ and \y \ 0\ by auto have \{norm (f x) |x. norm x = 1} \ {}\ - using \y \ {norm (f x) |x. norm x = 1}\ by blast + using \y \ {norm (f x) |x. norm x = 1}\ by blast moreover have \bdd_above {norm (f x) |x. norm x = 1}\ using bdd_above_norm_f - by (metis (mono_tags, lifting) a1) + by (metis (mono_tags, lifting) a1) ultimately have \y \ Sup {norm (f x) |x. norm x = 1}\ using \y \ {norm (f x) |x. norm x = 1}\ - by (simp add: cSup_upper) + by (simp add: cSup_upper) thus ?thesis using \y \ 0\ by simp qed moreover have \Sup ({norm (f x) |x. norm x = 1} \ {0}) = Sup {norm (f x) |x. norm x = 1}\ proof- have \{norm (f x) |x. norm x = 1} \ {}\ by (simp add: assms(1) ex_norm1) moreover have \bdd_above {norm (f x) |x. norm x = 1}\ using a1 bdd_above_norm_f by force have \{0::real} \ {}\ by simp moreover have \bdd_above {0::real}\ by simp ultimately have \Sup ({norm (f x) |x. norm x = 1} \ {(0::real)}) = max (Sup {norm (f x) |x. norm x = 1}) (Sup {0::real})\ by (metis (lifting) \0 \ Sup {norm (f x) |x. norm x = 1}\ \bdd_above {0}\ \bdd_above {norm (f x) |x. norm x = 1}\ \{0} \ {}\ \{norm (f x) |x. norm x = 1} \ {}\ cSup_singleton cSup_union_distrib max.absorb_iff1 sup.absorb_iff1) moreover have \Sup {(0::real)} = (0::real)\ - by simp + by simp moreover have \Sup {norm (f x) |x. norm x = 1} \ 0\ by (simp add: \0 \ Sup {norm (f x) |x. norm x = 1}\) ultimately show ?thesis by simp qed moreover have \Sup ( {norm (f x) |x. norm x = 1} \ {0}) = max (Sup {norm (f x) |x. norm x = 1}) (Sup {0}) \ using calculation(2) calculation(3) by auto ultimately have w1: "Sup {norm (f x) / norm x | x. True} = Sup {norm (f x) | x. norm x = 1}" - by simp + by simp have \(SUP x. norm (f x) / (norm x)) = Sup {norm (f x) / norm x | x. True}\ by (simp add: full_SetCompr_eq) also have \... = Sup {norm (f x) | x. norm x = 1}\ using w1 by auto ultimately have \(SUP x. norm (f x) / (norm x)) = Sup {norm (f x) | x. norm x = 1}\ by linarith thus ?thesis unfolding onorm_def by blast qed lemma onorm_Inf_bound: fixes f :: \'a::{real_normed_vector,not_singleton} \ 'b::real_normed_vector\ assumes a1: "bounded_linear f" shows "onorm f = Inf {K. (\x\0. norm (f x) \ norm x * K)}" proof- have a2: \(UNIV::'a set) \ {0}\ by simp define A where \A = {norm (f x) / (norm x) | x. x \ 0}\ have \A \ {}\ proof- have \\ x::'a. x \ 0\ using a2 by auto thus ?thesis using A_def - by simp + by simp qed moreover have \bdd_above A\ proof- have \\ M. \ x. norm (f x) / (norm x) \ M\ using \bounded_linear f\ le_onorm by auto thus ?thesis using A_def - by auto + by auto qed - ultimately have \Sup A = Inf {b. \a\A. a \ b}\ - by (simp add: cSup_cInf) + ultimately have \Sup A = Inf {b. \a\A. a \ b}\ + by (simp add: cSup_cInf) moreover have \{b. \a\A. a \ b} = {K. (\x\0. norm (f x)/ norm x \ K)}\ proof- have \{b. \a\A. a \ b} = {b. \a\{norm (f x) / (norm x) | x. x \ 0}. a \ b}\ using A_def by blast also have \... = {b. \x\{x | x. x \ 0}. norm (f x) / (norm x) \ b}\ by auto also have \... = {b. \x\0. norm (f x) / (norm x) \ b}\ by auto finally show ?thesis by blast qed - ultimately have \Sup {norm (f x) / (norm x) | x. x \ 0} + ultimately have \Sup {norm (f x) / (norm x) | x. x \ 0} = Inf {K. (\x\0. norm (f x)/ norm x \ K)}\ using A_def - by simp + by simp moreover have \(\x\0. norm (f x) \ norm x * K) \ (\x\0. norm (f x)/ norm x \ K)\ for K proof show "\x\0. norm (f x) / norm x \ K" if "\x\0. norm (f x) \ norm x * K" using divide_le_eq nonzero_mult_div_cancel_left norm_le_zero_iff that by (simp add: divide_le_eq mult.commute) show "\x\0. norm (f x) \ norm x * K" if "\x\0. norm (f x) / norm x \ K" using divide_le_eq nonzero_mult_div_cancel_left norm_le_zero_iff that by (simp add: divide_le_eq mult.commute) qed ultimately have f1: \Sup {norm (f x) / (norm x) | x. x \ 0} = Inf {K. (\x\0. norm (f x) \ norm x * K)}\ by simp - moreover + moreover have t1: \{norm (f x) / (norm x) | x. x \ 0} \ {norm (f x) / (norm x) | x. x = 0} = {norm (f x) / (norm x) | x. True}\ using Collect_cong by blast have \{norm (f x) / (norm x) | x. x \ 0} \ {}\ proof- have \\ x::'a. x \ 0\ using \UNIV\{0}\ by auto thus ?thesis - by simp + by simp qed moreover have \bdd_above {norm (f x) / (norm x) | x. x \ 0}\ proof- have \\ M. \ x. norm (f x) / (norm x) \ M\ - using \bounded_linear f\ bounded_linear.nonneg_bounded + using \bounded_linear f\ bounded_linear.nonneg_bounded mult_divide_mult_cancel_left_if norm_zero real_divide_square_eq using le_onorm by blast thus ?thesis - by auto + by auto qed moreover have \{norm (f x) / (norm x) | x. x = 0} \ {}\ by simp moreover have \bdd_above {norm (f x) / (norm x) | x. x = 0}\ by simp - ultimately + ultimately have d1: \Sup ({norm (f x) / (norm x) | x. x \ 0} \ {norm (f x) / (norm x) | x. x = 0}) = max (Sup {norm (f x) / (norm x) | x. x \ 0}) (Sup {norm (f x) / (norm x) | x. x = 0})\ by (metis (no_types, lifting) cSup_union_distrib sup_max) have g1: \Sup {norm (f x) / (norm x) | x. x \ 0} \ 0\ proof- have t2: \{norm (f x) / (norm x) | x. x \ 0} \ {}\ proof- have \\ x::'a. x \ 0\ using \UNIV\{0}\ by auto - thus ?thesis + thus ?thesis by auto qed have \\ M. \ x. norm (f x) / (norm x) \ M\ using \bounded_linear f\ by (metis \\K. (\x. x \ 0 \ norm (f x) \ norm x * K) = (\x. x \ 0 \ norm (f x) / norm x \ K)\ bounded_linear.nonneg_bounded mult_divide_mult_cancel_left_if norm_zero real_divide_square_eq) hence t3: \bdd_above {norm (f x) / (norm x) | x. x \ 0}\ by auto have \norm (f x) / (norm x) \ 0\ for x by simp hence \\ y\{norm (f x) / (norm x) | x. x \ 0}. y \ 0\ by blast show ?thesis by (metis (lifting) \\y\{norm (f x) / norm x |x. x \ 0}. 0 \ y\ \bdd_above {norm (f x) / norm x |x. x \ 0}\ \{norm (f x) / norm x |x. x \ 0} \ {}\ bot.extremum_uniqueI cSup_upper2 subset_emptyI) qed - hence r: \Sup ({norm (f x) / (norm x) | x. x \ 0} \ {norm (f x) / (norm x) | x. x = 0}) + hence r: \Sup ({norm (f x) / (norm x) | x. x \ 0} \ {norm (f x) / (norm x) | x. x = 0}) = Sup {norm (f x) / (norm x) | x. True}\ using t1 by auto have \{norm (f x) / (norm x) | x. x = 0} = {norm (f 0) / (norm 0)}\ by simp hence \Sup {norm (f x) / (norm x) | x. x = 0} = 0\ by simp have h1: \Sup {norm (f x) / (norm x) | x. x \ 0} = Sup {norm (f x) / (norm x) | x. True}\ - using d1 r g1 by auto + using d1 r g1 by auto have \(SUP x. norm (f x) / (norm x)) = Inf {K. (\x\0. norm (f x) \ norm x * K)}\ using full_SetCompr_eq by (metis f1 h1) thus ?thesis by (simp add: onorm_def) qed lemma onormI: assumes "\x. norm (f x) \ b * norm x" and "x \ 0" and "norm (f x) = b * norm x" shows "onorm f = b" apply (unfold onorm_def, rule cSup_eq_maximum) apply (smt (verit) UNIV_I assms(2) assms(3) image_iff nonzero_mult_div_cancel_right norm_eq_zero) by (smt (verit, del_insts) assms(1) assms(2) divide_nonneg_nonpos norm_ge_zero norm_le_zero_iff pos_divide_le_eq rangeE zero_le_mult_iff) lemma norm_unit_sphere: fixes f::\'a::{real_normed_vector,not_singleton} \\<^sub>L 'b::real_normed_vector\ - assumes a1: "bounded_linear f" and a2: "e > 0" + assumes a1: "bounded_linear f" and a2: "e > 0" shows \\x\(sphere 0 1). norm (norm (blinfun_apply f x) - norm f) < e\ proof- define S::"real set" where \S = { norm (f x)| x. x \ sphere 0 1 }\ have "\x::'a. norm x = 1" by (metis (full_types) Collect_empty_eq Extra_General.UNIV_not_singleton UNIV_I equalityI mem_Collect_eq norm_sgn singleton_conv subsetI) hence \\x::'a. x \ sphere 0 1\ - by simp - hence \S\{}\unfolding S_def - by auto + by simp + hence \S\{}\unfolding S_def + by auto hence t1: \e > 0 \ \ y \ S. Sup S - e < y\ for e by (simp add: less_cSupD) have \onorm f = Sup { norm (f x)| x. norm x = 1 }\ using \bounded_linear f\ onorm_sphere - by auto + by auto hence \onorm f = Sup { norm (f x)| x. x \ sphere 0 1 }\ unfolding sphere_def by simp - hence t2: \Sup S = onorm f\ unfolding S_def + hence t2: \Sup S = onorm f\ unfolding S_def by auto have s1: \\y\{norm (f x) |x. x \ sphere 0 1}. norm (onorm f - y) < e\ if "0 < e" for e proof- have \\ y \ S. (onorm f) - e < y\ using t1 t2 that by auto hence \\ y \ S. (onorm f) - y < e\ using that by force have \\ y \ S. (onorm f) - y < e\ using \0 < e\ \\y\S. onorm f - y < e\ by auto then obtain y where \y \ S\ and \(onorm f) - y < e\ by blast have \y \ {norm (f x) |x. x \ sphere 0 1} \ y \ onorm f\ proof- assume \y \ {norm (f x) |x. x \ sphere 0 1}\ hence \\ x \ sphere 0 1. y = norm (f x)\ by blast then obtain x where \x \ sphere 0 1\ and \y = norm (f x)\ by blast from \y = norm (f x)\ have \y \ onorm f * norm x\ using a1 onorm by auto moreover have \norm x = 1\ using \x \ sphere 0 1\ unfolding sphere_def by auto ultimately show ?thesis by simp qed hence \bdd_above {norm (f x) |x. x \ sphere 0 1}\ using a1 bdd_above_norm_f by force - hence \bdd_above S\ unfolding S_def + hence \bdd_above S\ unfolding S_def by blast hence \y \ Sup S\ using \y \ S\ \S \ {}\ cSup_upper by blast hence \0 \ Sup S - y\ by simp hence \0 \ onorm f - y\ using \Sup S = onorm f\ by simp hence \\ (onorm f - y) \ = onorm f - y\ by simp hence \norm (onorm f - y) = onorm f - y\ by auto hence \\ y \ S. norm ((onorm f) - y) < e\ - using \onorm f - y < e\ \y \ S\ by force + using \onorm f - y < e\ \y \ S\ by force show ?thesis unfolding S_def - using S_def \\y\S. norm (onorm (blinfun_apply f) - y) < e\ by blast + using S_def \\y\S. norm (onorm (blinfun_apply f) - y) < e\ by blast qed have f2: "onorm (blinfun_apply f) = Sup S" using S_def \onorm (blinfun_apply f) = Sup {norm (blinfun_apply f x) |x. x \ sphere 0 1}\ by blast hence "\a. norm (norm (blinfun_apply f a) - Sup S) < e \ a \ sphere 0 1" - using a1 a2 s1 a2 t2 - by force + using a1 a2 s1 a2 t2 + by force thus ?thesis - using f2 by (metis (full_types) norm_blinfun.rep_eq) + using f2 by (metis (full_types) norm_blinfun.rep_eq) qed end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy b/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy @@ -1,958 +1,958 @@ section \\Extra_Ordered_Fields\ -- Additional facts about ordered fields\ theory Extra_Ordered_Fields - imports Complex_Main + imports Complex_Main Jordan_Normal_Form.Conjugate (* Defines ordering for complex. We have to use theirs, otherwise there will be conflicts *) begin subsection\Ordered Fields\ text \In this section we introduce some type classes for ordered rings/fields/etc. -that are weakenings of existing classes. Most theorems in this section are -copies of the eponymous theorems from Isabelle/HOL, except that they are now proven +that are weakenings of existing classes. Most theorems in this section are +copies of the eponymous theorems from Isabelle/HOL, except that they are now proven requiring weaker type classes (usually the need for a total order is removed). -Since the lemmas are identical to the originals except for weaker type constraints, +Since the lemmas are identical to the originals except for weaker type constraints, we use the same names as for the original lemmas. (In fact, the new lemmas could replace the original ones in Isabelle/HOL with at most minor incompatibilities.\ subsection \Missing from Orderings.thy\ text \This class is analogous to \<^class>\unbounded_dense_linorder\, except that it does not require a total order\ class unbounded_dense_order = dense_order + no_top + no_bot instance unbounded_dense_linorder \ unbounded_dense_order .. subsection \Missing from Rings.thy\ text \The existing class \<^class>\abs_if\ requires \<^term>\\a\ = (if a < 0 then - a else a)\. -However, if \<^term>\(<)\ is not a total order, this condition is too strong when \<^term>\a\ +However, if \<^term>\(<)\ is not a total order, this condition is too strong when \<^term>\a\ is incomparable with \<^term>\0\. (Namely, it requires the absolute value to be -the identity on such elements. E.g., the absolute value for complex numbers does not +the identity on such elements. E.g., the absolute value for complex numbers does not satisfy this.) The following class \partial_abs_if\ is analogous to \<^class>\abs_if\ but does not require anything if \<^term>\a\ is incomparable with \<^term>\0\.\ class partial_abs_if = minus + uminus + ord + zero + abs + assumes abs_neg: "a \ 0 \ abs a = -a" assumes abs_pos: "a \ 0 \ abs a = a" class ordered_semiring_1 = ordered_semiring + semiring_1 \ \missing class analogous to \<^class>\linordered_semiring_1\ without requiring a total order\ begin lemma convex_bound_le: assumes "x \ a" and "y \ a" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end subclass (in linordered_semiring_1) ordered_semiring_1 .. class ordered_semiring_strict = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add + \ \missing class analogous to \<^class>\linordered_semiring_strict\ without requiring a total order\ assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass ordered_semiring proof fix a b c :: 'a assume t1: "a \ b" and t2: "0 \ c" thus "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from t2 show "a * c \ b * c" unfolding le_less - by (metis local.antisym_conv2 local.mult_not_zero local.mult_strict_right_mono t1) + by (metis local.antisym_conv2 local.mult_not_zero local.mult_strict_right_mono t1) qed lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes t1: "a < b" and t2: "c < d" and t3: "0 < b" and t4: "0 \ c" shows "a * c < b * d" proof- have "a * c < b * d" - by (metis local.dual_order.order_iff_strict local.dual_order.strict_trans2 - local.mult_strict_left_mono local.mult_strict_right_mono local.mult_zero_right t1 t2 t3 t4) + by (metis local.dual_order.order_iff_strict local.dual_order.strict_trans2 + local.mult_strict_left_mono local.mult_strict_right_mono local.mult_zero_right t1 t2 t3 t4) thus ?thesis using assms by blast qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes t1: "a < b" and t2: "c \ d" and t3: "0 \ a" and t4: "0 < c" shows "a * c < b * d" - using local.mult_strict_mono' local.mult_strict_right_mono local.order.order_iff_strict + using local.mult_strict_mono' local.mult_strict_right_mono local.order.order_iff_strict t1 t2 t3 t4 by auto lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" - by (metis assms(1) assms(2) assms(3) assms(4) local.antisym_conv2 local.dual_order.strict_trans1 + by (metis assms(1) assms(2) assms(3) assms(4) local.antisym_conv2 local.dual_order.strict_trans1 local.mult_strict_left_mono local.mult_strict_mono) end -subclass (in linordered_semiring_strict) ordered_semiring_strict +subclass (in linordered_semiring_strict) ordered_semiring_strict proof show "c * a < c * b" if "a < b" and "0 < c" for a :: 'a - and b - and c + and b + and c using that - by (simp add: local.mult_strict_left_mono) + by (simp add: local.mult_strict_left_mono) show "a * c < b * c" if "a < b" and "0 < c" for a :: 'a - and b - and c + and b + and c using that - by (simp add: local.mult_strict_right_mono) + by (simp add: local.mult_strict_right_mono) qed class ordered_semiring_1_strict = ordered_semiring_strict + semiring_1 - \ \missing class analogous to \<^class>\linordered_semiring_1_strict\ without requiring + \ \missing class analogous to \<^class>\linordered_semiring_1_strict\ without requiring a total order\ begin subclass ordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" and "y < a" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end -subclass (in linordered_semiring_1_strict) ordered_semiring_1_strict .. +subclass (in linordered_semiring_1_strict) ordered_semiring_1_strict .. class ordered_comm_semiring_strict = comm_semiring_0 + ordered_cancel_ab_semigroup_add + \ \missing class analogous to \<^class>\linordered_comm_semiring_strict\ without requiring a total order\ assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass ordered_semiring_strict proof fix a b c :: 'a assume "a < b" and "0 < c" thus "c * a < c * b" by (rule comm_mult_strict_left_mono) thus "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" and "0 \ c" thus "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end -subclass (in linordered_comm_semiring_strict) ordered_comm_semiring_strict +subclass (in linordered_comm_semiring_strict) ordered_comm_semiring_strict apply standard by (simp add: local.mult_strict_left_mono) class ordered_ring_strict = ring + ordered_semiring_strict + ordered_ab_group_add + partial_abs_if \ \missing class analogous to \<^class>\linordered_ring_strict\ without requiring a total order\ begin subclass ordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg subsection \Ordered fields\ -class ordered_field = field + order + ordered_comm_semiring_strict + ordered_ab_group_add - + partial_abs_if +class ordered_field = field + order + ordered_comm_semiring_strict + ordered_ab_group_add + + partial_abs_if \ \missing class analogous to \<^class>\linordered_field\ without requiring a total order\ begin lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) lemmas sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff lemmas (in -) sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff text\Simplify expressions equated with 1\ lemma zero_eq_1_divide_iff [simp]: "0 = 1 / a \ a = 0" by (cases "a = 0") (auto simp: field_simps) lemma one_divide_eq_0_iff [simp]: "1 / a = 0 \ a = 0" using zero_eq_1_divide_iff[of a] by simp text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ text\Simplify quotients that are compared with the value 1.\ text \Conditional Simplification Rules: No Case Splits\ lemma eq_divide_eq_1 [simp]: "(1 = b/a) = ((a \ 0 & a = b))" by (auto simp add: eq_divide_eq) lemma divide_eq_eq_1 [simp]: "(b/a = 1) = ((a \ 0 & a = b))" by (auto simp add: divide_eq_eq) end (* class ordered_field *) -text \The following type class intends to capture some important properties +text \The following type class intends to capture some important properties that are common both to the real and the complex numbers. The purpose is - to be able to state and prove lemmas that apply both to the real and the complex + to be able to state and prove lemmas that apply both to the real and the complex numbers without needing to state the lemma twice. \ class nice_ordered_field = ordered_field + zero_less_one + idom_abs_sgn + assumes positive_imp_inverse_positive: "0 < a \ 0 < inverse a" and inverse_le_imp_le: "inverse a \ inverse b \ 0 < a \ b \ a" and dense_le: "(\x. x < y \ x \ z) \ y \ z" and nn_comparable: "0 \ a \ 0 \ b \ a \ b \ b \ a" and abs_nn: "\x\ \ 0" begin subclass (in linordered_field) nice_ordered_field proof show "\a\ = - a" if "a \ 0" for a :: 'a using that - by simp + by simp show "\a\ = a" if "0 \ a" for a :: 'a using that - by simp + by simp show "0 < inverse a" if "0 < a" for a :: 'a using that - by simp + by simp show "b \ a" if "inverse a \ inverse b" and "0 < a" for a :: 'a and b using that - using local.inverse_le_imp_le by blast + using local.inverse_le_imp_le by blast show "y \ z" if "\x::'a. x < y \ x \ z" for y and z using that - using local.dense_le by blast + using local.dense_le by blast show "a \ b \ b \ a" if "0 \ a" and "0 \ b" for a :: 'a and b using that - by auto + by auto show "0 \ \x\" for x :: 'a - by simp + by simp qed lemma comparable: assumes h1: "a \ c \ a \ c" and h2: "b \ c \ b \ c" shows "a \ b \ b \ a" proof- have "a \ b" if t1: "\ b \ a" and t2: "a \ c" and t3: "b \ c" proof- have "0 \ c-a" - by (simp add: t2) + by (simp add: t2) moreover have "0 \ c-b" - by (simp add: t3) + by (simp add: t3) ultimately have "c-a \ c-b \ c-a \ c-b" by (rule nn_comparable) hence "-a \ -b \ -a \ -b" using local.add_le_imp_le_right local.uminus_add_conv_diff by presburger thus ?thesis by (simp add: t1) qed moreover have "a \ b" if t1: "\ b \ a" and t2: "c \ a" and t3: "b \ c" proof- - have "b \ a" - using local.dual_order.trans t2 t3 by blast + have "b \ a" + using local.dual_order.trans t2 t3 by blast thus ?thesis using t1 by auto qed moreover have "a \ b" if t1: "\ b \ a" and t2: "c \ a" and t3: "c \ b" proof- have "0 \ a-c" - by (simp add: t2) + by (simp add: t2) moreover have "0 \ b-c" - by (simp add: t3) + by (simp add: t3) ultimately have "a-c \ b-c \ a-c \ b-c" by (rule nn_comparable) hence "a \ b \ a \ b" by (simp add: local.le_diff_eq) thus ?thesis by (simp add: t1) qed ultimately show ?thesis using assms by auto qed lemma negative_imp_inverse_negative: "a < 0 \ inverse a < 0" by (insert positive_imp_inverse_positive [of "-a"], simp add: nonzero_inverse_minus_eq less_imp_not_eq) lemma inverse_positive_imp_positive: assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" shows "0 < a" proof - have "0 < inverse (inverse a)" using inv_gt_0 by (rule positive_imp_inverse_positive) thus "0 < a" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_negative_imp_negative: assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" shows "a < 0" proof- have "inverse (inverse a) < 0" using inv_less_0 by (rule negative_imp_inverse_negative) thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma linordered_field_no_lb: "\x. \y. y < x" proof fix x::'a have m1: "- (1::'a) < 0" by simp from add_strict_right_mono[OF m1, where c=x] have "(- 1) + x < x" by simp thus "\y. y < x" by blast qed lemma linordered_field_no_ub: "\x. \y. y > x" proof fix x::'a have m1: " (1::'a) > 0" by simp from add_strict_right_mono[OF m1, where c=x] have "1 + x > x" by simp thus "\y. y > x" by blast qed lemma less_imp_inverse_less: assumes less: "a < b" and apos: "0 < a" shows "inverse b < inverse a" - using assms by (metis local.dual_order.strict_iff_order + using assms by (metis local.dual_order.strict_iff_order local.inverse_inverse_eq local.inverse_le_imp_le local.positive_imp_inverse_positive) lemma inverse_less_imp_less: "inverse a < inverse b \ 0 < a \ b < a" using local.inverse_le_imp_le local.order.strict_iff_order by blast text\Both premises are essential. Consider -1 and 1.\ lemma inverse_less_iff_less [simp]: "0 < a \ 0 < b \ inverse a < inverse b \ b < a" by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) lemma le_imp_inverse_le: "a \ b \ 0 < a \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less) lemma inverse_le_iff_le [simp]: "0 < a \ 0 < b \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) text\These results refer to both operands being negative. The opposite-sign case is trivial, since inverse preserves signs.\ lemma inverse_le_imp_le_neg: "inverse a \ inverse b \ b < 0 \ b \ a" - by (metis local.inverse_le_imp_le local.inverse_minus_eq local.neg_0_less_iff_less + by (metis local.inverse_le_imp_le local.inverse_minus_eq local.neg_0_less_iff_less local.neg_le_iff_le) lemma inverse_less_imp_less_neg: "inverse a < inverse b \ b < 0 \ b < a" using local.dual_order.strict_iff_order local.inverse_le_imp_le_neg by blast lemma inverse_less_iff_less_neg [simp]: "a < 0 \ b < 0 \ inverse a < inverse b \ b < a" - by (metis local.antisym_conv2 local.inverse_less_imp_less_neg local.negative_imp_inverse_negative + by (metis local.antisym_conv2 local.inverse_less_imp_less_neg local.negative_imp_inverse_negative local.nonzero_inverse_inverse_eq local.order.strict_implies_order) lemma le_imp_inverse_le_neg: "a \ b \ b < 0 \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less_neg) lemma inverse_le_iff_le_neg [simp]: "a < 0 \ b < 0 \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) lemma one_less_inverse: "0 < a \ a < 1 \ 1 < inverse a" using less_imp_inverse_less [of a 1, unfolded inverse_1] . lemma one_le_inverse: "0 < a \ a \ 1 \ 1 \ inverse a" using le_imp_inverse_le [of a 1, unfolded inverse_1] . lemma pos_le_divide_eq [field_simps]: assumes "0 < c" shows "a \ b / c \ a * c \ b" - using assms by (metis local.divide_eq_imp local.divide_inverse_commute - local.dual_order.order_iff_strict local.dual_order.strict_iff_order - local.mult_right_mono local.mult_strict_left_mono local.nonzero_divide_eq_eq + using assms by (metis local.divide_eq_imp local.divide_inverse_commute + local.dual_order.order_iff_strict local.dual_order.strict_iff_order + local.mult_right_mono local.mult_strict_left_mono local.nonzero_divide_eq_eq local.order.strict_implies_order local.positive_imp_inverse_positive) lemma pos_less_divide_eq [field_simps]: assumes "0 < c" shows "a < b / c \ a * c < b" - using assms local.dual_order.strict_iff_order local.nonzero_divide_eq_eq local.pos_le_divide_eq + using assms local.dual_order.strict_iff_order local.nonzero_divide_eq_eq local.pos_le_divide_eq by auto lemma neg_less_divide_eq [field_simps]: assumes "c < 0" shows "a < b / c \ b < a * c" - by (metis assms local.minus_divide_divide local.mult_minus_right local.neg_0_less_iff_less + by (metis assms local.minus_divide_divide local.mult_minus_right local.neg_0_less_iff_less local.neg_less_iff_less local.pos_less_divide_eq) lemma neg_le_divide_eq [field_simps]: assumes "c < 0" shows "a \ b / c \ b \ a * c" - by (metis assms local.dual_order.order_iff_strict local.dual_order.strict_iff_order + by (metis assms local.dual_order.order_iff_strict local.dual_order.strict_iff_order local.neg_less_divide_eq local.nonzero_divide_eq_eq) lemma pos_divide_le_eq [field_simps]: assumes "0 < c" shows "b / c \ a \ b \ a * c" - by (metis assms local.dual_order.strict_iff_order local.nonzero_eq_divide_eq + by (metis assms local.dual_order.strict_iff_order local.nonzero_eq_divide_eq local.pos_le_divide_eq) lemma pos_divide_less_eq [field_simps]: assumes "0 < c" shows "b / c < a \ b < a * c" - by (metis assms local.minus_divide_left local.mult_minus_left local.neg_less_iff_less + by (metis assms local.minus_divide_left local.mult_minus_left local.neg_less_iff_less local.pos_less_divide_eq) lemma neg_divide_le_eq [field_simps]: assumes "c < 0" shows "b / c \ a \ a * c \ b" - by (metis assms local.minus_divide_left local.mult_minus_left local.neg_le_divide_eq + by (metis assms local.minus_divide_left local.mult_minus_left local.neg_le_divide_eq local.neg_le_iff_le) lemma neg_divide_less_eq [field_simps]: assumes "c < 0" shows "b / c < a \ a * c < b" using assms local.dual_order.strict_iff_order local.neg_divide_le_eq by auto text\The following \field_simps\ rules are necessary, as minus is always moved atop of division but we want to get rid of division.\ lemma pos_le_minus_divide_eq [field_simps]: "0 < c \ a \ - (b / c) \ a * c \ - b" unfolding minus_divide_left by (rule pos_le_divide_eq) lemma neg_le_minus_divide_eq [field_simps]: "c < 0 \ a \ - (b / c) \ - b \ a * c" unfolding minus_divide_left by (rule neg_le_divide_eq) lemma pos_less_minus_divide_eq [field_simps]: "0 < c \ a < - (b / c) \ a * c < - b" unfolding minus_divide_left by (rule pos_less_divide_eq) lemma neg_less_minus_divide_eq [field_simps]: "c < 0 \ a < - (b / c) \ - b < a * c" unfolding minus_divide_left by (rule neg_less_divide_eq) lemma pos_minus_divide_less_eq [field_simps]: "0 < c \ - (b / c) < a \ - b < a * c" unfolding minus_divide_left by (rule pos_divide_less_eq) lemma neg_minus_divide_less_eq [field_simps]: "c < 0 \ - (b / c) < a \ a * c < - b" unfolding minus_divide_left by (rule neg_divide_less_eq) lemma pos_minus_divide_le_eq [field_simps]: "0 < c \ - (b / c) \ a \ - b \ a * c" unfolding minus_divide_left by (rule pos_divide_le_eq) lemma neg_minus_divide_le_eq [field_simps]: "c < 0 \ - (b / c) \ a \ a * c \ - b" unfolding minus_divide_left by (rule neg_divide_le_eq) lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) text\Lemmas \sign_simps\ is a first attempt to automate proofs -of positivity/negativity needed for \field_simps\. Have not added \sign_simps\ to \field_simps\ +of positivity/negativity needed for \field_simps\. Have not added \sign_simps\ to \field_simps\ because the former can lead to case explosions.\ lemma divide_pos_pos[simp]: "0 < x \ 0 < y \ 0 < x / y" by(simp add:field_simps) lemma divide_nonneg_pos: "0 \ x \ 0 < y \ 0 \ x / y" by(simp add:field_simps) lemma divide_neg_pos: "x < 0 \ 0 < y \ x / y < 0" by(simp add:field_simps) lemma divide_nonpos_pos: "x \ 0 \ 0 < y \ x / y \ 0" by(simp add:field_simps) lemma divide_pos_neg: "0 < x \ y < 0 \ x / y < 0" by(simp add:field_simps) lemma divide_nonneg_neg: "0 \ x \ y < 0 \ x / y \ 0" by(simp add:field_simps) lemma divide_neg_neg: "x < 0 \ y < 0 \ 0 < x / y" by(simp add:field_simps) lemma divide_nonpos_neg: "x \ 0 \ y < 0 \ 0 \ x / y" by(simp add:field_simps) lemma divide_strict_right_mono: "a < b \ 0 < c \ a / c < b / c" by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono positive_imp_inverse_positive) lemma divide_strict_right_mono_neg: "b < a \ c < 0 \ a / c < b / c" by (simp add: local.neg_less_divide_eq) text\The last premise ensures that \<^term>\a\ and \<^term>\b\ have the same sign\ lemma divide_strict_left_mono: "b < a \ 0 < c \ 0 < a*b \ c / a < c / b" by (metis local.divide_neg_pos local.dual_order.strict_iff_order local.frac_less_eq local.less_iff_diff_less_0 local.mult_not_zero local.mult_strict_left_mono) lemma divide_left_mono: "b \ a \ 0 \ c \ 0 < a*b \ c / a \ c / b" using local.divide_cancel_left local.divide_strict_left_mono local.dual_order.order_iff_strict by auto lemma divide_strict_left_mono_neg: "a < b \ c < 0 \ 0 < a*b \ c / a < c / b" by (metis local.divide_strict_left_mono local.minus_divide_left local.neg_0_less_iff_less local.neg_less_iff_less mult_commute) lemma mult_imp_div_pos_le: "0 < y \ x \ z * y \ x / y \ z" by (subst pos_divide_le_eq, assumption+) lemma mult_imp_le_div_pos: "0 < y \ z * y \ x \ z \ x / y" by(simp add:field_simps) lemma mult_imp_div_pos_less: "0 < y \ x < z * y \ x / y < z" by(simp add:field_simps) lemma mult_imp_less_div_pos: "0 < y \ z * y < x \ z < x / y" by(simp add:field_simps) lemma frac_le: "0 \ x \ x \ y \ 0 < w \ w \ z \ x / z \ y / w" using local.mult_imp_div_pos_le local.mult_imp_le_div_pos local.mult_mono by auto lemma frac_less: "0 \ x \ x < y \ 0 < w \ w \ z \ x / z < y / w" proof- assume a1: "w \ z" assume a2: "0 < w" assume a3: "0 \ x" assume a4: "x < y" have f5: "a = 0 \ (b = c / a) = (b * a = c)" for a b c::'a by (meson local.nonzero_eq_divide_eq) have f6: "0 < z" - using a2 a1 by (meson local.order.ordering_axioms ordering.strict_trans2) + using a2 a1 by (meson local.order.ordering_axioms order.strict_trans2) have "z \ 0" using a2 a1 by (meson local.leD) moreover have "x / z \ y / w" using a1 a2 a3 a4 local.frac_eq_eq local.mult_less_le_imp_less by fastforce ultimately have "x / z \ y / w" using f5 by (metis (no_types)) thus ?thesis - using a4 a3 a2 a1 by (meson local.frac_le local.order.not_eq_order_implies_strict + using a4 a3 a2 a1 by (meson local.frac_le local.order.not_eq_order_implies_strict local.order.strict_implies_order) qed lemma frac_less2: "0 < x \ x \ y \ 0 < w \ w < z \ x / z < y / w" - by (metis local.antisym_conv2 local.divide_cancel_left local.dual_order.strict_implies_order + by (metis local.antisym_conv2 local.divide_cancel_left local.dual_order.strict_implies_order local.frac_le local.frac_less) lemma less_half_sum: "a < b \ a < (a+b) / (1+1)" by (metis local.add_pos_pos local.add_strict_left_mono local.mult_imp_less_div_pos local.semiring_normalization_rules(4) local.zero_less_one mult_commute) lemma gt_half_sum: "a < b \ (a+b)/(1+1) < b" by (metis local.add_pos_pos local.add_strict_left_mono local.mult_imp_div_pos_less local.semiring_normalization_rules(24) local.semiring_normalization_rules(4) local.zero_less_one mult_commute) subclass unbounded_dense_order proof fix x y :: 'a have less_add_one: "a < a + 1" for a::'a by auto from less_add_one show "\y. x < y" - by blast + by blast from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono) hence "x - 1 < x + 1 - 1" by simp hence "x - 1 < x" by (simp add: algebra_simps) thus "\y. y < x" .. show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) qed lemma dense_le_bounded: fixes x y z :: 'a assumes "x < y" and *: "\w. \ x < w ; w < y \ \ w \ z" shows "y \ z" proof (rule dense_le) fix w assume "w < y" from dense[OF \x < y\] obtain u where "x < u" "u < y" by safe have "u \ w \ w \ u" using \u < y\ \w < y\ comparable local.order.strict_implies_order by blast thus "w \ z" using "*" \u < y\ \w < y\ \x < u\ local.dual_order.trans local.order.strict_trans2 by blast qed subclass field_abs_sgn .. lemma nonzero_abs_inverse: "a \ 0 \ \inverse a\ = inverse \a\" by (rule abs_inverse) lemma nonzero_abs_divide: "b \ 0 \ \a / b\ = \a\ / \b\" by (rule abs_divide) lemma field_le_epsilon: assumes e: "\e. 0 < e \ x \ y + e" shows "x \ y" proof (rule dense_le) fix t assume "t < x" hence "0 < x - t" by (simp add: less_diff_eq) from e [OF this] have "x + 0 \ x + (y - t)" by (simp add: algebra_simps) hence "0 \ y - t" by (simp only: add_le_cancel_left) thus "t \ y" by (simp add: algebra_simps) qed lemma inverse_positive_iff_positive [simp]: "(0 < inverse a) = (0 < a)" using local.positive_imp_inverse_positive by fastforce lemma inverse_negative_iff_negative [simp]: "(inverse a < 0) = (a < 0)" using local.negative_imp_inverse_negative by fastforce lemma inverse_nonnegative_iff_nonnegative [simp]: "0 \ inverse a \ 0 \ a" by (simp add: local.dual_order.order_iff_strict) lemma inverse_nonpositive_iff_nonpositive [simp]: "inverse a \ 0 \ a \ 0" using local.inverse_nonnegative_iff_nonnegative local.neg_0_le_iff_le by fastforce lemma one_less_inverse_iff: "1 < inverse x \ 0 < x \ x < 1" using less_trans[of 1 x 0 for x] by (metis local.dual_order.strict_trans local.inverse_1 local.inverse_less_imp_less local.inverse_positive_iff_positive local.one_less_inverse local.zero_less_one) lemma one_le_inverse_iff: "1 \ inverse x \ 0 < x \ x \ 1" by (metis local.dual_order.strict_trans1 local.inverse_1 local.inverse_le_imp_le local.inverse_positive_iff_positive local.one_le_inverse local.zero_less_one) lemma inverse_less_1_iff: "inverse x < 1 \ x \ 0 \ 1 < x" proof (rule) assume invx1: "inverse x < 1" have "inverse x \ 0 \ inverse x \ 0" using comparable invx1 local.order.strict_implies_order local.zero_less_one by blast then consider (leq0) "inverse x \ 0" | (pos) "inverse x > 0" | (zero) "inverse x = 0" using local.antisym_conv1 by blast thus "x \ 0 \ 1 < x" - by (metis invx1 local.eq_iff local.inverse_1 local.inverse_less_imp_less - local.inverse_nonpositive_iff_nonpositive local.inverse_positive_imp_positive) + by (metis invx1 local.eq_refl local.inverse_1 inverse_less_imp_less + inverse_nonpositive_iff_nonpositive inverse_positive_iff_positive) next assume "x \ 0 \ 1 < x" then consider (neg) "x \ 0" | (g1) "1 < x" by auto thus "inverse x < 1" by (metis local.dual_order.not_eq_order_implies_strict local.dual_order.strict_trans - local.inverse_1 local.inverse_negative_iff_negative local.inverse_zero - local.less_imp_inverse_less local.zero_less_one) + local.inverse_1 local.inverse_negative_iff_negative local.inverse_zero + local.less_imp_inverse_less local.zero_less_one) qed lemma inverse_le_1_iff: "inverse x \ 1 \ x \ 0 \ 1 \ x" - by (metis local.dual_order.order_iff_strict local.inverse_1 local.inverse_le_iff_le + by (metis local.dual_order.order_iff_strict local.inverse_1 local.inverse_le_iff_le local.inverse_less_1_iff local.one_le_inverse_iff) text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ lemma zero_le_divide_1_iff [simp]: "0 \ 1 / a \ 0 \ a" - using local.dual_order.order_iff_strict local.inverse_eq_divide + using local.dual_order.order_iff_strict local.inverse_eq_divide local.inverse_positive_iff_positive by auto lemma zero_less_divide_1_iff [simp]: "0 < 1 / a \ 0 < a" by (simp add: local.dual_order.strict_iff_order) lemma divide_le_0_1_iff [simp]: "1 / a \ 0 \ a \ 0" - by (smt local.abs_0 local.abs_1 local.abs_divide local.abs_neg local.abs_nn + by (smt local.abs_0 local.abs_1 local.abs_divide local.abs_neg local.abs_nn local.divide_cancel_left local.le_minus_iff local.minus_divide_right local.zero_neq_one) lemma divide_less_0_1_iff [simp]: "1 / a < 0 \ a < 0" using local.dual_order.strict_iff_order by auto lemma divide_right_mono: "a \ b \ 0 \ c \ a/c \ b/c" using local.divide_cancel_right local.divide_strict_right_mono local.dual_order.order_iff_strict by blast lemma divide_right_mono_neg: "a \ b \ c \ 0 \ b / c \ a / c" by (metis local.divide_cancel_right local.divide_strict_right_mono_neg local.dual_order.strict_implies_order local.eq_refl local.le_imp_less_or_eq) lemma divide_left_mono_neg: "a \ b - \ c \ 0 \ 0 < a * b \ c / a \ c / b" + \ c \ 0 \ 0 < a * b \ c / a \ c / b" by (metis local.divide_left_mono local.minus_divide_left local.neg_0_le_iff_le local.neg_le_iff_le mult_commute) lemma divide_nonneg_nonneg [simp]: "0 \ x \ 0 \ y \ 0 \ x / y" using local.divide_eq_0_iff local.divide_nonneg_pos local.dual_order.order_iff_strict by blast lemma divide_nonpos_nonpos: "x \ 0 \ y \ 0 \ 0 \ x / y" using local.divide_nonpos_neg local.dual_order.order_iff_strict by auto lemma divide_nonneg_nonpos: "0 \ x \ y \ 0 \ x / y \ 0" by (metis local.divide_eq_0_iff local.divide_nonneg_neg local.dual_order.order_iff_strict) lemma divide_nonpos_nonneg: "x \ 0 \ 0 \ y \ x / y \ 0" using local.divide_nonpos_pos local.dual_order.order_iff_strict by auto text \Conditional Simplification Rules: No Case Splits\ lemma le_divide_eq_1_pos [simp]: "0 < a \ (1 \ b/a) = (a \ b)" by (simp add: local.pos_le_divide_eq) lemma le_divide_eq_1_neg [simp]: "a < 0 \ (1 \ b/a) = (b \ a)" by (metis local.le_divide_eq_1_pos local.minus_divide_divide local.neg_0_less_iff_less local.neg_le_iff_le) lemma divide_le_eq_1_pos [simp]: "0 < a \ (b/a \ 1) = (b \ a)" using local.pos_divide_le_eq by auto lemma divide_le_eq_1_neg [simp]: "a < 0 \ (b/a \ 1) = (a \ b)" - by (metis local.divide_le_eq_1_pos local.minus_divide_divide local.neg_0_less_iff_less + by (metis local.divide_le_eq_1_pos local.minus_divide_divide local.neg_0_less_iff_less local.neg_le_iff_le) lemma less_divide_eq_1_pos [simp]: "0 < a \ (1 < b/a) = (a < b)" by (simp add: local.dual_order.strict_iff_order) lemma less_divide_eq_1_neg [simp]: "a < 0 \ (1 < b/a) = (b < a)" using local.dual_order.strict_iff_order by auto lemma divide_less_eq_1_pos [simp]: "0 < a \ (b/a < 1) = (b < a)" using local.divide_le_eq_1_pos local.dual_order.strict_iff_order by auto lemma divide_less_eq_1_neg [simp]: "a < 0 \ b/a < 1 \ a < b" using local.dual_order.strict_iff_order by auto lemma abs_div_pos: "0 < y \ \x\ / y = \x / y\" by (simp add: local.abs_pos) lemma zero_le_divide_abs_iff [simp]: "(0 \ a / \b\) = (0 \ a | b = 0)" -proof +proof assume assm: "0 \ a / \b\" have absb: "abs b \ 0" by (fact abs_nn) thus "0 \ a \ b = 0" using absb assm local.abs_eq_0_iff local.mult_nonneg_nonneg by fastforce next assume "0 \ a \ b = 0" then consider (a) "0 \ a" | (b) "b = 0" by atomize_elim auto thus "0 \ a / \b\" by (metis local.abs_eq_0_iff local.abs_nn local.divide_eq_0_iff local.divide_nonneg_nonneg) qed lemma divide_le_0_abs_iff [simp]: "(a / \b\ \ 0) = (a \ 0 | b = 0)" by (metis local.minus_divide_left local.neg_0_le_iff_le local.zero_le_divide_abs_iff) text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: assumes "u \ v" and "0 \ r" and "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms - by (metis local.divide_le_eq_1_pos local.division_ring_divide_zero + by (metis local.divide_le_eq_1_pos local.division_ring_divide_zero local.dual_order.order_iff_strict local.dual_order.trans local.zero_less_one) hence "(r/s) * (v - u) \ 1 * (v - u)" using assms(1) local.diff_ge_0_iff_ge local.mult_right_mono by blast thus ?thesis by (simp add: field_simps) qed end (* class nice_ordered_field *) code_identifier code_module Ordered_Fields \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection\Ordered Complex\ declare Conjugate.less_eq_complex_def[simp del] declare Conjugate.less_complex_def[simp del] subsection \Ordering on complex numbers\ instantiation complex :: nice_ordered_field begin instance proof intro_classes note defs = less_eq_complex_def less_complex_def abs_complex_def fix x y z a b c :: complex show "a \ 0 \ \a\ = - a" unfolding defs by (simp add: cmod_eq_Re complex_is_Real_iff) show "0 \ a \ \a\ = a" unfolding defs by (metis abs_of_nonneg cmod_eq_Re comp_apply complex.exhaust_sel complex_of_real_def zero_complex.simps(1) zero_complex.simps(2)) show "a < b \ 0 < c \ c * a < c * b" unfolding defs by auto show "0 < (1::complex)" unfolding defs by simp show "0 < a \ 0 < inverse a" unfolding defs by auto define ra ia rb ib rc ic where "ra = Re a" "ia = Im a" "rb = Re b" "ib = Im b" "rc = Re c" "ic = Im c" note ri = this[symmetric] hence "a = Complex ra ia" "b = Complex rb ib" "c = Complex rc ic" by auto note ri = this ri have "rb \ ra" - if "1 / ra \ (if rb = 0 then 0 else 1 / rb)" + if "1 / ra \ (if rb = 0 then 0 else 1 / rb)" and "ia = 0" and "0 < ra" and "ib = 0" proof(cases "rb = 0") case True thus ?thesis - using that(3) by auto + using that(3) by auto next case False thus ?thesis - by (smt nice_ordered_field_class.frac_less2 that(1) that(3)) + by (smt nice_ordered_field_class.frac_less2 that(1) that(3)) qed thus "inverse a \ inverse b \ 0 < a \ b \ a" unfolding defs ri - by (auto simp: power2_eq_square) + by (auto simp: power2_eq_square) show "(\a. a < b \ a \ c) \ b \ c" unfolding defs ri - by (metis complex.sel(1) complex.sel(2) dense less_le_not_le - nice_ordered_field_class.linordered_field_no_lb not_le_imp_less) + by (metis complex.sel(1) complex.sel(2) dense less_le_not_le + nice_ordered_field_class.linordered_field_no_lb not_le_imp_less) show "0 \ a \ 0 \ b \ a \ b \ b \ a" unfolding defs by auto show "0 \ \x\" unfolding defs by auto qed end -lemma less_eq_complexI: "Re x \ Re y \ Im x = Im y \ x\y" unfolding less_eq_complex_def +lemma less_eq_complexI: "Re x \ Re y \ Im x = Im y \ x\y" unfolding less_eq_complex_def by simp -lemma less_complexI: "Re x < Re y \ Im x = Im y \ x Im x = Im y \ x y \ complex_of_real x \ complex_of_real y" unfolding less_eq_complex_def by auto lemma complex_of_real_mono_iff[simp]: "complex_of_real x \ complex_of_real y \ x \ y" unfolding less_eq_complex_def by auto lemma complex_of_real_strict_mono_iff[simp]: "complex_of_real x < complex_of_real y \ x < y" unfolding less_complex_def by auto lemma complex_of_real_nn_iff[simp]: "0 \ complex_of_real y \ 0 \ y" unfolding less_eq_complex_def by auto lemma complex_of_real_pos_iff[simp]: "0 < complex_of_real y \ 0 < y" unfolding less_complex_def by auto lemma Re_mono: "x \ y \ Re x \ Re y" unfolding less_eq_complex_def by simp lemma comp_Im_same: "x \ y \ Im x = Im y" unfolding less_eq_complex_def by simp lemma Re_strict_mono: "x < y \ Re x < Re y" unfolding less_complex_def by simp lemma complex_of_real_cmod: assumes "x \ 0" shows "complex_of_real (cmod x) = x" by (metis Reals_cases abs_of_nonneg assms comp_Im_same complex_is_Real_iff complex_of_real_nn_iff norm_of_real zero_complex.simps(2)) end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Pretty_Code_Examples.thy b/thys/Complex_Bounded_Operators/extra/Extra_Pretty_Code_Examples.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Pretty_Code_Examples.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Pretty_Code_Examples.thy @@ -1,73 +1,73 @@ section \\Extra_Pretty_Code_Examples\ -- Setup for nicer output of \value\\ theory Extra_Pretty_Code_Examples imports - "HOL-ex.Sqrt" + "HOL-Examples.Sqrt" Real_Impl.Real_Impl "HOL-Library.Code_Target_Numeral" Jordan_Normal_Form.Matrix_Impl begin text \Some setup that makes the output of the \value\ command more readable if matrices and complex numbers are involved. It is not recommended to import this theory in theories that get included in actual developments (because of the changes to the code generation setup). It is meant for inclusion in example theories only.\ lemma two_sqrt_irrat[simp]: "2 \ sqrt_irrat" using sqrt_prime_irrational[OF two_is_prime_nat] unfolding Rats_def sqrt_irrat_def image_def apply auto proof - (* Sledgehammer proof *) fix p :: rat assume "p * p = 2" hence f1: "(Ratreal p)\<^sup>2 = real 2" by (metis Ratreal_def of_nat_numeral of_rat_numeral_eq power2_eq_square real_times_code) have "\r. if 0 \ r then sqrt (r\<^sup>2) = r else r + sqrt (r\<^sup>2) = 0" by simp hence f2: "Ratreal p + sqrt ((Ratreal p)\<^sup>2) = 0" using f1 by (metis Ratreal_def Rats_def \sqrt (real 2) \ \\ range_eqI) have f3: "sqrt (real 2) + - 1 * sqrt ((Ratreal p)\<^sup>2) \ 0" using f1 by fastforce have f4: "0 \ sqrt (real 2) + - 1 * sqrt ((Ratreal p)\<^sup>2)" using f1 by force have f5: "(- 1 * sqrt (real 2) = real_of_rat p) = (sqrt (real 2) + real_of_rat p = 0)" by linarith have "\x0. - (x0::real) = - 1 * x0" by auto hence "sqrt (real 2) + real_of_rat p \ 0" using f5 by (metis (no_types) Rats_def Rats_minus_iff \sqrt (real 2) \ \\ range_eqI) thus False using f4 f3 f2 by simp qed (* Ensure that complex numbers with zero-imaginary part are rendered as reals *) lemma complex_number_code_post[code_post]: shows "Complex a 0 = complex_of_real a" and "complex_of_real 0 = 0" and "complex_of_real 1 = 1" and "complex_of_real (a/b) = complex_of_real a / complex_of_real b" and "complex_of_real (numeral n) = numeral n" and "complex_of_real (-r) = - complex_of_real r" using complex_eq_cancel_iff2 by auto (* Make real number implementation more readable *) lemma real_number_code_post[code_post]: shows "real_of (Abs_mini_alg (p, 0, b)) = real_of_rat p" and "real_of (Abs_mini_alg (p, q, 2)) = real_of_rat p + real_of_rat q * sqrt 2" and "sqrt 0 = 0" and "sqrt (real 0) = 0" and "x * (0::real) = 0" and "(0::real) * x = 0" and "(0::real) + x = x" and "x + (0::real) = x" and "(1::real) * x = x" and "x * (1::real) = x" by (auto simp add: eq_onp_same_args real_of.abs_eq) (* Hide IArray in output *) translations "x" \ "CONST IArray x" end