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,4403 +1,4530 @@ section \\Complex_Bounded_Linear_Function\ -- Complex bounded linear functions (bounded operators)\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Bounded_Linear_Function imports "HOL-Types_To_Sets.Types_To_Sets" Banach_Steinhaus.Banach_Steinhaus Complex_Inner_Product One_Dimensional_Spaces Complex_Bounded_Linear_Function0 "HOL-Library.Function_Algebras" begin unbundle lattice_syntax subsection \Misc basic facts and declarations\ notation cblinfun_apply (infixr "*\<^sub>V" 70) lemma id_cblinfun_apply[simp]: "id_cblinfun *\<^sub>V \ = \" apply transfer by simp lemma apply_id_cblinfun[simp]: \(*\<^sub>V) id_cblinfun = id\ by auto lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" apply transfer by (simp add: clinear_continuous_at) declare cblinfun.scaleC_left[simp] lemma cblinfun_apply_clinear[simp]: \clinear (cblinfun_apply A)\ using bounded_clinear.axioms(1) cblinfun_apply by blast lemma cblinfun_cinner_eqI: fixes A B :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ assumes \\\. cinner \ (A *\<^sub>V \) = cinner \ (B *\<^sub>V \)\ shows \A = B\ proof - define C where \C = A - B\ have C0[simp]: \cinner \ (C \) = 0\ for \ by (simp add: C_def assms cblinfun.diff_left cinner_diff_right) { fix f g \ have \0 = cinner (f + \ *\<^sub>C g) (C *\<^sub>V (f + \ *\<^sub>C g))\ by (simp add: cinner_diff_right minus_cblinfun.rep_eq) also have \\ = \ *\<^sub>C cinner f (C g) + cnj \ *\<^sub>C cinner g (C f)\ by (smt (z3) C0 add.commute add.right_neutral cblinfun.add_right cblinfun.scaleC_right cblinfun_cinner_right.rep_eq cinner_add_left cinner_scaleC_left complex_scaleC_def) finally have \\ *\<^sub>C cinner f (C g) = - cnj \ *\<^sub>C cinner g (C f)\ by (simp add: eq_neg_iff_add_eq_0) } then have \cinner f (C g) = 0\ for f g by (metis complex_cnj_i complex_cnj_one complex_vector.scale_cancel_right complex_vector.scale_left_imp_eq equation_minus_iff i_squared mult_eq_0_iff one_neq_neg_one) then have \C g = 0\ for g using cinner_eq_zero_iff by blast then have \C = 0\ by (simp add: cblinfun_eqI) then show \A = B\ using C_def by auto qed lemma id_cblinfun_not_0[simp]: \(id_cblinfun :: 'a::{complex_normed_vector, not_singleton} \\<^sub>C\<^sub>L _) \ 0\ by (metis (full_types) Extra_General.UNIV_not_singleton cblinfun.zero_left cblinfun_id_cblinfun_apply ex_norm1 norm_zero one_neq_zero) lemma cblinfun_norm_geqI: assumes \norm (f *\<^sub>V x) / norm x \ K\ shows \norm f \ K\ using assms apply transfer by (smt (z3) bounded_clinear.bounded_linear le_onorm) (* This lemma is proven in Complex_Bounded_Linear_Function0 but we add the [simp] only here because we try to keep Complex_Bounded_Linear_Function0 as close to Bounded_Linear_Function as possible. *) declare scaleC_conv_of_complex[simp] lemma cblinfun_eq_0_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = 0" shows \F *\<^sub>V x = 0\ apply (rule complex_vector.linear_eq_0_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = G *\<^sub>V s" shows \F *\<^sub>V x = G *\<^sub>V x\ apply (rule complex_vector.linear_eq_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_0_on_UNIV_span: fixes basis::\'a::complex_normed_vector set\ assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = 0" shows \F = 0\ by (metis cblinfun_eq_0_on_span UNIV_I assms cblinfun.zero_left cblinfun_eqI) lemma cblinfun_eq_on_UNIV_span: fixes basis::"'a::complex_normed_vector set" and \::"'a \ 'b::complex_normed_vector" assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = G *\<^sub>V s" shows \F = G\ proof- have "F - G = 0" apply (rule cblinfun_eq_0_on_UNIV_span[where basis=basis]) using assms by (auto simp add: cblinfun.diff_left) thus ?thesis by simp qed lemma cblinfun_eq_on_canonical_basis: fixes f g::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = g *\<^sub>V u" shows "f = g" apply (rule cblinfun_eq_on_UNIV_span[where basis=basis]) using assms is_generator_set is_cindependent_set by auto lemma cblinfun_eq_0_on_canonical_basis: fixes f ::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = 0" shows "f = 0" by (simp add: assms cblinfun_eq_on_canonical_basis) lemma cinner_canonical_basis_eq_0: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = 0" shows "F = 0" proof- have "F *\<^sub>V u = 0" if "u\basisA" for u proof- have "\v. v\basisB \ \v, F *\<^sub>V u\ = 0" by (simp add: assms(3) that) moreover have "(\v. v\basisB \ \v, x\ = 0) \ x = 0" for x proof- assume r1: "\v. v\basisB \ \v, x\ = 0" have "\v, x\ = 0" for v proof- have "cspan basisB = UNIV" using basisB_def is_generator_set by auto hence "v \ cspan basisB" by (smt iso_tuple_UNIV_I) hence "\t s. v = (\a\t. s a *\<^sub>C a) \ finite t \ t \ basisB" using complex_vector.span_explicit by (smt mem_Collect_eq) then obtain t s where b1: "v = (\a\t. s a *\<^sub>C a)" and b2: "finite t" and b3: "t \ basisB" by blast have "\v, x\ = \(\a\t. s a *\<^sub>C a), x\" by (simp add: b1) also have "\ = (\a\t. \s a *\<^sub>C a, x\)" using cinner_sum_left by blast also have "\ = (\a\t. cnj (s a) * \a, x\)" by auto also have "\ = 0" using b3 r1 subsetD by force finally show ?thesis by simp qed thus ?thesis by (simp add: \\v. \v, x\ = 0\ cinner_extensionality) qed ultimately show ?thesis by simp qed thus ?thesis using basisA_def cblinfun_eq_0_on_canonical_basis by auto qed lemma cinner_canonical_basis_eq: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = \v, G *\<^sub>V u\" shows "F = G" proof- define H where "H = F - G" have "\u v. u\basisA \ v\basisB \ \v, H *\<^sub>V u\ = 0" unfolding H_def by (simp add: assms(3) cinner_diff_right minus_cblinfun.rep_eq) hence "H = 0" by (simp add: basisA_def basisB_def cinner_canonical_basis_eq_0) thus ?thesis unfolding H_def by simp qed lemma cinner_canonical_basis_eq': defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \F *\<^sub>V u, v\ = \G *\<^sub>V u, v\" shows "F = G" using cinner_canonical_basis_eq assms by (metis cinner_commute') lemma cblinfun_norm_approx_witness: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ > 0\ shows \\\. norm (A *\<^sub>V \) \ norm A - \ \ norm \ = 1\ proof (transfer fixing: \) fix A :: \'a \ 'b\ assume [simp]: \bounded_clinear A\ have \\y\{norm (A x) |x. norm x = 1}. y > \ {norm (A x) |x. norm x = 1} - \\ apply (rule Sup_real_close) using assms by (auto simp: ex_norm1 bounded_clinear.bounded_linear bdd_above_norm_f) also have \\ {norm (A x) |x. norm x = 1} = onorm A\ by (simp add: bounded_clinear.bounded_linear onorm_sphere) finally show \\\. onorm A - \ \ norm (A \) \ norm \ = 1\ by force qed lemma cblinfun_norm_approx_witness_mult: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ < 1\ shows \\\. norm (A *\<^sub>V \) \ norm A * \ \ norm \ = 1\ proof (cases \norm A = 0\) case True then show ?thesis apply auto by (simp add: ex_norm1) next case False then have \(1 - \) * norm A > 0\ using assms by fastforce then obtain \ where geq: \norm (A *\<^sub>V \) \ norm A - ((1 - \) * norm A)\ and \norm \ = 1\ using cblinfun_norm_approx_witness by blast have \norm A * \ = norm A - (1 - \) * norm A\ by (simp add: mult.commute right_diff_distrib') also have \\ \ norm (A *\<^sub>V \)\ by (rule geq) finally show ?thesis using \norm \ = 1\ by auto qed lemma cblinfun_to_CARD_1_0[simp]: \(A :: _ \\<^sub>C\<^sub>L _::CARD_1) = 0\ apply (rule cblinfun_eqI) by auto lemma cblinfun_from_CARD_1_0[simp]: \(A :: _::CARD_1 \\<^sub>C\<^sub>L _) = 0\ apply (rule cblinfun_eqI) apply (subst CARD_1_vec_0) by auto lemma cblinfun_cspan_UNIV: fixes basis :: \('a::{complex_normed_vector,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_normed_vector) set\ and basisA :: \'a set\ and basisB :: \'b set\ assumes \cspan basisA = UNIV\ and \cspan basisB = UNIV\ assumes basis: \\a b. a\basisA \ b\basisB \ \F\basis. \a'\basisA. F *\<^sub>V a' = (if a'=a then b else 0)\ shows \cspan basis = UNIV\ proof - obtain basisA' where \basisA' \ basisA\ and \cindependent basisA'\ and \cspan basisA' = UNIV\ by (metis assms(1) complex_vector.maximal_independent_subset complex_vector.span_eq top_greatest) then have [simp]: \finite basisA'\ by (simp add: cindependent_cfinite_dim_finite) have basis': \\a b. a\basisA' \ b\basisB \ \F\basis. \a'\basisA'. F *\<^sub>V a' = (if a'=a then b else 0)\ using basis \basisA' \ basisA\ by fastforce obtain F where F: \F a b \ basis \ F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using basis' by metis then have F_apply: \F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' using that by auto have F_basis: \F a b \ basis\ if \a\basisA'\ \b\basisB\ for a b using that F by auto have b_span: \\G\cspan {F a b|b. b\basisB}. \a'\basisA'. G *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ for a b proof - from \cspan basisB = UNIV\ obtain r t where \finite t\ and \t \ basisB\ and b_lincom: \b = (\a\t. r a *\<^sub>C a)\ unfolding complex_vector.span_alt apply atomize_elim by blast define G where \G = (\i\t. r i *\<^sub>C F a i)\ have \G \ cspan {F a b|b. b\basisB}\ using \finite t\ \t \ basisB\ unfolding G_def by (smt (verit, ccfv_threshold) complex_vector.span_base complex_vector.span_scale complex_vector.span_sum mem_Collect_eq subset_eq) moreover have \G *\<^sub>V a' = (if a'=a then b else 0)\ if \a'\basisA'\ for a' apply (cases \a'=a\) using \t \ basisB\ \a\basisA'\ \a'\basisA'\ by (auto simp: b_lincom G_def cblinfun.sum_left F_apply intro!: sum.neutral sum.cong) ultimately show ?thesis by blast qed have a_span: \cspan (\a\basisA'. cspan {F a b|b. b\basisB}) = UNIV\ proof (intro equalityI subset_UNIV subsetI, rename_tac H) fix H obtain G where G: \G a b \ cspan {F a b|b. b\basisB} \ G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using b_span by blast then have G_cspan: \G a b \ cspan {F a b|b. b\basisB}\ if \a\basisA'\ for a b using that by auto from G have G: \G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' using that by auto define H' where \H' = (\a\basisA'. G a (H *\<^sub>V a))\ have \H' \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ unfolding H'_def using G_cspan by (smt (verit, del_insts) UN_iff complex_vector.span_clauses(1) complex_vector.span_sum) moreover have \H' = H\ using \cspan basisA' = UNIV\ apply (rule cblinfun_eq_on_UNIV_span) apply (auto simp: H'_def cblinfun.sum_left) apply (subst sum_single) by (auto simp: G) ultimately show \H \ cspan (\a\basisA'. cspan {F a b |b. b \ basisB})\ by simp qed moreover have \cspan basis \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ using F_basis by (smt (z3) UN_subset_iff complex_vector.span_alt complex_vector.span_minimal complex_vector.subspace_span mem_Collect_eq subset_iff) ultimately show \cspan basis = UNIV\ by auto qed instance cblinfun :: (\{cfinite_dim,complex_normed_vector}\, \{cfinite_dim,complex_normed_vector}\) cfinite_dim proof intro_classes obtain basisA :: \'a set\ where [simp]: \cspan basisA = UNIV\ \cindependent basisA\ \finite basisA\ using finite_basis by blast obtain basisB :: \'b set\ where [simp]: \cspan basisB = UNIV\ \cindependent basisB\ \finite basisB\ using finite_basis by blast define f where \f a b = cconstruct basisA (\x. if x=a then b else 0)\ for a :: 'a and b :: 'b have f_a: \f a b a = b\ if \a : basisA\ for a b by (simp add: complex_vector.construct_basis f_def that) have f_not_a: \f a b c = 0\ if \a : basisA\ and \c : basisA\ and \a \ c\for a b c using that by (simp add: complex_vector.construct_basis f_def) define F where \F a b = CBlinfun (f a b)\ for a b have \clinear (f a b)\ for a b by (auto intro: complex_vector.linear_construct simp: f_def) then have \bounded_clinear (f a b)\ for a b by auto then have F_apply: \cblinfun_apply (F a b) = f a b\ for a b by (simp add: F_def bounded_clinear_CBlinfun_apply) define basis where \basis = {F a b| a b. a\basisA \ b\basisB}\ have \cspan basis = UNIV\ apply (rule cblinfun_cspan_UNIV[where basisA=basisA and basisB=basisB]) apply (auto simp: basis_def) by (metis F_apply f_a f_not_a) moreover have \finite basis\ unfolding basis_def apply (rule finite_image_set2) by auto ultimately show \\S :: ('a \\<^sub>C\<^sub>L 'b) set. finite S \ cspan S = UNIV\ by auto qed lemma norm_cblinfun_bound_dense: assumes \0 \ b\ assumes S: \closure S = UNIV\ assumes bound: \\x. x\S \ norm (cblinfun_apply f x) \ b * norm x\ shows \norm f \ b\ proof - have 1: \continuous_on UNIV (\a. norm (f *\<^sub>V a))\ apply (intro continuous_on_norm linear_continuous_on) by (simp add: Complex_Vector_Spaces.bounded_clinear.bounded_linear cblinfun.bounded_clinear_right) have 2: \continuous_on UNIV (\a. b * norm a)\ using continuous_on_mult_left continuous_on_norm_id by blast have \norm (cblinfun_apply f x) \ b * norm x\ for x apply (rule on_closure_leI[where x=x and S=S]) using S bound 1 2 by auto then show \norm f \ b\ apply (rule_tac norm_cblinfun_bound) using \0 \ b\ by auto qed lemma dense_span_separating: \closure (cspan S) = UNIV \ bounded_clinear F \ bounded_clinear G \ (\x\S. F x = G x) \ F = G\ proof - fix F G :: \'a \ 'b\ assume dense: \closure (cspan S) = UNIV\ assume [simp]: \bounded_clinear F\ \bounded_clinear G\ assume \\x\S. F x = G x\ then have \F x = G x\ if \x \ cspan S\ for x apply (rule_tac complex_vector.linear_eq_on[of F G _ S]) using that by (auto simp: bounded_clinear.clinear) then show \F = G\ apply (rule_tac ext) apply (rule on_closure_eqI[of \cspan S\ F G]) using dense by (auto intro!: continuous_at_imp_continuous_on clinear_continuous_at) qed lemma infsum_cblinfun_apply: assumes \g summable_on S\ shows \infsum (\x. A *\<^sub>V g x) S = A *\<^sub>V (infsum g S)\ apply (rule infsum_bounded_linear[unfolded o_def, of \cblinfun_apply A\]) using assms by (auto simp add: bounded_clinear.bounded_linear bounded_cbilinear.bounded_clinear_right bounded_cbilinear_cblinfun_apply) lemma has_sum_cblinfun_apply: assumes \has_sum g S x\ shows \has_sum (\x. A *\<^sub>V g x) S (A *\<^sub>V x)\ apply (rule has_sum_bounded_linear[unfolded o_def, of \cblinfun_apply A\]) using assms by (auto simp add: bounded_clinear.bounded_linear cblinfun.bounded_clinear_right) lemma abs_summable_on_cblinfun_apply: assumes \g abs_summable_on S\ shows \(\x. A *\<^sub>V g x) abs_summable_on S\ using bounded_clinear.bounded_linear[OF cblinfun.bounded_clinear_right] assms by (rule abs_summable_on_bounded_linear[unfolded o_def]) text \The next eight lemmas logically belong in \<^theory>\Complex_Bounded_Operators.Complex_Inner_Product\ but the proofs use facts from this theory.\ lemma has_sum_cinner_left: assumes \has_sum f I x\ shows \has_sum (\i. cinner a (f i)) I (cinner a x)\ by (metis assms cblinfun_cinner_right.rep_eq has_sum_cblinfun_apply) lemma summable_on_cinner_left: assumes \f summable_on I\ shows \(\i. cinner a (f i)) summable_on I\ by (metis assms has_sum_cinner_left summable_on_def) lemma infsum_cinner_left: assumes \\ summable_on I\ shows \cinner \ (\\<^sub>\i\I. \ i) = (\\<^sub>\i\I. cinner \ (\ i))\ by (metis assms has_sum_cinner_left has_sum_infsum infsumI) lemma has_sum_cinner_right: assumes \has_sum f I x\ shows \has_sum (\i. f i \\<^sub>C a) I (x \\<^sub>C a)\ apply (rule has_sum_bounded_linear[where f=\\x. x \\<^sub>C a\, unfolded o_def]) using assms by (simp_all add: bounded_antilinear.bounded_linear bounded_antilinear_cinner_left) lemma summable_on_cinner_right: assumes \f summable_on I\ shows \(\i. f i \\<^sub>C a) summable_on I\ by (metis assms has_sum_cinner_right summable_on_def) lemma infsum_cinner_right: assumes \\ summable_on I\ shows \(\\<^sub>\i\I. \ i) \\<^sub>C \ = (\\<^sub>\i\I. \ i \\<^sub>C \)\ by (metis assms has_sum_cinner_right has_sum_infsum infsumI) lemma Cauchy_cinner_product_summable: assumes asum: \a summable_on UNIV\ assumes bsum: \b summable_on UNIV\ assumes \finite X\ \finite Y\ assumes pos: \\x y. x \ X \ y \ Y \ cinner (a x) (b y) \ 0\ shows absum: \(\(x, y). cinner (a x) (b y)) summable_on UNIV\ proof - have \(\(x,y)\F. norm (cinner (a x) (b y))) \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) + norm (infsum b (-Y)) + 1\ if \finite F\ and \F \ (-X)\(-Y)\ for F proof - from asum \finite X\ have \a summable_on (-X)\ by (simp add: Compl_eq_Diff_UNIV summable_on_cofin_subset) then obtain MA where \finite MA\ and \MA \ -X\ and MA: \G \ MA \ G \ -X \ finite G \ norm (sum a G - infsum a (-X)) \ 1\ for G apply (simp add: summable_iff_has_sum_infsum has_sum_metric dist_norm) by (meson less_eq_real_def zero_less_one) from bsum \finite Y\ have \b summable_on (-Y)\ by (simp add: Compl_eq_Diff_UNIV summable_on_cofin_subset) then obtain MB where \finite MB\ and \MB \ -Y\ and MB: \G \ MB \ G \ -Y \ finite G \ norm (sum b G - infsum b (-Y)) \ 1\ for G apply (simp add: summable_iff_has_sum_infsum has_sum_metric dist_norm) by (meson less_eq_real_def zero_less_one) define F1 F2 where \F1 = fst ` F \ MA\ and \F2 = snd ` F \ MB\ define t1 t2 where \t1 = sum a F1 - infsum a (-X)\ and \t2 = sum b F2 - infsum b (-Y)\ have [simp]: \finite F1\ \finite F2\ using F1_def F2_def \finite MA\ \finite MB\ that by auto have [simp]: \F1 \ -X\ \F2 \ -Y\ using \F \ (-X)\(-Y)\ \MA \ -X\ \MB \ -Y\ by (auto simp: F1_def F2_def) from MA[OF _ \F1 \ -X\ \finite F1\] have \norm t1 \ 1\ by (auto simp: t1_def F1_def) from MB[OF _ \F2 \ -Y\ \finite F2\] have \norm t2 \ 1\ by (auto simp: t2_def F2_def) have [simp]: \F \ F1 \ F2\ apply (auto simp: F1_def F2_def image_def) by force+ have \(\(x,y)\F. norm (cinner (a x) (b y))) \ (\(x,y)\F1\F2. norm (cinner (a x) (b y)))\ apply (rule sum_mono2) by auto also from pos have \\ = norm (\(x,y)\F1\F2. cinner (a x) (b y))\ apply (auto intro!: of_real_eq_iff[THEN iffD1] simp: case_prod_beta) apply (subst abs_complex_def[unfolded o_def, symmetric, THEN fun_cong])+ apply (subst (2) abs_pos) apply (rule sum_nonneg) apply (metis Compl_eq_Diff_UNIV Diff_iff SigmaE \F1 \ - X\ \F2 \ - Y\ fst_conv prod.sel(2) subsetD) apply (rule sum.cong) apply simp by (metis Compl_iff SigmaE \F1 \ - X\ \F2 \ - Y\ abs_pos fst_conv prod.sel(2) subset_eq) also have \\ = norm (cinner (sum a F1) (sum b F2))\ by (simp add: sum.cartesian_product sum_cinner) also have \\ = norm (cinner (infsum a (-X) + t1) (infsum b (-Y) + t2))\ by (simp add: t1_def t2_def) also have \\ \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) * norm t2 + norm t1 * norm (infsum b (-Y)) + norm t1 * norm t2\ apply (simp add: cinner_add_right cinner_add_left) by (smt (verit, del_insts) complex_inner_class.Cauchy_Schwarz_ineq2 norm_triangle_ineq) also from \norm t1 \ 1\ \norm t2 \ 1\ have \\ \ norm (cinner (infsum a (-X)) (infsum b (-Y))) + norm (infsum a (-X)) + norm (infsum b (-Y)) + 1\ by (auto intro!: add_mono mult_left_le mult_left_le_one_le mult_le_one) finally show ?thesis by - qed then have \(\(x, y). cinner (a x) (b y)) abs_summable_on (-X)\(-Y)\ apply (rule_tac nonneg_bdd_above_summable_on) by (auto intro!: bdd_aboveI2 simp: case_prod_unfold) then have 1: \(\(x, y). cinner (a x) (b y)) summable_on (-X)\(-Y)\ using abs_summable_summable by blast from bsum have \(\y. b y) summable_on (-Y)\ by (simp add: Compl_eq_Diff_UNIV assms(4) summable_on_cofin_subset) then have \(\y. cinner (a x) (b y)) summable_on (-Y)\ for x using summable_on_cinner_left by blast with \finite X\ have 2: \(\(x, y). cinner (a x) (b y)) summable_on X\(-Y)\ apply (rule_tac summable_on_product_finite_left) by auto from asum have \(\x. a x) summable_on (-X)\ by (simp add: Compl_eq_Diff_UNIV assms(3) summable_on_cofin_subset) then have \(\x. cinner (a x) (b y)) summable_on (-X)\ for y using summable_on_cinner_right by blast with \finite Y\ have 3: \(\(x, y). cinner (a x) (b y)) summable_on (-X)\Y\ apply (rule_tac summable_on_product_finite_right) by auto have 4: \(\(x, y). cinner (a x) (b y)) summable_on X\Y\ by (simp add: \finite X\ \finite Y\) show ?thesis apply (subst asm_rl[of \UNIV = (-X)\(-Y) \ X\(-Y) \ (-X)\Y \ X\Y\]) using 1 2 3 4 by (auto intro!: summable_on_Un_disjoint) qed text \A variant of @{thm [source] Series.Cauchy_product_sums} with \<^term>\(*)\ replaced by \<^term>\cinner\. Differently from @{thm [source] Series.Cauchy_product_sums}, we do not require absolute summability of \<^term>\a\ and \<^term>\b\ individually but only unconditional summability of \<^term>\a\, \<^term>\b\, and their product. While on, e.g., reals, unconditional summability is equivalent to absolute summability, in general unconditional summability is a weaker requirement.\ lemma fixes a b :: "nat \ 'a::complex_inner" assumes asum: \a summable_on UNIV\ assumes bsum: \b summable_on UNIV\ assumes absum: \(\(x, y). cinner (a x) (b y)) summable_on UNIV\ (* \ \See @{thm [source] Cauchy_cinner_product_summable} or @{thm [source] Cauchy_cinner_product_summable'} for a way to rewrite this premise.\ *) shows Cauchy_cinner_product_infsum: \(\\<^sub>\k. \i\k. cinner (a i) (b (k - i))) = cinner (\\<^sub>\k. a k) (\\<^sub>\k. b k)\ and Cauchy_cinner_product_infsum_exists: \(\k. \i\k. cinner (a i) (b (k - i))) summable_on UNIV\ proof - have img: \(\(k::nat, i). (i, k - i)) ` {(k, i). i \ k} = UNIV\ apply (auto simp: image_def) by (metis add.commute add_diff_cancel_right' diff_le_self) have inj: \inj_on (\(k::nat, i). (i, k - i)) {(k, i). i \ k}\ by (smt (verit, del_insts) Pair_inject case_prodE case_prod_conv eq_diff_iff inj_onI mem_Collect_eq) have sigma: \(SIGMA k:UNIV. {i. i \ k}) = {(k, i). i \ k}\ by auto from absum have \(\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ by (rule Cauchy_cinner_product_summable'[THEN iffD1]) then have \(\k. \\<^sub>\i|i\k. cinner (a i) (b (k-i))) summable_on UNIV\ by (metis (mono_tags, lifting) sigma summable_on_Sigma_banach summable_on_cong) then show \(\k. \i\k. cinner (a i) (b (k - i))) summable_on UNIV\ by (metis (mono_tags, lifting) atMost_def finite_Collect_le_nat infsum_finite summable_on_cong) have \cinner (\\<^sub>\k. a k) (\\<^sub>\k. b k) = (\\<^sub>\k. \\<^sub>\l. cinner (a k) (b l))\ apply (subst infsum_cinner_right) apply (rule asum) apply (subst infsum_cinner_left) apply (rule bsum) by simp also have \\ = (\\<^sub>\(k,l). cinner (a k) (b l))\ apply (subst infsum_Sigma'_banach) using absum by auto also have \\ = (\\<^sub>\(k, l)\(\(k, i). (i, k - i)) ` {(k, i). i \ k}. cinner (a k) (b l))\ by (simp only: img) also have \\ = infsum ((\(k, l). a k \\<^sub>C b l) \ (\(k, i). (i, k - i))) {(k, i). i \ k}\ using inj by (rule infsum_reindex) also have \\ = (\\<^sub>\(k,i)|i\k. a i \\<^sub>C b (k-i))\ by (simp add: o_def case_prod_unfold) also have \\ = (\\<^sub>\k. \\<^sub>\i|i\k. a i \\<^sub>C b (k-i))\ apply (subst infsum_Sigma'_banach) using absum by (auto simp: sigma Cauchy_cinner_product_summable') also have \\ = (\\<^sub>\k. \i\k. a i \\<^sub>C b (k-i))\ apply (subst infsum_finite[symmetric]) by (auto simp add: atMost_def) finally show \(\\<^sub>\k. \i\k. a i \\<^sub>C b (k - i)) = (\\<^sub>\k. a k) \\<^sub>C (\\<^sub>\k. b k)\ by simp qed lemma CBlinfun_plus: assumes [simp]: \bounded_clinear f\ \bounded_clinear g\ shows \CBlinfun (f + g) = CBlinfun f + CBlinfun g\ by (auto intro!: cblinfun_eqI simp: plus_fun_def bounded_clinear_add CBlinfun_inverse cblinfun.add_left) lemma CBlinfun_scaleC: assumes \bounded_clinear f\ shows \CBlinfun (\y. c *\<^sub>C f y) = c *\<^sub>C CBlinfun f\ by (simp add: assms eq_onp_same_args scaleC_cblinfun.abs_eq) lemma bounded_clinear_inv: assumes [simp]: \bounded_clinear f\ assumes b: \b > 0\ assumes bound: \\x. norm (f x) \ b * norm x\ assumes \surj f\ shows \bounded_clinear (inv f)\ proof (rule bounded_clinear_intro) fix x y :: 'b and r :: complex define x' y' where \x' = inv f x\ and \y' = inv f y\ have [simp]: \clinear f\ by (simp add: bounded_clinear.clinear) have [simp]: \inj f\ proof (rule injI) fix x y assume \f x = f y\ then have \norm (f (x - y)) = 0\ by (simp add: complex_vector.linear_diff) with bound b have \norm (x - y) = 0\ by (metis linorder_not_le mult_le_0_iff nle_le norm_ge_zero) then show \x = y\ by simp qed from \surj f\ have [simp]: \x = f x'\ \y = f y'\ by (simp_all add: surj_f_inv_f x'_def y'_def) show "inv f (x + y) = inv f x + inv f y" by (simp flip: complex_vector.linear_add) show "inv f (r *\<^sub>C x) = r *\<^sub>C inv f x" by (simp flip: clinear.scaleC) from bound have "b * norm (inv f x) \ norm x" by (simp flip: clinear.scaleC) with b show "norm (inv f x) \ norm x * inverse b" by (smt (verit, ccfv_threshold) left_inverse mult.commute mult_cancel_right1 mult_le_cancel_left_pos vector_space_over_itself.scale_scale) qed +lemma cinner_sup_norm_cblinfun: + fixes A :: \'a::{complex_normed_vector,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ + shows \norm A = (SUP (\,\). cmod (cinner \ (A *\<^sub>V \)) / (norm \ * norm \))\ + apply transfer + apply (rule cinner_sup_onorm) + by (simp add: bounded_clinear.bounded_linear) + + +lemma norm_cblinfun_Sup: \norm A = (SUP \. norm (A *\<^sub>V \) / norm \)\ + by (simp add: norm_cblinfun.rep_eq onorm_def) subsection \Relationship to real bounded operators (\<^typ>\_ \\<^sub>L _\)\ instantiation blinfun :: (real_normed_vector, complex_normed_vector) "complex_normed_vector" begin lift_definition scaleC_blinfun :: \complex \ ('a::real_normed_vector, 'b::complex_normed_vector) blinfun \ ('a, 'b) blinfun\ is \\ c::complex. \ f::'a\'b. (\ x. c *\<^sub>C (f x) )\ proof fix c::complex and f :: \'a\'b\ and b1::'a and b2::'a assume \bounded_linear f\ show \c *\<^sub>C f (b1 + b2) = c *\<^sub>C f b1 + c *\<^sub>C f b2\ by (simp add: \bounded_linear f\ linear_simps scaleC_add_right) fix c::complex and f :: \'a\'b\ and b::'a and r::real assume \bounded_linear f\ show \c *\<^sub>C f (r *\<^sub>R b) = r *\<^sub>R (c *\<^sub>C f b)\ by (simp add: \bounded_linear f\ linear_simps(5) scaleR_scaleC) fix c::complex and f :: \'a\'b\ assume \bounded_linear f\ have \\ K. \ x. norm (f x) \ norm x * K\ using \bounded_linear f\ by (simp add: bounded_linear.bounded) then obtain K where \\ x. norm (f x) \ norm x * K\ by blast have \cmod c \ 0\ by simp hence \\ x. (cmod c) * norm (f x) \ (cmod c) * norm x * K\ using \\ x. norm (f x) \ norm x * K\ by (metis ordered_comm_semiring_class.comm_mult_left_mono vector_space_over_itself.scale_scale) moreover have \norm (c *\<^sub>C f x) = (cmod c) * norm (f x)\ for x by simp ultimately show \\K. \x. norm (c *\<^sub>C f x) \ norm x * K\ by (metis ab_semigroup_mult_class.mult_ac(1) mult.commute) qed instance proof have "r *\<^sub>R x = complex_of_real r *\<^sub>C x" for x :: "('a, 'b) blinfun" and r apply transfer by (simp add: scaleR_scaleC) thus "((*\<^sub>R) r::'a \\<^sub>L 'b \ _) = (*\<^sub>C) (complex_of_real r)" for r by auto show "a *\<^sub>C (x + y) = a *\<^sub>C x + a *\<^sub>C y" for a :: complex and x y :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_right) show "(a + b) *\<^sub>C x = a *\<^sub>C x + b *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_left) show "a *\<^sub>C b *\<^sub>C x = (a * b) *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by simp have \1 *\<^sub>C f x = f x\ for f :: \'a\'b\ and x by auto thus "1 *\<^sub>C x = x" for x :: "'a \\<^sub>L 'b" by (simp add: scaleC_blinfun.rep_eq blinfun_eqI) have \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ if \bounded_linear f\ for f :: \'a \ 'b\ and a :: complex proof- have \cmod a \ 0\ by simp have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \\ x. (cmod a) *(\ ereal ((norm (f x)) / (norm x)) \) \ (cmod a) * K\ using \cmod a \ 0\ by (metis abs_ereal.simps(1) abs_ereal_pos abs_pos ereal_mult_left_mono times_ereal.simps(1)) hence \\ x. (\ ereal ((cmod a) * (norm (f x)) / (norm x)) \) \ (cmod a) * K\ by simp hence \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have p1: \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ using \\ x. \ ereal (cmod a * (norm (f x)) / (norm x)) \ \ cmod a * K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) have p2: \\i. i \ UNIV \ 0 \ ereal (cmod a * norm (f i) / norm i)\ by simp hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ (SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\)\ using \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I p2 abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ cmod a * K\ using \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ by simp hence \\ ( SUP i\UNIV::'a set. ereal ((\ x. (cmod a) * (norm (f x)) / norm x) i)) \ \ \\ by auto hence w2: \( SUP i\UNIV::'a set. ereal ((\ x. cmod a * (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. cmod a * (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) have \(UNIV::('a set)) \ {}\ by simp moreover have \\ i. i \ (UNIV::('a set)) \ (\ x. (norm (f x)) / norm x :: ereal) i \ 0\ by simp moreover have \cmod a \ 0\ by simp ultimately have \(SUP i\(UNIV::('a set)). ((cmod a)::ereal) * (\ x. (norm (f x)) / norm x :: ereal) i ) = ((cmod a)::ereal) * ( SUP i\(UNIV::('a set)). (\ x. (norm (f x)) / norm x :: ereal) i )\ by (simp add: Sup_ereal_mult_left') hence \(SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp hence z1: \real_of_ereal ( (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) ) = real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) )\ by simp have z2: \real_of_ereal (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = (SUP x. cmod a * (norm (f x) / norm x))\ using w2 by auto have \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = (cmod a) * real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp moreover have \real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) ) = ( SUP x. ((norm (f x)) / norm x) )\ proof- have \\ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) \ \ \\ proof- have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ using \\ x. \ ereal ((norm (f x)) / (norm x)) \ \ K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ (SUP x. \ereal ((norm (f x)) / (norm x))\)\ using \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I \\i. i \ UNIV \ 0 \ ereal (norm (f i) / norm i)\ abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ K\ using \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ by simp thus ?thesis by auto qed hence \ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) thus ?thesis by simp qed have z3: \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = cmod a * (SUP x. norm (f x) / norm x)\ by (simp add: \real_of_ereal (SUP x. ereal (norm (f x) / norm x)) = (SUP x. norm (f x) / norm x)\) hence w1: \(SUP x. cmod a * (norm (f x) / norm x)) = cmod a * (SUP x. norm (f x) / norm x)\ using z1 z2 by linarith have v1: \onorm (\x. a *\<^sub>C f x) = (SUP x. norm (a *\<^sub>C f x) / norm x)\ by (simp add: onorm_def) have v2: \(SUP x. norm (a *\<^sub>C f x) / norm x) = (SUP x. ((cmod a) * norm (f x)) / norm x)\ by simp have v3: \(SUP x. ((cmod a) * norm (f x)) / norm x) = (SUP x. (cmod a) * ((norm (f x)) / norm x))\ by simp have v4: \(SUP x. (cmod a) * ((norm (f x)) / norm x)) = (cmod a) * (SUP x. ((norm (f x)) / norm x))\ using w1 by blast show \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ using v1 v2 v3 v4 by (metis (mono_tags, lifting) onorm_def) qed thus \norm (a *\<^sub>C x) = cmod a * norm x\ for a::complex and x::\('a, 'b) blinfun\ apply transfer by blast qed end (* We do not have clinear_blinfun_compose_right *) lemma clinear_blinfun_compose_left: \clinear (\x. blinfun_compose x y)\ by (auto intro!: clinearI simp: blinfun_eqI scaleC_blinfun.rep_eq bounded_bilinear.add_left bounded_bilinear_blinfun_compose) instantiation blinfun :: (real_normed_vector, cbanach) "cbanach" begin instance.. end lemma blinfun_compose_assoc: "(A o\<^sub>L B) o\<^sub>L C = A o\<^sub>L (B o\<^sub>L C)" by (simp add: blinfun_eqI) lift_definition blinfun_of_cblinfun::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a \\<^sub>L 'b\ is "id" apply transfer by (simp add: bounded_clinear.bounded_linear) lift_definition blinfun_cblinfun_eq :: \'a \\<^sub>L 'b \ 'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ bool\ is "(=)" . lemma blinfun_cblinfun_eq_bi_unique[transfer_rule]: \bi_unique blinfun_cblinfun_eq\ unfolding bi_unique_def apply transfer by auto lemma blinfun_cblinfun_eq_right_total[transfer_rule]: \right_total blinfun_cblinfun_eq\ unfolding right_total_def apply transfer by (simp add: bounded_clinear.bounded_linear) named_theorems cblinfun_blinfun_transfer lemma cblinfun_blinfun_transfer_0[cblinfun_blinfun_transfer]: "blinfun_cblinfun_eq (0::(_,_) blinfun) (0::(_,_) cblinfun)" apply transfer by simp lemma cblinfun_blinfun_transfer_plus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (+) (+)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_minus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (-) (-)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_uminus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (uminus) (uminus)" unfolding rel_fun_def apply transfer by auto definition "real_complex_eq r c \ complex_of_real r = c" lemma bi_unique_real_complex_eq[transfer_rule]: \bi_unique real_complex_eq\ unfolding real_complex_eq_def bi_unique_def by auto lemma left_total_real_complex_eq[transfer_rule]: \left_total real_complex_eq\ unfolding real_complex_eq_def left_total_def by auto lemma cblinfun_blinfun_transfer_scaleC[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(real_complex_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (scaleR) (scaleC)" unfolding rel_fun_def apply transfer by (simp add: real_complex_eq_def scaleR_scaleC) lemma cblinfun_blinfun_transfer_CBlinfun[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(eq_onp bounded_clinear ===> blinfun_cblinfun_eq) Blinfun CBlinfun" unfolding rel_fun_def blinfun_cblinfun_eq.rep_eq eq_onp_def by (auto simp: CBlinfun_inverse Blinfun_inverse bounded_clinear.bounded_linear) lemma cblinfun_blinfun_transfer_norm[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=)) norm norm" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_dist[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> (=)) dist dist" unfolding rel_fun_def dist_norm apply transfer by auto lemma cblinfun_blinfun_transfer_sgn[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) sgn sgn" unfolding rel_fun_def sgn_blinfun_def sgn_cblinfun_def apply transfer by (auto simp: scaleR_scaleC) lemma cblinfun_blinfun_transfer_Cauchy[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> (=)) Cauchy Cauchy" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding Cauchy_def by transfer_prover qed lemma cblinfun_blinfun_transfer_tendsto[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> blinfun_cblinfun_eq ===> (=) ===> (=)) tendsto tendsto" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding tendsto_iff by transfer_prover qed lemma cblinfun_blinfun_transfer_compose[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (o\<^sub>L) (o\<^sub>C\<^sub>L)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_apply[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=) ===> (=)) blinfun_apply cblinfun_apply" unfolding rel_fun_def apply transfer by auto lemma blinfun_of_cblinfun_inj: \blinfun_of_cblinfun f = blinfun_of_cblinfun g \ f = g\ by (metis cblinfun_apply_inject blinfun_of_cblinfun.rep_eq) lemma blinfun_of_cblinfun_inv: assumes "\c. \x. f *\<^sub>v (c *\<^sub>C x) = c *\<^sub>C (f *\<^sub>v x)" shows "\g. blinfun_of_cblinfun g = f" using assms proof transfer show "\g\Collect bounded_clinear. id g = f" if "bounded_linear f" and "\c x. f (c *\<^sub>C x) = c *\<^sub>C f x" for f :: "'a \ 'b" using that bounded_linear_bounded_clinear by auto qed lemma blinfun_of_cblinfun_zero: \blinfun_of_cblinfun 0 = 0\ apply transfer by simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ apply transfer by auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_norm: fixes f::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ shows \norm f = norm (blinfun_of_cblinfun f)\ apply transfer by auto subsection \Composition\ lemma blinfun_of_cblinfun_cblinfun_compose: fixes f::\'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector\ and g::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b\ shows \blinfun_of_cblinfun (f o\<^sub>C\<^sub>L g) = (blinfun_of_cblinfun f) o\<^sub>L (blinfun_of_cblinfun g)\ apply transfer by auto lemma cblinfun_compose_assoc: shows "(A o\<^sub>C\<^sub>L B) o\<^sub>C\<^sub>L C = A o\<^sub>C\<^sub>L (B o\<^sub>C\<^sub>L C)" by (metis (no_types, lifting) cblinfun_apply_inject fun.map_comp cblinfun_compose.rep_eq) lemma cblinfun_compose_zero_right[simp]: "U o\<^sub>C\<^sub>L 0 = 0" using bounded_cbilinear.zero_right bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_zero_left[simp]: "0 o\<^sub>C\<^sub>L U = 0" using bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_scaleC_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>C A) o\<^sub>C\<^sub>L B = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ by (simp add: bounded_cbilinear.scaleC_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_scaleR_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>R A) o\<^sub>C\<^sub>L B = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_scaleC_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>C B) = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ apply transfer by (auto intro!: ext bounded_clinear.clinear complex_vector.linear_scale) lemma cblinfun_compose_scaleR_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>R B) = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_id_right[simp]: shows "U o\<^sub>C\<^sub>L id_cblinfun = U" apply transfer by auto lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" apply transfer by auto lemma cblinfun_eq_on: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \t \ closure (cspan G)\ shows "A *\<^sub>V t = B *\<^sub>V t" using assms apply transfer using bounded_clinear_eq_on by blast lemma cblinfun_eq_gen_eqI: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \ccspan G = \\ shows "A = B" apply (rule cblinfun_eqI) apply (rule cblinfun_eq_on[where G=G]) using assms apply auto by (metis ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) lemma cblinfun_compose_add_left: \(a + b) o\<^sub>C\<^sub>L c = (a o\<^sub>C\<^sub>L c) + (b o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_add_right: \a o\<^sub>C\<^sub>L (b + c) = (a o\<^sub>C\<^sub>L b) + (a o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) lemma cbilinear_cblinfun_compose[simp]: "cbilinear cblinfun_compose" by (auto intro!: clinearI simp add: cbilinear_def bounded_cbilinear.add_left bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_sum_left: \(\i\S. g i) o\<^sub>C\<^sub>L x = (\i\S. g i o\<^sub>C\<^sub>L x)\ apply (induction S rule:infinite_finite_induct) by (auto simp: cblinfun_compose_add_left) lemma cblinfun_compose_sum_right: \x o\<^sub>C\<^sub>L (\i\S. g i) = (\i\S. x o\<^sub>C\<^sub>L g i)\ apply (induction S rule:infinite_finite_induct) by (auto simp: cblinfun_compose_add_right) lemma cblinfun_compose_minus_right: \a o\<^sub>C\<^sub>L (b - c) = (a o\<^sub>C\<^sub>L b) - (a o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.diff_right bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_minus_left: \(a - b) o\<^sub>C\<^sub>L c = (a o\<^sub>C\<^sub>L c) - (b o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.diff_left bounded_cbilinear_cblinfun_compose) lemma simp_a_oCL_b: \a o\<^sub>C\<^sub>L b = c \ a o\<^sub>C\<^sub>L (b o\<^sub>C\<^sub>L d) = c o\<^sub>C\<^sub>L d\ \ \A convenience lemma to transform simplification rules of the form \<^term>\a o\<^sub>C\<^sub>L b = c\. E.g., \simp_a_oCL_b[OF isometryD, simp]\ declares a simp-rule for simplifying \<^term>\adj x o\<^sub>C\<^sub>L (x o\<^sub>C\<^sub>L y) = id_cblinfun o\<^sub>C\<^sub>L y\.\ by (auto simp: cblinfun_compose_assoc) lemma simp_a_oCL_b': \a o\<^sub>C\<^sub>L b = c \ a *\<^sub>V (b *\<^sub>V d) = c *\<^sub>V d\ \ \A convenience lemma to transform simplification rules of the form \<^term>\a o\<^sub>C\<^sub>L b = c\. E.g., \simp_a_oCL_b'[OF isometryD, simp]\ declares a simp-rule for simplifying \<^term>\adj x *\<^sub>V x *\<^sub>V y = id_cblinfun *\<^sub>V y\.\ by auto lemma cblinfun_compose_uminus_left: \(- a) o\<^sub>C\<^sub>L b = - (a o\<^sub>C\<^sub>L b)\ by (simp add: bounded_cbilinear.minus_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_uminus_right: \a o\<^sub>C\<^sub>L (- b) = - (a o\<^sub>C\<^sub>L b)\ by (simp add: bounded_cbilinear.minus_right bounded_cbilinear_cblinfun_compose) subsection \Adjoint\ lift_definition adj :: "'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ 'b \\<^sub>C\<^sub>L 'a" ("_*" [99] 100) is cadjoint by (fact cadjoint_bounded_clinear) lemma id_cblinfun_adjoint[simp]: "id_cblinfun* = id_cblinfun" apply transfer using cadjoint_id by (metis eq_id_iff) lemma double_adj[simp]: "(A*)* = A" apply transfer using double_cadjoint by blast lemma adj_cblinfun_compose[simp]: fixes B::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and A::\'b \\<^sub>C\<^sub>L 'c::complex_inner\ shows "(A o\<^sub>C\<^sub>L B)* = (B*) o\<^sub>C\<^sub>L (A*)" proof transfer fix A :: \'b \ 'c\ and B :: \'a \ 'b\ assume \bounded_clinear A\ and \bounded_clinear B\ hence \bounded_clinear (A \ B)\ by (simp add: comp_bounded_clinear) have \\ (A \ B) u, v \ = \ u, (B\<^sup>\ \ A\<^sup>\) v \\ for u v by (metis (no_types, lifting) cadjoint_univ_prop \bounded_clinear A\ \bounded_clinear B\ cinner_commute' comp_def) thus \(A \ B)\<^sup>\ = B\<^sup>\ \ A\<^sup>\\ using \bounded_clinear (A \ B)\ by (metis cadjoint_eqI cinner_commute') qed lemma scaleC_adj[simp]: "(a *\<^sub>C A)* = (cnj a) *\<^sub>C (A*)" apply transfer by (simp add: bounded_clinear.bounded_linear bounded_clinear_def complex_vector.linear_scale scaleC_cadjoint) lemma scaleR_adj[simp]: "(a *\<^sub>R A)* = a *\<^sub>R (A*)" by (simp add: scaleR_scaleC) lemma adj_plus: \(A + B)* = (A*) + (B*)\ proof transfer fix A B::\'b \ 'a\ assume a1: \bounded_clinear A\ and a2: \bounded_clinear B\ define F where \F = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ define G where \G = (\x. A x + B x)\ have \bounded_clinear G\ unfolding G_def by (simp add: a1 a2 bounded_clinear_add) moreover have \\F u, v\ = \u, G v\\ for u v unfolding F_def G_def using cadjoint_univ_prop a1 a2 cinner_add_left by (simp add: cadjoint_univ_prop cinner_add_left cinner_add_right) ultimately have \F = G\<^sup>\ \ using cadjoint_eqI by blast thus \(\x. A x + B x)\<^sup>\ = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ unfolding F_def G_def by auto qed -lemma cinner_sup_norm_cblinfun: - fixes A :: \'a::{complex_normed_vector,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ - shows \norm A = (SUP (\,\). cmod (cinner \ (A *\<^sub>V \)) / (norm \ * norm \))\ - apply transfer - apply (rule cinner_sup_onorm) - by (simp add: bounded_clinear.bounded_linear) - lemma cinner_adj_left: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\G* *\<^sub>V x, y\ = \x, G *\<^sub>V y\\ apply transfer using cadjoint_univ_prop by blast lemma cinner_adj_right: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\x, G* *\<^sub>V y\ = \G *\<^sub>V x, y\\ apply transfer using cadjoint_univ_prop' by blast lemma adj_0[simp]: \0* = 0\ by (metis add_cancel_right_left adj_plus) lemma norm_adj[simp]: \norm (A*) = norm A\ for A :: \'b::chilbert_space \\<^sub>C\<^sub>L 'c::complex_inner\ proof (cases \(\x y :: 'b. x \ y) \ (\x y :: 'c. x \ y)\) case True then have c1: \class.not_singleton TYPE('b)\ apply intro_classes by simp from True have c2: \class.not_singleton TYPE('c)\ apply intro_classes by simp have normA: \norm A = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c1) have normAadj: \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A* *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c2) have \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ unfolding normAadj apply (subst cinner_adj_right) apply (subst cinner_commute) apply (subst complex_mod_cnj) by rule also have \\ = Sup ((\(\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \)) ` prod.swap ` UNIV)\ by auto also have \\ = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (subst image_image) by auto also have \\ = norm A\ unfolding normA by (simp add: mult.commute) finally show ?thesis by - next case False then consider (b) \\x::'b. x = 0\ | (c) \\x::'c. x = 0\ by auto then have \A = 0\ apply (cases; transfer) apply (metis (full_types) bounded_clinear_def complex_vector.linear_0) by auto then show \norm (A*) = norm A\ by simp qed lemma antilinear_adj[simp]: \antilinear adj\ apply (rule antilinearI) by (auto simp add: adj_plus) lemma bounded_antilinear_adj[bounded_antilinear, simp]: \bounded_antilinear adj\ by (auto intro!: antilinearI exI[of _ 1] simp: bounded_antilinear_def bounded_antilinear_axioms_def adj_plus) lemma adjoint_eqI: fixes G:: \'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner\ 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 adj_uminus: \(-A)* = - (A*)\ apply (rule adjoint_eqI[symmetric]) by (simp add: cblinfun.minus_left cinner_adj_left) lemma cinner_real_hermiteanI: \ \Prop. II.2.12 in @{cite conway2013course}\ assumes \\\. cinner \ (A *\<^sub>V \) \ \\ shows \A = A*\ proof - { fix g h :: 'a { fix \ :: complex have \cinner h (A h) + cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) + (abs \)\<^sup>2 * cinner g (A g) = cinner (h + \ *\<^sub>C g) (A *\<^sub>V (h + \ *\<^sub>C g))\ (is \?sum4 = _\) apply (auto simp: cinner_add_right cinner_add_left cblinfun.add_right cblinfun.scaleC_right ring_class.ring_distribs) by (metis cnj_x_x mult.commute) also have \\ \ \\ using assms by auto finally have \?sum4 = cnj ?sum4\ using Reals_cnj_iff by fastforce then have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner (A h) g + cnj \ *\<^sub>C cinner (A g) h\ using Reals_cnj_iff abs_complex_real assms by force also have \\ = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by (simp add: cinner_adj_right) finally have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by - } from this[where \2=1] this[where \2=\] have 1: \cinner g (A h) + cinner h (A g) = cinner h (A* *\<^sub>V g) + cinner g (A* *\<^sub>V h)\ and i: \- \ * cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) - \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by auto from arg_cong2[OF 1 arg_cong[OF i, where f=\(*) (-\)\], where f=plus] have \cinner h (A g) = cinner h (A* *\<^sub>V g)\ by (auto simp: ring_class.ring_distribs) } then show "A = A*" by (simp add: adjoint_eqI cinner_adj_right) qed lemma norm_AAadj[simp]: \norm (A o\<^sub>C\<^sub>L A*) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof (cases \class.not_singleton TYPE('b)\) case True then have [simp]: \class.not_singleton TYPE('b)\ by - have 1: \(norm A)\<^sup>2 * \ \ norm (A o\<^sub>C\<^sub>L A*)\ if \\ < 1\ and \\ \ 0\ for \ proof - obtain \ where \: \norm ((A*) *\<^sub>V \) \ norm (A*) * sqrt \\ and [simp]: \norm \ = 1\ apply atomize_elim apply (rule cblinfun_norm_approx_witness_mult[internalize_sort' 'a]) using \\ < 1\ by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) have \complex_of_real ((norm A)\<^sup>2 * \) = (norm (A*) * sqrt \)\<^sup>2\ by (simp add: ordered_field_class.sign_simps(23) that(2)) also have \\ \ (norm ((A* *\<^sub>V \)))\<^sup>2\ apply (rule complex_of_real_mono) using \ apply (rule power_mono) using \\ \ 0\ by auto also have \\ \ cinner (A* *\<^sub>V \) (A* *\<^sub>V \)\ by (auto simp flip: power2_norm_eq_cinner) also have \\ = cinner \ (A *\<^sub>V A* *\<^sub>V \)\ by (simp add: cinner_adj_left) also have \\ = cinner \ ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ by auto also have \\ \ norm (A o\<^sub>C\<^sub>L A*)\ using \norm \ = 1\ by (smt (verit, best) Im_complex_of_real Re_complex_of_real \(A* *\<^sub>V \) \\<^sub>C (A* *\<^sub>V \) = \ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \)\ \\ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \) = \ \\<^sub>C ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ cdot_square_norm cinner_ge_zero cmod_Re complex_inner_class.Cauchy_Schwarz_ineq2 less_eq_complex_def mult_cancel_left1 mult_cancel_right1 norm_cblinfun) finally show ?thesis by (auto simp: less_eq_complex_def) qed then have 1: \(norm A)\<^sup>2 \ norm (A o\<^sub>C\<^sub>L A*)\ by (metis field_le_mult_one_interval less_eq_real_def ordered_field_class.sign_simps(5)) have 2: \norm (A o\<^sub>C\<^sub>L A*) \ (norm A)\<^sup>2\ proof (rule norm_cblinfun_bound) show \0 \ (norm A)\<^sup>2\ by simp fix \ have \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) = norm (A *\<^sub>V A* *\<^sub>V \)\ by auto also have \\ \ norm A * norm (A* *\<^sub>V \)\ by (simp add: norm_cblinfun) also have \\ \ norm A * norm (A*) * norm \\ by (metis mult.assoc norm_cblinfun norm_imp_pos_and_ge ordered_comm_semiring_class.comm_mult_left_mono) also have \\ = (norm A)\<^sup>2 * norm \\ by (simp add: power2_eq_square) finally show \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) \ (norm A)\<^sup>2 * norm \\ by - qed from 1 2 show ?thesis by simp next case False then have [simp]: \class.CARD_1 TYPE('b)\ by (rule not_singleton_vs_CARD_1) have \A = 0\ apply (rule cblinfun_to_CARD_1_0[internalize_sort' 'b]) by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) then show ?thesis by auto qed lemma sum_adj: \(sum a F)* = sum (\i. (a i)*) F\ apply (induction rule:infinite_finite_induct) by (auto simp add: adj_plus) lemma has_sum_adj: assumes \has_sum f I x\ shows \has_sum (\x. adj (f x)) I (adj x)\ apply (rule has_sum_comm_additive[where f=adj, unfolded o_def]) apply (simp add: antilinear.axioms(1)) apply (metis (no_types, lifting) LIM_eq adj_plus adj_uminus norm_adj uminus_add_conv_diff) by (simp add: assms) lemma adj_minus: \(A - B)* = (A*) - (B*)\ by (metis add_implies_diff adj_plus diff_add_cancel) lemma cinner_hermitian_real: \x \\<^sub>C (A *\<^sub>V x) \ \\ if \A* = A\ by (metis Reals_cnj_iff cinner_adj_right cinner_commute' that) lemma adj_inject: \adj a = adj b \ a = b\ by (metis (no_types, opaque_lifting) adj_minus eq_iff_diff_eq_0 norm_adj norm_eq_zero) lemma norm_AadjA[simp]: \norm (A* o\<^sub>C\<^sub>L A) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ by (metis double_adj norm_AAadj norm_adj) subsection \Powers of operators\ lift_definition cblinfun_power :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'a \ nat \ 'a \\<^sub>C\<^sub>L 'a\ is \\(a::'a\'a) n. a ^^ n\ apply (rename_tac f n, induct_tac n, auto simp: Nat.funpow_code_def) by (simp add: bounded_clinear_compose) lemma cblinfun_power_0[simp]: \cblinfun_power A 0 = id_cblinfun\ apply transfer by auto lemma cblinfun_power_Suc': \cblinfun_power A (Suc n) = A o\<^sub>C\<^sub>L cblinfun_power A n\ apply transfer by auto lemma cblinfun_power_Suc: \cblinfun_power A (Suc n) = cblinfun_power A n o\<^sub>C\<^sub>L A\ apply (induction n) by (auto simp: cblinfun_power_Suc' simp flip: cblinfun_compose_assoc) lemma cblinfun_power_compose[simp]: \cblinfun_power A n o\<^sub>C\<^sub>L cblinfun_power A m = cblinfun_power A (n+m)\ apply (induction n) by (auto simp: cblinfun_power_Suc' cblinfun_compose_assoc) lemma cblinfun_power_scaleC: \cblinfun_power (c *\<^sub>C a) n = c^n *\<^sub>C cblinfun_power a n\ apply (induction n) by (auto simp: cblinfun_power_Suc) lemma cblinfun_power_scaleR: \cblinfun_power (c *\<^sub>R a) n = c^n *\<^sub>R cblinfun_power a n\ apply (induction n) by (auto simp: cblinfun_power_Suc) lemma cblinfun_power_uminus: \cblinfun_power (-a) n = (-1)^n *\<^sub>R cblinfun_power a n\ apply (subst asm_rl[of \-a = (-1) *\<^sub>R a\]) apply simp by (rule cblinfun_power_scaleR) lemma cblinfun_power_adj: \(cblinfun_power S n)* = cblinfun_power (S*) n\ apply (induction n) apply simp apply (subst cblinfun_power_Suc) apply (subst cblinfun_power_Suc') by auto subsection \Unitaries / isometries\ definition isometry::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun\ definition unitary::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \unitary U \ (U* o\<^sub>C\<^sub>L U = id_cblinfun) \ (U o\<^sub>C\<^sub>L U* = id_cblinfun)\ lemma unitary_twosided_isometry: "unitary U \ isometry U \ isometry (U*)" unfolding unitary_def isometry_def by simp lemma isometryD[simp]: "isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding isometry_def by simp (* Not [simp] because isometryD[simp] + unitary_isometry[simp] already have the same effect *) lemma unitaryD1: "unitary U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding unitary_def by simp lemma unitaryD2[simp]: "unitary U \ U o\<^sub>C\<^sub>L U* = id_cblinfun" unfolding unitary_def by simp lemma unitary_isometry[simp]: "unitary U \ isometry U" unfolding unitary_def isometry_def by simp lemma unitary_adj[simp]: "unitary (U*) = unitary U" unfolding unitary_def by auto lemma isometry_cblinfun_compose[simp]: assumes "isometry A" and "isometry B" shows "isometry (A o\<^sub>C\<^sub>L B)" proof- have "B* o\<^sub>C\<^sub>L A* o\<^sub>C\<^sub>L (A o\<^sub>C\<^sub>L B) = id_cblinfun" if "A* o\<^sub>C\<^sub>L A = id_cblinfun" and "B* o\<^sub>C\<^sub>L B = id_cblinfun" using that by (smt (verit, del_insts) adjoint_eqI cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply) thus ?thesis using assms unfolding isometry_def by simp qed lemma unitary_cblinfun_compose[simp]: "unitary (A o\<^sub>C\<^sub>L B)" if "unitary A" and "unitary B" using that by (smt (z3) adj_cblinfun_compose cblinfun_compose_assoc cblinfun_compose_id_right double_adj isometryD isometry_cblinfun_compose unitary_def unitary_isometry) lemma unitary_surj: assumes "unitary U" shows "surj (cblinfun_apply U)" apply (rule surjI[where f=\cblinfun_apply (U*)\]) using assms unfolding unitary_def apply transfer using comp_eq_dest_lhs by force lemma unitary_id[simp]: "unitary id_cblinfun" by (simp add: unitary_def) lemma orthogonal_on_basis_is_isometry: assumes spanB: \ccspan B = \\ assumes orthoU: \\b c. b\B \ c\B \ cinner (U *\<^sub>V b) (U *\<^sub>V c) = cinner b c\ shows \isometry U\ proof - have [simp]: \b \ closure (cspan B)\ for b using spanB apply transfer by simp have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ and \\\B\ for \ \ by (simp add: cinner_adj_left orthoU that(1) that(2)) have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ for \ \ apply (rule bounded_clinear_eq_on[where t=\ and G=B]) using bounded_clinear_cinner_right *[OF that] by auto have \U* *\<^sub>V U *\<^sub>V \ = \\ if \\\B\ for \ apply (rule cinner_extensionality) apply (subst cinner_eq_flip) by (simp add: * that) then have \U* o\<^sub>C\<^sub>L U = id_cblinfun\ by (metis cblinfun_apply_cblinfun_compose cblinfun_eq_gen_eqI cblinfun_id_cblinfun_apply spanB) then show \isometry U\ using isometry_def by blast qed lemma isometry_preserves_norm: \isometry U \ norm (U *\<^sub>V \) = norm \\ by (metis (no_types, lifting) cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply cinner_adj_right cnorm_eq isometryD) +lemma norm_isometry_compose: + assumes \isometry U\ + shows \norm (U o\<^sub>C\<^sub>L A) = norm A\ +proof - + have *: \norm (U *\<^sub>V A *\<^sub>V \) = norm (A *\<^sub>V \)\ for \ + by (smt (verit, ccfv_threshold) assms cblinfun_apply_cblinfun_compose cinner_adj_right cnorm_eq id_cblinfun_apply isometryD) + + have \norm (U o\<^sub>C\<^sub>L A) = (SUP \. norm (U *\<^sub>V A *\<^sub>V \) / norm \)\ + unfolding norm_cblinfun_Sup by auto + also have \\ = (SUP \. norm (A *\<^sub>V \) / norm \)\ + using * by auto + also have \\ = norm A\ + unfolding norm_cblinfun_Sup by auto + finally show ?thesis + by simp +qed + +lemma norm_isometry: + fixes U :: \'a::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ + assumes \isometry U\ + shows \norm U = 1\ + apply (subst asm_rl[of \U = U o\<^sub>C\<^sub>L id_cblinfun\], simp) + apply (subst norm_isometry_compose, simp add: assms) + by simp + +lemma norm_preserving_isometry: \isometry U\ if \\\. norm (U *\<^sub>V \) = norm \\ + by (smt (verit, ccfv_SIG) cblinfun_cinner_eqI cblinfun_id_cblinfun_apply cinner_adj_right cnorm_eq isometry_def simp_a_oCL_b' that) + subsection \Product spaces\ lift_definition cblinfun_left :: \'a::complex_normed_vector \\<^sub>C\<^sub>L ('a\'b::complex_normed_vector)\ is \(\x. (x,0))\ by (auto intro!: bounded_clinearI[where K=1]) lift_definition cblinfun_right :: \'b::complex_normed_vector \\<^sub>C\<^sub>L ('a::complex_normed_vector\'b)\ is \(\x. (0,x))\ by (auto intro!: bounded_clinearI[where K=1]) lemma isometry_cblinfun_left[simp]: \isometry cblinfun_left\ apply (rule orthogonal_on_basis_is_isometry[of some_chilbert_basis]) apply simp apply transfer by simp lemma isometry_cblinfun_right[simp]: \isometry cblinfun_right\ apply (rule orthogonal_on_basis_is_isometry[of some_chilbert_basis]) apply simp apply transfer by simp lemma cblinfun_left_right_ortho[simp]: \cblinfun_left* o\<^sub>C\<^sub>L cblinfun_right = 0\ proof - have \x \\<^sub>C ((cblinfun_left* o\<^sub>C\<^sub>L cblinfun_right) *\<^sub>V y) = 0\ for x :: 'b and y :: 'a apply (simp add: cinner_adj_right) apply transfer by auto then show ?thesis by (metis cblinfun.zero_left cblinfun_eqI cinner_eq_zero_iff) qed lemma cblinfun_right_left_ortho[simp]: \cblinfun_right* o\<^sub>C\<^sub>L cblinfun_left = 0\ proof - have \x \\<^sub>C ((cblinfun_right* o\<^sub>C\<^sub>L cblinfun_left) *\<^sub>V y) = 0\ for x :: 'b and y :: 'a apply (simp add: cinner_adj_right) apply transfer by auto then show ?thesis by (metis cblinfun.zero_left cblinfun_eqI cinner_eq_zero_iff) qed lemma cblinfun_left_apply[simp]: \cblinfun_left *\<^sub>V \ = (\,0)\ apply transfer by simp lemma cblinfun_left_adj_apply[simp]: \cblinfun_left* *\<^sub>V \ = fst \\ apply (cases \) by (auto intro!: cinner_extensionality[of \_ *\<^sub>V _\] simp: cinner_adj_right) lemma cblinfun_right_apply[simp]: \cblinfun_right *\<^sub>V \ = (0,\)\ apply transfer by simp lemma cblinfun_right_adj_apply[simp]: \cblinfun_right* *\<^sub>V \ = snd \\ apply (cases \) by (auto intro!: cinner_extensionality[of \_ *\<^sub>V _\] simp: cinner_adj_right) lift_definition ccsubspace_Times :: \'a::complex_normed_vector ccsubspace \ 'b::complex_normed_vector ccsubspace \ ('a\'b) ccsubspace\ is Product_Type.Times proof - fix S :: \'a set\ and T :: \'b set\ assume [simp]: \closed_csubspace S\ \closed_csubspace T\ have \csubspace (S \ T)\ by (simp add: complex_vector.subspace_Times) moreover have \closed (S \ T)\ by (simp add: closed_Times closed_csubspace.closed) ultimately show \closed_csubspace (S \ T)\ by (rule closed_csubspace.intro) qed lemma ccspan_Times: \ccspan (S \ T) = ccsubspace_Times (ccspan S) (ccspan T)\ if \0 \ S\ and \0 \ T\ proof (transfer fixing: S T) from that have \closure (cspan (S \ T)) = closure (cspan S \ cspan T)\ by (simp add: cspan_Times) also have \\ = closure (cspan S) \ closure (cspan T)\ using closure_Times by blast finally show \closure (cspan (S \ T)) = closure (cspan S) \ closure (cspan T)\ by - qed lemma ccspan_Times_sing1: \ccspan ({0::'a::complex_normed_vector} \ B) = ccsubspace_Times 0 (ccspan B)\ proof (transfer fixing: B) have \closure (cspan ({0::'a} \ B)) = closure ({0} \ cspan B)\ by (simp add: complex_vector.span_Times_sing1) also have \\ = closure {0} \ closure (cspan B)\ using closure_Times by blast also have \\ = {0} \ closure (cspan B)\ by simp finally show \closure (cspan ({0::'a} \ B)) = {0} \ closure (cspan B)\ by - qed lemma ccspan_Times_sing2: \ccspan (B \ {0::'a::complex_normed_vector}) = ccsubspace_Times (ccspan B) 0\ proof (transfer fixing: B) have \closure (cspan (B \ {0::'a})) = closure (cspan B \ {0})\ by (simp add: complex_vector.span_Times_sing2) also have \\ = closure (cspan B) \ closure {0}\ using closure_Times by blast also have \\ = closure (cspan B) \ {0}\ by simp finally show \closure (cspan (B \ {0::'a})) = closure (cspan B) \ {0}\ by - qed lemma ccsubspace_Times_sup: \sup (ccsubspace_Times A B) (ccsubspace_Times C D) = ccsubspace_Times (sup A C) (sup B D)\ proof transfer fix A C :: \'a set\ and B D :: \'b set\ have \A \ B +\<^sub>M C \ D = closure ((A \ B) + (C \ D))\ using closed_sum_def by blast also have \\ = closure ((A + C) \ (B + D))\ by (simp add: set_Times_plus_distrib) also have \\ = closure (A + C) \ closure (B + D)\ by (simp add: closure_Times) also have \\ = (A +\<^sub>M C) \ (B +\<^sub>M D)\ by (simp add: closed_sum_def) finally show \A \ B +\<^sub>M C \ D = (A +\<^sub>M C) \ (B +\<^sub>M D)\ by - qed lemma ccsubspace_Times_top_top[simp]: \ccsubspace_Times top top = top\ apply transfer by simp lemma is_onb_prod: assumes \is_onb B\ \is_onb B'\ shows \is_onb ((B \ {0}) \ ({0} \ B'))\ proof - from assms have 1: \is_ortho_set ((B \ {0}) \ ({0} \ B'))\ unfolding is_ortho_set_def apply (auto simp: is_onb_def is_ortho_set_def zero_prod_def) by (meson is_onb_def is_ortho_set_def)+ have 2: \(l, r) \ B \ {0} \ norm (l, r) = 1\ for l :: 'a and r :: 'b using \is_onb B\ is_onb_def by auto have 3: \(l, r) \ {0} \ B' \ norm (l, r) = 1\ for l :: 'a and r :: 'b using \is_onb B'\ is_onb_def by auto have [simp]: \ccspan B = top\ \ccspan B' = top\ using assms is_onb_def by auto have 4: \ccspan ((B \ {0}) \ ({0} \ B')) = top\ by (auto simp: ccspan_Times_sing1 ccspan_Times_sing2 ccsubspace_Times_sup simp flip: ccspan_union) from 1 2 3 4 show \is_onb ((B \ {0}) \ ({0} \ B'))\ by (auto simp add: is_onb_def) qed subsection \Images\ (* Closure is necessary. See email 47a3bb3d-3cc3-0934-36eb-3ef0f7b70a85@ut.ee *) lift_definition cblinfun_image :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a ccsubspace \ 'b ccsubspace\ (infixr "*\<^sub>S" 70) is "\A S. closure (A ` S)" using bounded_clinear_def closed_closure closed_csubspace.intro by (simp add: bounded_clinear_def complex_vector.linear_subspace_image closure_is_closed_csubspace) lemma cblinfun_image_mono: assumes a1: "S \ T" shows "A *\<^sub>S S \ A *\<^sub>S T" using a1 by (simp add: cblinfun_image.rep_eq closure_mono image_mono less_eq_ccsubspace.rep_eq) lemma cblinfun_image_0[simp]: shows "U *\<^sub>S 0 = 0" thm zero_ccsubspace_def apply transfer by (simp add: bounded_clinear_def complex_vector.linear_0) lemma cblinfun_image_bot[simp]: "U *\<^sub>S bot = bot" using cblinfun_image_0 by auto lemma cblinfun_image_sup[simp]: fixes A B :: \'a::chilbert_space ccsubspace\ and U :: "'a \\<^sub>C\<^sub>L'b::chilbert_space" shows \U *\<^sub>S (sup A B) = sup (U *\<^sub>S A) (U *\<^sub>S B)\ apply transfer using bounded_clinear.bounded_linear closure_image_closed_sum by blast lemma scaleC_cblinfun_image[simp]: fixes A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b :: chilbert_space\ and S :: \'a ccsubspace\ and \ :: complex shows \(\ *\<^sub>C A) *\<^sub>S S = \ *\<^sub>C (A *\<^sub>S S)\ proof- have \closure ( ( ((*\<^sub>C) \) \ (cblinfun_apply A) ) ` space_as_set S) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis closure_scaleC image_comp) hence \(closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis (mono_tags, lifting) comp_apply image_cong scaleC_cblinfun.rep_eq) hence \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by (metis space_as_set_inverse cblinfun_image.rep_eq scaleC_ccsubspace.rep_eq) have x1: "Abs_clinear_space (closure ((*\<^sub>V) (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure ((*\<^sub>V) A ` space_as_set S))" using \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by blast show ?thesis unfolding cblinfun_image_def using x1 by force qed lemma cblinfun_image_id[simp]: "id_cblinfun *\<^sub>S \ = \" apply transfer by (simp add: closed_csubspace.closed) lemma cblinfun_compose_image: \(A o\<^sub>C\<^sub>L B) *\<^sub>S S = A *\<^sub>S (B *\<^sub>S S)\ apply transfer unfolding image_comp[symmetric] apply (rule closure_bounded_linear_image_subset_eq[symmetric]) by (simp add: bounded_clinear.bounded_linear) lemmas cblinfun_assoc_left = cblinfun_compose_assoc[symmetric] cblinfun_compose_image[symmetric] add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space", symmetric] lemmas cblinfun_assoc_right = cblinfun_compose_assoc cblinfun_compose_image add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space"] lemma cblinfun_image_INF_leq[simp]: fixes U :: "'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::cbanach" and V :: "'a \ 'b ccsubspace" shows \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ apply transfer by (simp add: INT_greatest Inter_lower closure_mono image_mono) lemma isometry_cblinfun_image_inf_distrib': fixes U::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::cbanach\ and B C::"'a ccsubspace" shows "U *\<^sub>S (inf B C) \ inf (U *\<^sub>S B) (U *\<^sub>S C)" proof - define V where \V b = (if b then B else C)\ for b have \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ by auto then show ?thesis unfolding V_def by (metis (mono_tags, lifting) INF_UNIV_bool_expand) qed lemma cblinfun_image_eq: fixes S :: "'a::cbanach ccsubspace" and A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::cbanach" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and "ccspan G \ S" shows "A *\<^sub>S S = B *\<^sub>S S" proof (use assms in transfer) fix G :: "'a set" and A :: "'a \ 'b" and B :: "'a \ 'b" and S :: "'a set" assume a1: "bounded_clinear A" assume a2: "bounded_clinear B" assume a3: "\x. x \ G \ A x = B x" assume a4: "S \ closure (cspan G)" have "A ` closure S = B ` closure S" by (smt (verit, best) UnCI a1 a2 a3 a4 bounded_clinear_eq_on closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis bounded_clinear.bounded_linear a1 a2 closure_bounded_linear_image_subset_eq) qed lemma cblinfun_fixes_range: assumes "A o\<^sub>C\<^sub>L B = B" and "\ \ space_as_set (B *\<^sub>S top)" shows "A *\<^sub>V \ = \" proof- define rangeB rangeB' where "rangeB = space_as_set (B *\<^sub>S top)" and "rangeB' = range (cblinfun_apply B)" from assms have "\ \ closure rangeB'" by (simp add: cblinfun_image.rep_eq rangeB'_def top_ccsubspace.rep_eq) then obtain \i where \i_lim: "\i \ \" and \i_B: "\i i \ rangeB'" for i using closure_sequential by blast have A_invariant: "A *\<^sub>V \i i = \i i" for i proof- from \i_B obtain \ where \: "\i i = B *\<^sub>V \" using rangeB'_def by blast hence "A *\<^sub>V \i i = (A o\<^sub>C\<^sub>L B) *\<^sub>V \" by (simp add: cblinfun_compose.rep_eq) also have "\ = B *\<^sub>V \" by (simp add: assms) also have "\ = \i i" by (simp add: \) finally show ?thesis. qed from \i_lim have "(\i. A *\<^sub>V (\i i)) \ A *\<^sub>V \" by (rule isCont_tendsto_compose[rotated], simp) with A_invariant have "(\i. \i i) \ A *\<^sub>V \" by auto with \i_lim show "A *\<^sub>V \ = \" using LIMSEQ_unique by blast qed lemma zero_cblinfun_image[simp]: "0 *\<^sub>S S = (0::_ ccsubspace)" apply transfer by (simp add: complex_vector.subspace_0 image_constant[where x=0]) lemma cblinfun_image_INF_eq_general: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L'c::chilbert_space" and Uinv :: "'c \\<^sub>C\<^sub>L'b" assumes UinvUUinv: "Uinv o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L Uinv = Uinv" and UUinvU: "U o\<^sub>C\<^sub>L Uinv o\<^sub>C\<^sub>L U = U" \ \Meaning: \<^term>\Uinv\ is a Pseudoinverse of \<^term>\U\\ and V: "\i. V i \ Uinv *\<^sub>S top" shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof (rule antisym) show "U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S V i)" by (rule cblinfun_image_INF_leq) next define rangeU rangeUinv where "rangeU = U *\<^sub>S top" and "rangeUinv = Uinv *\<^sub>S top" define INFUV INFV where INFUV_def: "INFUV = (INF i. U *\<^sub>S V i)" and INFV_def: "INFV = (INF i. V i)" from assms have "V i \ rangeUinv" for i unfolding rangeUinv_def by simp moreover have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set rangeUinv" for \ using UinvUUinv cblinfun_fixes_range rangeUinv_def that by fastforce ultimately have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set (V i)" for \ i using less_eq_ccsubspace.rep_eq that by blast hence d1: "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>S (V i) = (V i)" for i proof transfer show "closure ((Uinv \ U) ` V i) = V i" if "pred_fun \ closed_csubspace V" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ (Uinv \ U) \ = \" for V :: "'a \ 'b set" and Uinv :: "'c \ 'b" and U :: "'b \ 'c" and i :: 'a using that proof auto show "x \ V i" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ closure (V i)" for x :: 'b using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure (V i)" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ V i" for x :: 'b using that using setdist_eq_0_sing_1 setdist_sing_in_set by blast qed qed have "U *\<^sub>S V i \ rangeU" for i by (simp add: cblinfun_image_mono rangeU_def) hence "INFUV \ rangeU" unfolding INFUV_def by (meson INF_lower UNIV_I order_trans) moreover have "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set rangeU" for \ using UUinvU cblinfun_fixes_range rangeU_def that by fastforce ultimately have x: "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set INFUV" for \ by (simp add: in_mono less_eq_ccsubspace.rep_eq that) have "closure ((U \ Uinv) ` INFUV) = INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ (U \ Uinv) \ = \" for INFUV :: "'c set" and U :: "'b \ 'c" and Uinv :: "'c \ 'b" using that proof auto show "x \ INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ closure INFUV" for x :: 'c using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ INFUV" for x :: 'c using that using setdist_eq_0_sing_1 setdist_sing_in_set by (simp add: closed_csubspace.closed) qed hence "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>S INFUV = INFUV" by (metis (mono_tags, opaque_lifting) x cblinfun_image.rep_eq cblinfun_image_id id_cblinfun_apply image_cong space_as_set_inject) hence "INFUV = U *\<^sub>S Uinv *\<^sub>S INFUV" by (simp add: cblinfun_compose_image) also have "\ \ U *\<^sub>S (INF i. Uinv *\<^sub>S U *\<^sub>S V i)" unfolding INFUV_def by (metis cblinfun_image_mono cblinfun_image_INF_leq) also have "\ = U *\<^sub>S INFV" using d1 by (metis (no_types, lifting) INFV_def cblinfun_assoc_left(2) image_cong) finally show "INFUV \ U *\<^sub>S INFV". qed lemma unitary_range[simp]: assumes "unitary U" shows "U *\<^sub>S top = top" using assms unfolding unitary_def apply transfer by (metis closure_UNIV comp_apply surj_def) lemma range_adjoint_isometry: assumes "isometry U" shows "U* *\<^sub>S top = top" proof- from assms have "top = U* *\<^sub>S U *\<^sub>S top" by (simp add: cblinfun_assoc_left(2)) also have "\ \ U* *\<^sub>S top" by (simp add: cblinfun_image_mono) finally show ?thesis using top.extremum_unique by blast qed lemma cblinfun_image_INF_eq[simp]: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L 'c::chilbert_space" assumes \isometry U\ shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof - from \isometry U\ have "U* o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L U* = U*" unfolding isometry_def by simp moreover from \isometry U\ have "U o\<^sub>C\<^sub>L U* o\<^sub>C\<^sub>L U = U" unfolding isometry_def by (simp add: cblinfun_compose_assoc) moreover have "V i \ U* *\<^sub>S top" for i by (simp add: range_adjoint_isometry assms) ultimately show ?thesis by (rule cblinfun_image_INF_eq_general) qed lemma isometry_cblinfun_image_inf_distrib[simp]: fixes U::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and X Y::"'a ccsubspace" assumes "isometry U" shows "U *\<^sub>S (inf X Y) = inf (U *\<^sub>S X) (U *\<^sub>S Y)" using cblinfun_image_INF_eq[where V="\b. if b then X else Y" and U=U] unfolding INF_UNIV_bool_expand using assms by auto lemma cblinfun_image_ccspan: shows "A *\<^sub>S ccspan G = ccspan ((*\<^sub>V) A ` G)" apply transfer by (simp add: bounded_clinear.bounded_linear bounded_clinear_def closure_bounded_linear_image_subset_eq complex_vector.linear_span_image) lemma cblinfun_apply_in_image[simp]: "A *\<^sub>V \ \ space_as_set (A *\<^sub>S \)" by (metis cblinfun_image.rep_eq closure_subset in_mono range_eqI top_ccsubspace.rep_eq) lemma cblinfun_plus_image_distr: \(A + B) *\<^sub>S S \ A *\<^sub>S S \ B *\<^sub>S S\ apply transfer by (smt (verit, ccfv_threshold) closed_closure closed_sum_def closure_minimal closure_subset image_subset_iff set_plus_intro subset_eq) lemma cblinfun_sum_image_distr: \(\i\I. A i) *\<^sub>S S \ (SUP i\I. A i *\<^sub>S S)\ proof (cases \finite I\) case True then show ?thesis proof induction case empty then show ?case by auto next case (insert x F) then show ?case apply auto by (smt (z3) cblinfun_plus_image_distr inf_sup_aci(6) le_iff_sup) qed next case False then show ?thesis by auto qed lemma space_as_set_image_commute: assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ and VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ (* I think only one of them is needed, can the lemma be strengthened? *) shows \(*\<^sub>V) U ` space_as_set T = space_as_set (U *\<^sub>S T)\ proof - have \space_as_set (U *\<^sub>S T) = U ` V ` space_as_set (U *\<^sub>S T)\ by (simp add: image_image UV flip: cblinfun_apply_cblinfun_compose) also have \\ \ U ` space_as_set (V *\<^sub>S U *\<^sub>S T)\ by (metis cblinfun_image.rep_eq closure_subset image_mono) also have \\ = U ` space_as_set T\ by (simp add: VU cblinfun_assoc_left(2)) finally have 1: \space_as_set (U *\<^sub>S T) \ U ` space_as_set T\ by - have 2: \U ` space_as_set T \ space_as_set (U *\<^sub>S T)\ by (simp add: cblinfun_image.rep_eq closure_subset) from 1 2 show ?thesis by simp qed lemma right_total_rel_ccsubspace: fixes R :: \'a::complex_normed_vector \ 'b::complex_normed_vector \ bool\ assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ assumes VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ assumes R_def: \\x y. R x y \ x = U *\<^sub>V y\ shows \right_total (rel_ccsubspace R)\ proof (rule right_totalI) fix T :: \'b ccsubspace\ show \\S. rel_ccsubspace R S T\ apply (rule exI[of _ \U *\<^sub>S T\]) using UV VU by (auto simp add: rel_ccsubspace_def R_def rel_set_def simp flip: space_as_set_image_commute) qed lemma left_total_rel_ccsubspace: fixes R :: \'a::complex_normed_vector \ 'b::complex_normed_vector \ bool\ assumes UV: \U o\<^sub>C\<^sub>L V = id_cblinfun\ assumes VU: \V o\<^sub>C\<^sub>L U = id_cblinfun\ assumes R_def: \\x y. R x y \ y = U *\<^sub>V x\ shows \left_total (rel_ccsubspace R)\ proof - have \right_total (rel_ccsubspace (conversep R))\ apply (rule right_total_rel_ccsubspace) using assms by auto then show ?thesis by (auto intro!: right_total_conversep[THEN iffD1] simp: converse_rel_ccsubspace) qed lemma cblinfun_image_bot_zero[simp]: \A *\<^sub>S top = bot \ A = 0\ by (metis Complex_Bounded_Linear_Function.zero_cblinfun_image bot_ccsubspace.rep_eq cblinfun_apply_in_image cblinfun_eqI empty_iff insert_iff zero_ccsubspace_def) subsection \Sandwiches\ lift_definition sandwich :: \('a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner) \ (('a \\<^sub>C\<^sub>L 'a) \\<^sub>C\<^sub>L ('b \\<^sub>C\<^sub>L 'b))\ is \\(A::'a\\<^sub>C\<^sub>L'b) B. A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ proof fix A :: \'a \\<^sub>C\<^sub>L 'b\ and B B1 B2 :: \'a \\<^sub>C\<^sub>L 'a\ and c :: complex show \A o\<^sub>C\<^sub>L (B1 + B2) o\<^sub>C\<^sub>L A* = (A o\<^sub>C\<^sub>L B1 o\<^sub>C\<^sub>L A*) + (A o\<^sub>C\<^sub>L B2 o\<^sub>C\<^sub>L A*)\ by (simp add: cblinfun_compose_add_left cblinfun_compose_add_right) show \A o\<^sub>C\<^sub>L (c *\<^sub>C B) o\<^sub>C\<^sub>L A* = c *\<^sub>C (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*)\ by auto show \\K. \B. norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * K\ proof (rule exI[of _ \norm A * norm (A*)\], rule allI) fix B have \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by blast also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) finally show \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * (norm A * norm (A*))\ by (simp add: mult.assoc vector_space_over_itself.scale_left_commute) qed qed lemma sandwich_0[simp]: \sandwich 0 = 0\ by (simp add: cblinfun_eqI sandwich.rep_eq) lemma sandwich_apply: \sandwich A *\<^sub>V B = A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ apply (transfer fixing: A B) by auto - lemma norm_sandwich: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'a::{chilbert_space} \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof - have main: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'c::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'd::{complex_inner}\ proof (rule norm_cblinfun_eqI) show \(norm A)\<^sup>2 \ norm (sandwich A *\<^sub>V id_cblinfun) / norm (id_cblinfun :: 'c \\<^sub>C\<^sub>L _)\ apply (auto simp: sandwich_apply) by - fix B have \norm (sandwich A *\<^sub>V B) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by (auto simp: sandwich_apply simp del: norm_adj) also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) also have \\ \ (norm A)\<^sup>2 * norm B\ by (simp add: power2_eq_square mult.assoc vector_space_over_itself.scale_left_commute) finally show \norm (sandwich A *\<^sub>V B) \ (norm A)\<^sup>2 * norm B\ by - show \0 \ (norm A)\<^sup>2\ by auto qed show ?thesis proof (cases \class.not_singleton TYPE('a)\) case True show ?thesis apply (rule main[internalize_sort' 'c2]) apply standard[1] using True by simp next case False have \A = 0\ apply (rule cblinfun_from_CARD_1_0[internalize_sort' 'a]) apply (rule not_singleton_vs_CARD_1) apply (rule False) by standard then show ?thesis by simp qed qed lemma sandwich_apply_adj: \sandwich A (B*) = (sandwich A B)*\ by (simp add: cblinfun_assoc_left(1) sandwich_apply) lemma sandwich_id[simp]: "sandwich id_cblinfun = id_cblinfun" apply (rule cblinfun_eqI) by (auto simp: sandwich_apply) subsection \Projectors\ lift_definition Proj :: "('a::chilbert_space) ccsubspace \ 'a \\<^sub>C\<^sub>L'a" is \projection\ by (rule projection_bounded_clinear) lemma Proj_range[simp]: "Proj S *\<^sub>S top = S" proof transfer fix S :: \'a set\ assume \closed_csubspace S\ then have "closure (range (projection S)) \ S" by (metis closed_csubspace.closed closed_csubspace.subspace closure_closed complex_vector.subspace_0 csubspace_is_convex dual_order.eq_iff insert_absorb insert_not_empty projection_image) moreover have "S \ closure (range (projection S))" using \closed_csubspace S\ by (metis closed_csubspace_def closure_subset csubspace_is_convex equals0D projection_image subset_iff) ultimately show \closure (range (projection S)) = S\ by auto qed lemma adj_Proj: \(Proj M)* = Proj M\ apply transfer by (simp add: projection_cadjoint) lemma Proj_idempotent[simp]: \Proj M o\<^sub>C\<^sub>L Proj M = Proj M\ proof - have u1: \(cblinfun_apply (Proj M)) = projection (space_as_set M)\ apply transfer by blast have \closed_csubspace (space_as_set M)\ using space_as_set by auto hence u2: \(projection (space_as_set M))\(projection (space_as_set M)) = (projection (space_as_set M))\ using projection_idem by fastforce have \(cblinfun_apply (Proj M)) \ (cblinfun_apply (Proj M)) = cblinfun_apply (Proj M)\ using u1 u2 by simp hence \cblinfun_apply ((Proj M) o\<^sub>C\<^sub>L (Proj M)) = cblinfun_apply (Proj M)\ by (simp add: cblinfun_compose.rep_eq) thus ?thesis using cblinfun_apply_inject by auto qed (* Widen the type class *) 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_top[simp]: \Proj \ = id_cblinfun\ + by (metis Proj_idempotent Proj_range cblinfun_eqI cblinfun_fixes_range id_cblinfun_apply iso_tuple_UNIV_I space_as_set_top) + lemma Proj_on_own_range': fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \P o\<^sub>C\<^sub>L P = P\ and \P = P*\ shows \Proj (P *\<^sub>S top) = P\ proof- define M where "M = P *\<^sub>S top" have v3: "x \ (\x. x - P *\<^sub>V x) -` {0}" if "x \ range (cblinfun_apply P)" for x :: 'a proof- have v3_1: \cblinfun_apply P \ cblinfun_apply P = cblinfun_apply P\ by (metis \P o\<^sub>C\<^sub>L P = P\ cblinfun_compose.rep_eq) have \\t. P *\<^sub>V t = x\ using that by blast then obtain t where t_def: \P *\<^sub>V t = x\ by blast hence \x - P *\<^sub>V x = x - P *\<^sub>V (P *\<^sub>V t)\ by simp also have \\ = x - (P *\<^sub>V t)\ using v3_1 by (metis comp_apply) also have \\ = 0\ by (simp add: t_def) finally have \x - P *\<^sub>V x = 0\ by blast thus ?thesis by simp qed have v1: "range (cblinfun_apply P) \ (\x. x - cblinfun_apply P x) -` {0}" using v3 by blast have "x \ range (cblinfun_apply P)" if "x \ (\x. x - P *\<^sub>V x) -` {0}" for x :: 'a proof- have x1:\x - P *\<^sub>V x = 0\ using that by blast have \x = P *\<^sub>V x\ by (simp add: x1 eq_iff_diff_eq_0) thus ?thesis by blast qed hence v2: "(\x. x - cblinfun_apply P x) -` {0} \ range (cblinfun_apply P)" by blast have i1: \range (cblinfun_apply P) = (\ x. x - cblinfun_apply P x) -` {0}\ using v1 v2 by (simp add: v1 dual_order.antisym) have p1: \closed {(0::'a)}\ by simp have p2: \continuous (at x) (\ x. x - P *\<^sub>V x)\ for x proof- have \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by (simp add: id_cblinfun.rep_eq minus_cblinfun.rep_eq) hence \bounded_clinear (cblinfun_apply (id_cblinfun - P))\ using cblinfun_apply by blast hence \continuous (at x) (cblinfun_apply (id_cblinfun - P))\ by (simp add: clinear_continuous_at) thus ?thesis using \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by simp qed have i2: \closed ( (\ x. x - P *\<^sub>V x) -` {0} )\ using p1 p2 by (rule Abstract_Topology.continuous_closed_vimage) have \closed (range (cblinfun_apply P))\ using i1 i2 by simp have u2: \cblinfun_apply P x \ space_as_set M\ for x by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq top_ccsubspace.rep_eq) have xy: \\ x - P *\<^sub>V x, y \ = 0\ if y1: \y \ space_as_set M\ for x y proof- have \\t. y = P *\<^sub>V t\ using y1 by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq image_iff top_ccsubspace.rep_eq) then obtain t where t_def: \y = P *\<^sub>V t\ by blast have \\ x - P *\<^sub>V x, y \ = \ x - P *\<^sub>V x, P *\<^sub>V t \\ by (simp add: t_def) also have \\ = \ P *\<^sub>V (x - P *\<^sub>V x), t \\ by (metis \P = P*\ cinner_adj_left) also have \\ = \ P *\<^sub>V x - P *\<^sub>V (P *\<^sub>V x), t \\ by (simp add: cblinfun.diff_right) also have \\ = \ P *\<^sub>V x - P *\<^sub>V x, t \\ by (metis assms(1) comp_apply cblinfun_compose.rep_eq) also have \\ = \ 0, t \\ by simp also have \\ = 0\ by simp finally show ?thesis by blast qed hence u1: \x - P *\<^sub>V x \ orthogonal_complement (space_as_set M)\ for x by (simp add: orthogonal_complementI) have "closed_csubspace (space_as_set M)" using space_as_set by auto hence f1: "(Proj M) *\<^sub>V a = P *\<^sub>V a" for a by (simp add: Proj.rep_eq projection_eqI u1 u2) have "(+) ((P - Proj M) *\<^sub>V a) = id" for a using f1 by (auto intro!: ext simp add: minus_cblinfun.rep_eq) hence "b - b = cblinfun_apply (P - Proj M) a" for a b by (metis (no_types) add_diff_cancel_right' id_apply) hence "cblinfun_apply (id_cblinfun - (P - Proj M)) a = a" for a by (simp add: minus_cblinfun.rep_eq) thus ?thesis using u1 u2 cblinfun_apply_inject diff_diff_eq2 diff_eq_diff_eq eq_id_iff id_cblinfun.rep_eq by (metis (no_types, opaque_lifting) M_def) qed lemma Proj_range_closed: assumes "is_Proj P" shows "closed (range (cblinfun_apply P))" using assms apply transfer using closed_csubspace.closed is_projection_on_image by blast lemma Proj_is_Proj[simp]: fixes M::\'a::chilbert_space ccsubspace\ shows \is_Proj (Proj M)\ proof- have u1: "closed_csubspace (space_as_set M)" using space_as_set by blast have v1: "h - Proj M *\<^sub>V h \ orthogonal_complement (space_as_set M)" for h by (simp add: Proj.rep_eq orthogonal_complementI projection_orthogonal u1) have v2: "Proj M *\<^sub>V h \ space_as_set M" for h by (metis Proj.rep_eq mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) have u2: "is_projection_on ((*\<^sub>V) (Proj M)) (space_as_set M)" unfolding is_projection_on_def by (simp add: smallest_dist_is_ortho u1 v1 v2) show ?thesis using u1 u2 is_Proj.rep_eq by blast qed lemma is_Proj_algebraic: fixes P::\'a::chilbert_space \\<^sub>C\<^sub>L 'a\ shows \is_Proj P \ P o\<^sub>C\<^sub>L P = P \ P = P*\ proof have "P o\<^sub>C\<^sub>L P = P" if "is_Proj P" using that apply transfer using is_projection_on_idem by fastforce moreover have "P = P*" if "is_Proj P" using that apply transfer by (metis is_projection_on_cadjoint) ultimately show "P o\<^sub>C\<^sub>L P = P \ P = P*" if "is_Proj P" using that by blast show "is_Proj P" if "P o\<^sub>C\<^sub>L P = P \ P = P*" using that Proj_on_own_range' Proj_is_Proj by metis qed lemma Proj_on_own_range: fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \is_Proj P\ shows \Proj (P *\<^sub>S top) = P\ using Proj_on_own_range' assms is_Proj_algebraic by blast lemma Proj_image_leq: "(Proj S) *\<^sub>S A \ S" by (metis Proj_range inf_top_left le_inf_iff isometry_cblinfun_image_inf_distrib') lemma Proj_sandwich: fixes A::"'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space" assumes "isometry A" shows "sandwich A *\<^sub>V Proj S = Proj (A *\<^sub>S S)" proof- define P where \P = A o\<^sub>C\<^sub>L Proj S o\<^sub>C\<^sub>L (A*)\ have \P o\<^sub>C\<^sub>L P = P\ using assms unfolding P_def isometry_def by (metis (no_types, lifting) Proj_idempotent cblinfun_assoc_left(1) cblinfun_compose_id_left) moreover have \P = P*\ unfolding P_def by (metis adj_Proj adj_cblinfun_compose cblinfun_assoc_left(1) double_adj) ultimately have \\M. P = Proj M \ space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ using P_def Proj_on_own_range' by (metis Proj_is_Proj Proj_range_closed cblinfun_image.rep_eq closure_closed top_ccsubspace.rep_eq) then obtain M where \P = Proj M\ and \space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ by blast have f1: "A o\<^sub>C\<^sub>L Proj S = P o\<^sub>C\<^sub>L A" by (simp add: P_def assms cblinfun_compose_assoc) hence "P o\<^sub>C\<^sub>L A o\<^sub>C\<^sub>L A* = P" using P_def by presburger hence "(P o\<^sub>C\<^sub>L A) *\<^sub>S (c \ A* *\<^sub>S d) = P *\<^sub>S (A *\<^sub>S c \ d)" for c d by (simp add: cblinfun_assoc_left(2)) hence "P *\<^sub>S (A *\<^sub>S \ \ c) = (P o\<^sub>C\<^sub>L A) *\<^sub>S \" for c by (metis sup_top_left) hence \M = A *\<^sub>S S\ using f1 by (metis \P = Proj M\ cblinfun_assoc_left(2) Proj_range sup_top_right) thus ?thesis using \P = Proj M\ unfolding P_def sandwich_apply by blast qed lemma Proj_orthog_ccspan_union: assumes "\x y. x \ X \ y \ Y \ is_orthogonal x y" shows \Proj (ccspan (X \ Y)) = Proj (ccspan X) + Proj (ccspan Y)\ proof - have \x \ cspan X \ y \ cspan Y \ is_orthogonal x y\ for x y apply (rule is_orthogonal_closure_cspan[where X=X and Y=Y]) using closure_subset assms by auto then have \x \ closure (cspan X) \ y \ closure (cspan Y) \ is_orthogonal x y\ for x y by (metis orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') then show ?thesis apply (transfer fixing: X Y) apply (subst projection_plus[symmetric]) by auto qed abbreviation proj :: "'a::chilbert_space \ 'a \\<^sub>C\<^sub>L 'a" where "proj \ \ Proj (ccspan {\})" lemma proj_0[simp]: \proj 0 = 0\ apply transfer by auto lemma surj_isometry_is_unitary: fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes \isometry U\ assumes \U *\<^sub>S \ = \\ shows \unitary U\ by (metis Proj_sandwich sandwich_apply Proj_on_own_range' assms(1) assms(2) cblinfun_compose_id_right isometry_def unitary_def unitary_id unitary_range) lemma ccsubspace_supI_via_Proj: fixes A B C::"'a::chilbert_space ccsubspace" assumes a1: \Proj (- C) *\<^sub>S A \ B\ shows "A \ sup B C" proof- have x2: \x \ space_as_set B\ if "x \ closure ( (projection (orthogonal_complement (space_as_set C))) ` space_as_set A)" for x using that by (metis Proj.rep_eq cblinfun_image.rep_eq assms less_eq_ccsubspace.rep_eq subsetD uminus_ccsubspace.rep_eq) have q1: \x \ closure {\ + \ |\ \. \ \ space_as_set B \ \ \ space_as_set C}\ if \x \ space_as_set A\ for x proof- have p1: \closed_csubspace (space_as_set C)\ using space_as_set by auto hence \x = (projection (space_as_set C)) x + (projection (orthogonal_complement (space_as_set C))) x\ by simp hence \x = (projection (orthogonal_complement (space_as_set C))) x + (projection (space_as_set C)) x\ by (metis ordered_field_class.sign_simps(2)) moreover have \(projection (orthogonal_complement (space_as_set C))) x \ space_as_set B\ using x2 by (meson closure_subset image_subset_iff that) moreover have \(projection (space_as_set C)) x \ space_as_set C\ by (metis mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) ultimately show ?thesis using closure_subset by fastforce qed have x1: \x \ (space_as_set B +\<^sub>M space_as_set C)\ if "x \ space_as_set A" for x proof - have f1: "x \ closure {a + b |a b. a \ space_as_set B \ b \ space_as_set C}" by (simp add: q1 that) have "{a + b |a b. a \ space_as_set B \ b \ space_as_set C} = {a. \p. p \ space_as_set B \ (\q. q \ space_as_set C \ a = p + q)}" by blast hence "x \ closure {a. \b\space_as_set B. \c\space_as_set C. a = b + c}" using f1 by (simp add: Bex_def_raw) thus ?thesis using that unfolding closed_sum_def set_plus_def by blast qed hence \x \ space_as_set (Abs_clinear_space (space_as_set B +\<^sub>M space_as_set C))\ if "x \ space_as_set A" for x using that by (metis space_as_set_inverse sup_ccsubspace.rep_eq) thus ?thesis by (simp add: x1 less_eq_ccsubspace.rep_eq subset_eq sup_ccsubspace.rep_eq) qed lemma is_Proj_idempotent: assumes "is_Proj P" shows "P o\<^sub>C\<^sub>L P = P" using assms unfolding is_Proj_def using assms is_Proj_algebraic by auto lemma is_proj_selfadj: assumes "is_Proj P" shows "P* = P" using assms unfolding is_Proj_def by (metis is_Proj_algebraic is_Proj_def) lemma is_Proj_I: assumes "P o\<^sub>C\<^sub>L P = P" and "P* = P" shows "is_Proj P" using assms is_Proj_algebraic by metis lemma is_Proj_0[simp]: "is_Proj 0" by (metis add_left_cancel adj_plus bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose group_cancel.rule0 is_Proj_I) lemma is_Proj_complement[simp]: assumes a1: "is_Proj P" shows "is_Proj (id_cblinfun-P)" by (smt (z3) add_diff_cancel_left add_diff_cancel_left' adj_cblinfun_compose adj_plus assms bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel id_cblinfun_adjoint is_Proj_algebraic cblinfun_compose_id_left) lemma Proj_bot[simp]: "Proj bot = 0" by (metis zero_cblinfun_image Proj_on_own_range' is_Proj_0 is_Proj_algebraic zero_ccsubspace_def) lemma Proj_ortho_compl: "Proj (- X) = id_cblinfun - Proj X" by (transfer , auto) lemma Proj_inj: assumes "Proj X = Proj Y" shows "X = Y" by (metis assms Proj_range) lemma norm_Proj_leq1: \norm (Proj M) \ 1\ apply transfer by (metis (no_types, opaque_lifting) mult.left_neutral onorm_bound projection_reduces_norm zero_less_one_class.zero_le_one) lemma Proj_orthog_ccspan_insert: assumes "\y. y \ Y \ is_orthogonal x y" shows \Proj (ccspan (insert x Y)) = proj x + Proj (ccspan Y)\ apply (subst asm_rl[of \insert x Y = {x} \ Y\], simp) apply (rule Proj_orthog_ccspan_union) using assms by auto lemma cancel_apply_Proj: assumes \\ \ space_as_set S\ shows \Proj S *\<^sub>V \ = \\ by (metis Proj_idempotent Proj_range assms cblinfun_fixes_range) lemma Proj_fixes_image: \Proj S *\<^sub>V \ = \\ if \\ \ space_as_set S\ by (simp add: Proj.rep_eq closed_csubspace_def projection_fixes_image that) lemma norm_is_Proj: \norm P \ 1\ if \is_Proj P\ by (metis Proj_on_own_range norm_Proj_leq1 that) +lemma Proj_sup: \orthogonal_spaces S T \ Proj (sup S T) = Proj S + Proj T\ + unfolding orthogonal_spaces_def + apply transfer + by (simp add: projection_plus) + +lemma Proj_sum_spaces: + assumes \finite X\ + assumes \\x y. x\X \ y\X \ x\y \ orthogonal_spaces (J x) (J y)\ + shows \Proj (\x\X. J x) = (\x\X. Proj (J x))\ + using assms +proof induction + case empty + show ?case + by auto +next + case (insert x F) + have \Proj (sum J (insert x F)) = Proj (J x \ sum J F)\ + by (simp add: insert) + also have \\ = Proj (J x) + Proj (sum J F)\ + apply (rule Proj_sup) + apply (rule orthogonal_sum) + using insert by auto + also have \\ = (\x\insert x F. Proj (J x))\ + by (simp add: insert.IH insert.hyps(1) insert.hyps(2) insert.prems) + finally show ?case + by - +qed + +lemma is_Proj_reduces_norm: + assumes \is_Proj P\ + shows \norm (P *\<^sub>V \) \ norm \\ + using assms apply transfer + using is_projection_on_reduces_norm by blast + +lemma norm_Proj_apply: \norm (Proj T *\<^sub>V \) = norm \ \ \ \ space_as_set T\ +proof (rule iffI) + show \norm (Proj T *\<^sub>V \) = norm \\ if \\ \ space_as_set T\ + by (simp add: cancel_apply_Proj that) + assume assm: \norm (Proj T *\<^sub>V \) = norm \\ + have \_decomp: \\ = Proj T \ + Proj (-T) \\ + by (simp add: Proj_ortho_compl cblinfun.real.diff_left) + have \(norm (Proj (-T) \))\<^sup>2 = (norm (Proj T \))\<^sup>2 + (norm (Proj (-T) \))\<^sup>2 - (norm (Proj T \))\<^sup>2\ + by auto + also have \\ = (norm (Proj T \ + Proj (-T) \))\<^sup>2 - (norm (Proj T \))\<^sup>2\ + apply (subst pythagorean_theorem) + apply (metis (no_types, lifting) Proj_idempotent \_decomp add_cancel_right_left adj_Proj cblinfun.real.add_right cblinfun_apply_cblinfun_compose cinner_adj_left cinner_zero_left) + by simp + also with \_decomp have \\ = (norm \)\<^sup>2 - (norm (Proj T \))\<^sup>2\ + by metis + also with assm have \\ = 0\ + by simp + finally have \norm (Proj (-T) \) = 0\ + by auto + with \_decomp have \\ = Proj T \\ + by auto + then show \\ \ space_as_set T\ + by (metis Proj_range cblinfun_apply_in_image) +qed + +lemma norm_Proj_apply_1: \norm \ = 1 \ norm (Proj T *\<^sub>V \) = 1 \ \ \ space_as_set T\ + using norm_Proj_apply by metis + subsection \Kernel\ lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector \ 'a ccsubspace" is "\ f. f -` {0}" by (metis kernel_is_closed_csubspace) definition eigenspace :: "complex \ 'a::complex_normed_vector \\<^sub>C\<^sub>L'a \ 'a ccsubspace" where "eigenspace a A = kernel (A - a *\<^sub>C id_cblinfun)" lemma kernel_scaleC[simp]: "a\0 \ kernel (a *\<^sub>C A) = kernel A" for a :: complex and A :: "(_,_) cblinfun" apply transfer using complex_vector.scale_eq_0_iff by blast lemma kernel_0[simp]: "kernel 0 = top" apply transfer by auto lemma kernel_id[simp]: "kernel id_cblinfun = 0" apply transfer by simp lemma eigenspace_scaleC[simp]: assumes a1: "a \ 0" shows "eigenspace b (a *\<^sub>C A) = eigenspace (b/a) A" proof - have "b *\<^sub>C (id_cblinfun::('a, _) cblinfun) = a *\<^sub>C (b / a) *\<^sub>C id_cblinfun" using a1 by (metis ceq_vector_fraction_iff) hence "kernel (a *\<^sub>C A - b *\<^sub>C id_cblinfun) = kernel (A - (b / a) *\<^sub>C id_cblinfun)" using a1 by (metis (no_types) complex_vector.scale_right_diff_distrib kernel_scaleC) thus ?thesis unfolding eigenspace_def by blast qed lemma eigenspace_memberD: assumes "x \ space_as_set (eigenspace e A)" shows "A *\<^sub>V x = e *\<^sub>C x" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms apply transfer by auto lemma eigenspace_memberI: assumes "A *\<^sub>V x = e *\<^sub>C x" shows "x \ space_as_set (eigenspace e A)" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms apply transfer by auto lemma kernel_Proj[simp]: \kernel (Proj S) = - S\ apply transfer apply auto apply (metis diff_0_right is_projection_on_iff_orthog projection_is_projection_on') by (simp add: complex_vector.subspace_0 projection_eqI) lemma orthogonal_projectors_orthogonal_spaces: \ \Logically belongs in section "Projectors".\ fixes S T :: \'a::chilbert_space set\ shows \(\x\S. \y\T. is_orthogonal x y) \ Proj (ccspan S) o\<^sub>C\<^sub>L Proj (ccspan T) = 0\ proof (intro ballI iffI) fix x y assume \Proj (ccspan S) o\<^sub>C\<^sub>L Proj (ccspan T) = 0\ \x \ S\ \y \ T\ then show \is_orthogonal x y\ by (smt (verit, del_insts) Proj_idempotent Proj_range adj_Proj cblinfun.zero_left cblinfun_apply_cblinfun_compose cblinfun_fixes_range ccspan_superset cinner_adj_right cinner_zero_right in_mono) next assume \\x\S. \y\T. is_orthogonal x y\ then have \ccspan S \ - ccspan T\ by (simp add: ccspan_leq_ortho_ccspan) then show \Proj (ccspan S) o\<^sub>C\<^sub>L Proj (ccspan T) = 0\ by (metis (no_types, opaque_lifting) Proj_range adj_Proj adj_cblinfun_compose basic_trans_rules(31) cblinfun.zero_left cblinfun_apply_cblinfun_compose cblinfun_apply_in_image cblinfun_eqI kernel_Proj kernel_memberD less_eq_ccsubspace.rep_eq) qed lemma cblinfun_compose_Proj_kernel[simp]: \a o\<^sub>C\<^sub>L Proj (kernel a) = 0\ apply (rule cblinfun_eqI) apply simp by (metis Proj_range cblinfun_apply_in_image kernel_memberD) lemma kernel_compl_adj_range: shows \kernel a = - (a* *\<^sub>S top)\ proof (rule ccsubspace_eqI) fix x have \x \ space_as_set (kernel a) \ a x = 0\ apply transfer by simp also have \a x = 0 \ (\y. is_orthogonal y (a x))\ by (metis cinner_gt_zero_iff cinner_zero_right) also have \\ \ (\y. is_orthogonal (a* *\<^sub>V y) x)\ by (simp add: cinner_adj_left) also have \\ \ x \ space_as_set (- (a* *\<^sub>S top))\ apply transfer by (metis (mono_tags, opaque_lifting) UNIV_I image_iff is_orthogonal_sym orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') finally show \x \ space_as_set (kernel a) \ x \ space_as_set (- (a* *\<^sub>S top))\ by - qed subsection \Partial isometries\ definition partial_isometry where \partial_isometry A \ (\h \ space_as_set (- kernel A). norm (A h) = norm h)\ lemma partial_isometryI: assumes \\h. h \ space_as_set (- kernel A) \ norm (A h) = norm h\ shows \partial_isometry A\ using assms partial_isometry_def by blast lemma fixes A :: \'a :: chilbert_space \\<^sub>C\<^sub>L 'b :: complex_normed_vector\ assumes iso: \\\. \ \ space_as_set V \ norm (A *\<^sub>V \) = norm \\ assumes zero: \\\. \ \ space_as_set (- V) \ A *\<^sub>V \ = 0\ shows partial_isometryI': \partial_isometry A\ and partial_isometry_initial: \kernel A = - V\ proof - from zero have \- V \ kernel A\ by (simp add: kernel_memberI less_eq_ccsubspace.rep_eq subsetI) moreover have \kernel A \ -V\ by (smt (verit, ccfv_threshold) Proj_ortho_compl Proj_range assms(1) cblinfun.diff_left cblinfun.diff_right cblinfun_apply_in_image cblinfun_id_cblinfun_apply ccsubspace_leI kernel_Proj kernel_memberD kernel_memberI norm_eq_zero ortho_involution subsetI zero) ultimately show kerA: \kernel A = -V\ by simp show \partial_isometry A\ apply (rule partial_isometryI) by (simp add: kerA iso) qed lemma Proj_partial_isometry: \partial_isometry (Proj S)\ apply (rule partial_isometryI) by (simp add: cancel_apply_Proj) lemma is_Proj_partial_isometry: \is_Proj P \ partial_isometry P\ by (metis Proj_on_own_range Proj_partial_isometry) lemma isometry_partial_isometry: \isometry P \ partial_isometry P\ by (simp add: isometry_preserves_norm partial_isometry_def) lemma unitary_partial_isometry: \unitary P \ partial_isometry P\ using isometry_partial_isometry unitary_isometry by blast lemma norm_partial_isometry: fixes A :: \'a :: chilbert_space \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \partial_isometry A\ and \A \ 0\ shows \norm A = 1\ proof - from \A \ 0\ have \- (kernel A) \ 0\ by (metis cblinfun_eqI diff_zero id_cblinfun_apply kernel_id kernel_memberD ortho_involution orthog_proj_exists orthogonal_complement_closed_subspace uminus_ccsubspace.rep_eq zero_cblinfun.rep_eq) then obtain h where \h \ space_as_set (- kernel A)\ and \h \ 0\ by (metis cblinfun_id_cblinfun_apply ccsubspace_eqI closed_csubspace.subspace complex_vector.subspace_0 kernel_id kernel_memberD kernel_memberI orthogonal_complement_closed_subspace uminus_ccsubspace.rep_eq) with \partial_isometry A\ have \norm (A h) = norm h\ using partial_isometry_def by blast then have \norm A \ 1\ by (smt (verit) \h \ 0\ mult_cancel_right1 mult_left_le_one_le norm_cblinfun norm_eq_zero norm_ge_zero) have \norm A \ 1\ proof (rule norm_cblinfun_bound) show \0 \ (1::real)\ by simp fix \ :: 'a define g h where \g = Proj (kernel A) \\ and \h = Proj (- kernel A) \\ have \A g = 0\ by (metis Proj_range cblinfun_apply_in_image g_def kernel_memberD) moreover from \partial_isometry A\ have \norm (A h) = norm h\ by (metis Proj_range cblinfun_apply_in_image h_def partial_isometry_def) ultimately have \norm (A \) = norm h\ by (simp add: Proj_ortho_compl cblinfun.diff_left cblinfun.diff_right g_def h_def) also have \norm h \ norm \\ by (smt (verit, del_insts) h_def mult_left_le_one_le norm_Proj_leq1 norm_cblinfun norm_ge_zero) finally show \norm (A *\<^sub>V \) \ 1 * norm \\ by simp qed from \norm A \ 1\ and \norm A \ 1\ show \norm A = 1\ by auto qed lemma partial_isometry_adj_a_o_a: assumes \partial_isometry a\ shows \a* o\<^sub>C\<^sub>L a = Proj (- kernel a)\ proof (rule cblinfun_cinner_eqI) define P where \P = Proj (- kernel a)\ have aP: \a o\<^sub>C\<^sub>L P = a\ by (auto intro!: simp: cblinfun_compose_minus_right P_def Proj_ortho_compl) have is_Proj_P[simp]: \is_Proj P\ by (simp add: P_def) fix \ :: 'a have \\ \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V \) = a \ \\<^sub>C a \\ by (simp add: cinner_adj_right) also have \\ = a (P \) \\<^sub>C a (P \)\ by (metis aP cblinfun_apply_cblinfun_compose) also have \\ = P \ \\<^sub>C P \\ by (metis P_def Proj_range assms cblinfun_apply_in_image cdot_square_norm partial_isometry_def) also have \\ = \ \\<^sub>C P \\ by (simp flip: cinner_adj_right add: is_proj_selfadj is_Proj_idempotent[THEN simp_a_oCL_b']) finally show \\ \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V \) = \ \\<^sub>C P \\ by - qed lemma partial_isometry_square_proj: \is_Proj (a* o\<^sub>C\<^sub>L a)\ if \partial_isometry a\ by (simp add: partial_isometry_adj_a_o_a that) lemma partial_isometry_adj[simp]: \partial_isometry (a*)\ if \partial_isometry a\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ proof - have ran_ker: \a *\<^sub>S top = - kernel (a*)\ by (simp add: kernel_compl_adj_range) have \norm (a* *\<^sub>V h) = norm h\ if \h \ range a\ for h proof - from that obtain x where h: \h = a x\ by auto have \norm (a* *\<^sub>V h) = norm (a* *\<^sub>V a *\<^sub>V x)\ by (simp add: h) also have \\ = norm (Proj (- kernel a) *\<^sub>V x)\ by (simp add: \partial_isometry a\ partial_isometry_adj_a_o_a simp_a_oCL_b') also have \\ = norm (a *\<^sub>V Proj (- kernel a) *\<^sub>V x)\ by (metis Proj_range \partial_isometry a\ cblinfun_apply_in_image partial_isometry_def) also have \\ = norm (a *\<^sub>V x)\ by (smt (verit, best) Proj_idempotent \partial_isometry a\ adj_Proj cblinfun_apply_cblinfun_compose cinner_adj_right cnorm_eq partial_isometry_adj_a_o_a) also have \\ = norm h\ using h by auto finally show ?thesis by - qed then have norm_pres: \norm (a* *\<^sub>V h) = norm h\ if \h \ closure (range a)\ for h using that apply (rule on_closure_eqI) apply assumption by (intro continuous_intros)+ show ?thesis apply (rule partial_isometryI) by (auto simp: cblinfun_image.rep_eq norm_pres simp flip: ran_ker) qed subsection \Isomorphisms and inverses\ definition iso_cblinfun :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ bool\ where \iso_cblinfun A = (\ B. A o\<^sub>C\<^sub>L B = id_cblinfun \ B o\<^sub>C\<^sub>L A = id_cblinfun)\ definition cblinfun_inv :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ ('b,'a) cblinfun\ where \cblinfun_inv A = (SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun)\ lemma assumes \iso_cblinfun A\ shows cblinfun_inv_left: \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and cblinfun_inv_right: \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ proof - from assms obtain B where AB: \A o\<^sub>C\<^sub>L B = id_cblinfun\ and BA: \B o\<^sub>C\<^sub>L A = id_cblinfun\ using iso_cblinfun_def by blast from BA have \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ by (metis (mono_tags, lifting) cblinfun_inv_def someI_ex) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB BA show \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ by auto qed lemma cblinfun_inv_uniq: assumes "A o\<^sub>C\<^sub>L B = id_cblinfun" and "B o\<^sub>C\<^sub>L A = id_cblinfun" shows "cblinfun_inv A = B" using assms by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) subsection \One-dimensional spaces\ instantiation cblinfun :: (one_dim, one_dim) complex_inner begin text \Once we have a theory for the trace, we could instead define the Hilbert-Schmidt inner product and relax the \<^class>\one_dim\-sort constraint to (\<^class>\cfinite_dim\,\<^class>\complex_normed_vector\) or similar\ definition "cinner_cblinfun (A::'a \\<^sub>C\<^sub>L 'b) (B::'a \\<^sub>C\<^sub>L 'b) = cnj (one_dim_iso (A *\<^sub>V 1)) * one_dim_iso (B *\<^sub>V 1)" instance proof intro_classes fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "\A, B\ = cnj \B, A\" unfolding cinner_cblinfun_def by auto show "\A + B, C\ = \A, C\ + \B, C\" by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) show "\c *\<^sub>C A, B\ = cnj c * \A, B\" by (simp add: cblinfun.scaleC_left cinner_cblinfun_def) show "0 \ \A, A\" unfolding cinner_cblinfun_def by auto have "bounded_clinear A \ A 1 = 0 \ A = (\_. 0)" for A::"'a \ 'b" proof (rule one_dim_clinear_eqI [where x = 1] , auto) show "clinear A" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: bounded_clinear.clinear) show "clinear ((\_. 0)::'a \ 'b)" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: complex_vector.module_hom_zero) qed hence "A *\<^sub>V 1 = 0 \ A = 0" by transfer hence "one_dim_iso (A *\<^sub>V 1) = 0 \ A = 0" by (metis one_dim_iso_of_zero one_dim_iso_inj) thus "(\A, A\ = 0) = (A = 0)" by (auto simp: cinner_cblinfun_def) show "norm A = sqrt (cmod \A, A\)" unfolding cinner_cblinfun_def apply transfer by (simp add: norm_mult abs_complex_def one_dim_onorm' cnj_x_x power2_eq_square bounded_clinear.clinear) qed end instantiation cblinfun :: (one_dim, one_dim) one_dim begin lift_definition one_cblinfun :: "'a \\<^sub>C\<^sub>L 'b" is "one_dim_iso" by (rule bounded_clinear_one_dim_iso) lift_definition times_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f g. f o one_dim_iso o g" by (simp add: comp_bounded_clinear) lift_definition inverse_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f. ((*) (one_dim_iso (inverse (f 1)))) o one_dim_iso" by (auto intro!: comp_bounded_clinear bounded_clinear_mult_right) definition divide_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" where "divide_cblinfun A B = A * inverse B" definition "canonical_basis_cblinfun = [1 :: 'a \\<^sub>C\<^sub>L 'b]" instance proof intro_classes let ?basis = "canonical_basis :: ('a \\<^sub>C\<^sub>L 'b) list" fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "distinct ?basis" unfolding canonical_basis_cblinfun_def by simp have "(1::'a \\<^sub>C\<^sub>L 'b) \ (0::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_one zero_neq_one) thus "cindependent (set ?basis)" unfolding canonical_basis_cblinfun_def by simp have "A \ cspan (set ?basis)" for A proof - define c :: complex where "c = one_dim_iso (A *\<^sub>V 1)" have "A x = one_dim_iso (A 1) *\<^sub>C one_dim_iso x" for x by (smt (z3) cblinfun.scaleC_right complex_vector.scale_left_commute one_dim_iso_idem one_dim_scaleC_1) hence "A = one_dim_iso (A *\<^sub>V 1) *\<^sub>C 1" apply transfer by metis thus "A \ cspan (set ?basis)" unfolding canonical_basis_cblinfun_def by (smt complex_vector.span_base complex_vector.span_scale list.set_intros(1)) qed thus "cspan (set ?basis) = UNIV" by auto have "A = (1::'a \\<^sub>C\<^sub>L 'b) \ norm (1::'a \\<^sub>C\<^sub>L 'b) = (1::real)" apply transfer by simp thus "A \ set ?basis \ norm A = 1" unfolding canonical_basis_cblinfun_def by simp show "?basis = [1]" unfolding canonical_basis_cblinfun_def by simp show "c *\<^sub>C 1 * c' *\<^sub>C 1 = (c * c') *\<^sub>C (1::'a\\<^sub>C\<^sub>L'b)" apply transfer by auto have "(1::'a \\<^sub>C\<^sub>L 'b) = (0::'a \\<^sub>C\<^sub>L 'b) \ False" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_zero' zero_neq_neg_one) thus "is_ortho_set (set ?basis)" unfolding is_ortho_set_def canonical_basis_cblinfun_def by auto show "A div B = A * inverse B" by (simp add: divide_cblinfun_def) show "inverse (c *\<^sub>C 1) = (1::'a\\<^sub>C\<^sub>L'b) /\<^sub>C c" apply transfer by (simp add: o_def one_dim_inverse) qed end lemma id_cblinfun_eq_1[simp]: \id_cblinfun = 1\ apply transfer by auto lemma one_dim_apply_is_times[simp]: fixes A :: "'a::one_dim \\<^sub>C\<^sub>L 'a" and B :: "'a \\<^sub>C\<^sub>L 'a" shows "A o\<^sub>C\<^sub>L B = A * B" apply transfer by simp lemma one_comp_one_cblinfun[simp]: "1 o\<^sub>C\<^sub>L 1 = 1" apply transfer unfolding o_def by simp lemma one_cblinfun_adj[simp]: "1* = 1" apply transfer by simp lemma scaleC_1_right[simp]: \scaleC x (1::'a::one_dim) = of_complex x\ unfolding of_complex_def by simp lemma scaleC_of_complex[simp]: \scaleC x (of_complex y) = of_complex (x * y)\ unfolding of_complex_def using scaleC_scaleC by blast lemma scaleC_1_apply[simp]: \(x *\<^sub>C 1) *\<^sub>V y = x *\<^sub>C y\ by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_apply_1_left[simp]: \1 *\<^sub>V y = y\ by (metis cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma of_complex_cblinfun_apply[simp]: \of_complex x *\<^sub>V y = x *\<^sub>C y\ unfolding of_complex_def by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_compose_1_left[simp]: \1 o\<^sub>C\<^sub>L x = x\ apply transfer by auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ apply transfer by auto lemma one_dim_iso_id_cblinfun: \one_dim_iso id_cblinfun = id_cblinfun\ by simp lemma one_dim_iso_id_cblinfun_eq_1: \one_dim_iso id_cblinfun = 1\ by simp lemma one_dim_iso_comp_distr[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a o\<^sub>C\<^sub>L one_dim_iso b\ by (smt (z3) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC) lemma one_dim_iso_comp_distr_times[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a * one_dim_iso b\ by (smt (verit, del_insts) mult.left_neutral mult_scaleC_left one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left) lemma one_dim_iso_adjoint[simp]: \one_dim_iso (A*) = (one_dim_iso A)*\ by (smt (z3) one_cblinfun_adj one_cinner_a_scaleC_one one_dim_iso_of_one one_dim_iso_scaleC scaleC_adj) lemma one_dim_iso_adjoint_complex[simp]: \one_dim_iso (A*) = cnj (one_dim_iso A)\ by (metis (mono_tags, lifting) one_cblinfun_adj one_dim_iso_idem one_dim_scaleC_1 scaleC_adj) lemma one_dim_cblinfun_compose_commute: \a o\<^sub>C\<^sub>L b = b o\<^sub>C\<^sub>L a\ for a b :: \('a::one_dim,'a) cblinfun\ by (simp add: one_dim_iso_inj) lemma one_cblinfun_apply_one[simp]: \1 *\<^sub>V 1 = 1\ by (simp add: one_cblinfun.rep_eq) lemma ccspan_one_dim[simp]: \ccspan {x} = top\ if \x \ 0\ for x :: \_ :: one_dim\ proof - have \y \ cspan {x}\ for y using that by (metis complex_vector.span_base complex_vector.span_zero cspan_singleton_scaleC insertI1 one_dim_scaleC_1 scaleC_zero_left) then show ?thesis by (auto intro!: order.antisym ccsubspace_leI simp: top_ccsubspace.rep_eq ccspan.rep_eq) qed lemma is_onb_one_dim[simp]: \norm x = 1 \ is_onb {x}\ for x :: \_ :: one_dim\ by (auto simp: is_onb_def intro!: ccspan_one_dim) lemma one_dim_iso_cblinfun_comp: \one_dim_iso (a o\<^sub>C\<^sub>L b) = of_complex (cinner (a* *\<^sub>V 1) (b *\<^sub>V 1))\ for a :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::one_dim\ and b :: \'c::one_dim \\<^sub>C\<^sub>L 'a\ by (simp add: cinner_adj_left cinner_cblinfun_def one_dim_iso_def) subsection \Loewner order\ lift_definition heterogenous_cblinfun_id :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ is \if bounded_clinear (heterogenous_identity :: 'a::complex_normed_vector \ 'b::complex_normed_vector) then heterogenous_identity else (\_. 0)\ by auto lemma heterogenous_cblinfun_id_def'[simp]: "heterogenous_cblinfun_id = id_cblinfun" apply transfer by auto definition "heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'b::chilbert_space itself) \ unitary (heterogenous_cblinfun_id :: 'a \\<^sub>C\<^sub>L 'b) \ unitary (heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a)" lemma heterogenous_same_type_cblinfun[simp]: \heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'a::chilbert_space itself)\ unfolding heterogenous_same_type_cblinfun_def by auto instantiation cblinfun :: (chilbert_space, chilbert_space) ord begin definition less_eq_cblinfun :: \('a \\<^sub>C\<^sub>L 'b) \ ('a \\<^sub>C\<^sub>L 'b) \ bool\ where less_eq_cblinfun_def_heterogenous: \less_eq_cblinfun A B = (if heterogenous_same_type_cblinfun TYPE('a) TYPE('b) then \\::'b. cinner \ ((B-A) *\<^sub>V heterogenous_cblinfun_id *\<^sub>V \) \ 0 else (A=B))\ definition \less_cblinfun (A :: 'a \\<^sub>C\<^sub>L 'b) B \ A \ B \ \ B \ A\ instance.. end lemma less_eq_cblinfun_def: \A \ B \ (\\. cinner \ (A *\<^sub>V \) \ cinner \ (B *\<^sub>V \))\ unfolding less_eq_cblinfun_def_heterogenous by (auto simp del: less_eq_complex_def simp: cblinfun.diff_left cinner_diff_right) instantiation cblinfun :: (chilbert_space, chilbert_space) ordered_complex_vector begin instance proof intro_classes note less_eq_complex_def[simp del] fix x y z :: \'a \\<^sub>C\<^sub>L 'b\ fix a b :: complex define pos where \pos X \ (\\. cinner \ (X *\<^sub>V \) \ 0)\ for X :: \'b \\<^sub>C\<^sub>L 'b\ consider (unitary) \heterogenous_same_type_cblinfun TYPE('a) TYPE('b)\ \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B = pos ((B-A) o\<^sub>C\<^sub>L (heterogenous_cblinfun_id :: 'b\\<^sub>C\<^sub>L'a))\ | (trivial) \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B \ A = B\ apply atomize_elim by (auto simp: pos_def less_eq_cblinfun_def_heterogenous) note cases = this have [simp]: \pos 0\ unfolding pos_def by auto have pos_nondeg: \X = 0\ if \pos X\ and \pos (-X)\ for X apply (rule cblinfun_cinner_eqI, simp) using that by (metis (no_types, lifting) cblinfun.minus_left cinner_minus_right dual_order.antisym equation_minus_iff neg_le_0_iff_le pos_def) have pos_add: \pos (X+Y)\ if \pos X\ and \pos Y\ for X Y by (smt (z3) pos_def cblinfun.diff_left cinner_minus_right cinner_simps(3) diff_ge_0_iff_ge diff_minus_eq_add neg_le_0_iff_le order_trans that(1) that(2) uminus_cblinfun.rep_eq) have pos_scaleC: \pos (a *\<^sub>C X)\ if \a\0\ and \pos X\ for X a using that unfolding pos_def by (auto simp: cblinfun.scaleC_left) let ?id = \heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a\ show \x \ x\ apply (cases rule:cases) by auto show \(x < y) \ (x \ y \ \ y \ x)\ unfolding less_cblinfun_def by simp show \x \ z\ if \x \ y\ and \y \ z\ proof (cases rule:cases) case unitary define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ and \b = (z-y) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ with unitary that have \pos a\ and \pos b\ by auto then have \pos (a + b)\ by (rule pos_add) moreover have \a + b = (z - x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ unfolding a_def b_def by (metis (no_types, lifting) bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel ordered_field_class.sign_simps(2) ordered_field_class.sign_simps(8)) ultimately show ?thesis using unitary by auto next case trivial with that show ?thesis by auto qed show \x = y\ if \x \ y\ and \y \ x\ proof (cases rule:cases) case unitary then have \unitary ?id\ by (auto simp: heterogenous_same_type_cblinfun_def) define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L ?id\ and \b = (x-y) o\<^sub>C\<^sub>L ?id\ with unitary that have \pos a\ and \pos b\ by auto then have \a = 0\ apply (rule_tac pos_nondeg) apply (auto simp: a_def b_def) by (smt (verit, best) add.commute bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose cblinfun_compose_zero_left diff_0 diff_add_cancel group_cancel.rule0 group_cancel.sub1) then show ?thesis unfolding a_def using \unitary ?id\ by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_compose_zero_left eq_iff_diff_eq_0 unitaryD2) next case trivial with that show ?thesis by simp qed show \x + y \ x + z\ if \y \ z\ proof (cases rule:cases) case unitary with that show ?thesis by auto next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ a *\<^sub>C y\ if \x \ y\ and \0 \ a\ proof (cases rule:cases) case unitary with that pos_scaleC show ?thesis by (metis cblinfun_compose_scaleC_left complex_vector.scale_right_diff_distrib) next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ b *\<^sub>C x\ if \a \ b\ and \0 \ x\ proof (cases rule:cases) case unitary with that show ?thesis by (auto intro!: pos_scaleC simp flip: scaleC_diff_left) next case trivial with that show ?thesis by auto qed qed end lemma positive_id_cblinfun[simp]: "id_cblinfun \ 0" unfolding less_eq_cblinfun_def using cinner_ge_zero by auto lemma positive_hermitianI: \A = A*\ if \A \ 0\ apply (rule cinner_real_hermiteanI) using that by (auto simp: complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma cblinfun_leI: assumes \\x. norm x = 1 \ x \\<^sub>C (A *\<^sub>V x) \ x \\<^sub>C (B *\<^sub>V x)\ shows \A \ B\ proof (unfold less_eq_cblinfun_def, intro allI, case_tac \\ = 0\) fix \ :: 'a assume \\ = 0\ then show \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ by simp next fix \ :: 'a assume \\ \ 0\ define \ where \\ = \ /\<^sub>R norm \\ have \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ apply (rule assms) unfolding \_def by (simp add: \\ \ 0\) with \\ \ 0\ show \\ \\<^sub>C (A *\<^sub>V \) \ \ \\<^sub>C (B *\<^sub>V \)\ unfolding \_def by (smt (verit) cinner_adj_left cinner_scaleR_left cinner_simps(6) complex_of_real_nn_iff mult_cancel_right1 mult_left_mono norm_eq_zero norm_ge_zero of_real_1 right_inverse scaleR_scaleC scaleR_scaleR) qed lemma positive_cblinfunI: \A \ 0\ if \\x. norm x = 1 \ cinner x (A *\<^sub>V x) \ 0\ apply (rule cblinfun_leI) using that by simp (* Note: this does not require B to be a square operator *) lemma positive_cblinfun_squareI: \A = B* o\<^sub>C\<^sub>L B \ A \ 0\ apply (rule positive_cblinfunI) by (metis cblinfun_apply_cblinfun_compose cinner_adj_right cinner_ge_zero) lemma one_dim_loewner_order: \A \ B \ one_dim_iso A \ (one_dim_iso B :: complex)\ for A B :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ proof - note less_eq_complex_def[simp del] have A: \A = one_dim_iso A *\<^sub>C id_cblinfun\ by simp have B: \B = one_dim_iso B *\<^sub>C id_cblinfun\ by simp have \A \ B \ (\\. cinner \ (A \) \ cinner \ (B \))\ by (simp add: less_eq_cblinfun_def) also have \\ \ (\\::'a. one_dim_iso B * (\ \\<^sub>C \) \ one_dim_iso A * (\ \\<^sub>C \))\ apply (subst A, subst B) by (metis (no_types, opaque_lifting) cinner_scaleC_right id_cblinfun_apply scaleC_cblinfun.rep_eq) also have \\ \ one_dim_iso A \ (one_dim_iso B :: complex)\ by (auto intro!: mult_right_mono elim!: allE[where x=1]) finally show ?thesis by - qed lemma one_dim_positive: \A \ 0 \ one_dim_iso A \ (0::complex)\ for A :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ using one_dim_loewner_order[where B=0] by auto lemma op_square_nondegenerate: \a = 0\ if \a* o\<^sub>C\<^sub>L a = 0\ proof (rule cblinfun_eq_0_on_UNIV_span[where basis=UNIV]; simp) fix s from that have \s \\<^sub>C ((a* o\<^sub>C\<^sub>L a) *\<^sub>V s) = 0\ by simp then have \(a *\<^sub>V s) \\<^sub>C (a *\<^sub>V s) = 0\ by (simp add: cinner_adj_right) then show \a *\<^sub>V s = 0\ by simp qed lemma comparable_hermitean: assumes \a \ b\ assumes \a* = a\ shows \b* = b\ by (smt (verit, best) assms(1) assms(2) cinner_hermitian_real cinner_real_hermiteanI comparable complex_is_real_iff_compare0 less_eq_cblinfun_def) lemma comparable_hermitean': assumes \a \ b\ assumes \b* = b\ shows \a* = a\ by (smt (verit, best) assms(1) assms(2) cinner_hermitian_real cinner_real_hermiteanI comparable complex_is_real_iff_compare0 less_eq_cblinfun_def) +lemma Proj_mono: \Proj S \ Proj T \ S \ T\ +proof (rule iffI) + assume \S \ T\ + define D where \D = Proj T - Proj S\ + from \S \ T\ have TS_S[simp]: \Proj T o\<^sub>C\<^sub>L Proj S = Proj S\ + by (smt (verit, ccfv_threshold) Proj_idempotent Proj_range cblinfun_apply_cblinfun_compose cblinfun_apply_in_image cblinfun_eqI cblinfun_fixes_range less_eq_ccsubspace.rep_eq subset_iff) + then have ST_S[simp]: \Proj S o\<^sub>C\<^sub>L Proj T = Proj S\ + by (metis adj_Proj adj_cblinfun_compose) + have \D* o\<^sub>C\<^sub>L D = D\ + by (simp add: D_def cblinfun_compose_minus_left cblinfun_compose_minus_right adj_minus adj_Proj) + then have \D \ 0\ + by (metis positive_cblinfun_squareI) + then show \Proj S \ Proj T\ + by (simp add: D_def) +next + assume PS_PT: \Proj S \ Proj T\ + show \S \ T\ + proof (rule ccsubspace_leI_unit) + fix \ assume \\ \ space_as_set S\ and [simp]: \norm \ = 1\ + then have \1 = norm (Proj S *\<^sub>V \)\ + by (simp add: cancel_apply_Proj) + also from PS_PT have \\ \ norm (Proj T *\<^sub>V \)\ + by (metis (no_types, lifting) Proj_idempotent adj_Proj cblinfun_apply_cblinfun_compose cinner_adj_left cnorm_le less_eq_cblinfun_def) + also have \\ \ 1\ + by (metis Proj_is_Proj \norm \ = 1\ is_Proj_reduces_norm) + ultimately have \norm (Proj T *\<^sub>V \) = 1\ + by auto + then show \\ \ space_as_set T\ + by (simp add: norm_Proj_apply_1) + qed +qed + subsection \Embedding vectors to operators\ lift_definition vector_to_cblinfun :: \'a::complex_normed_vector \ 'b::one_dim \\<^sub>C\<^sub>L 'a\ is \\\ \. one_dim_iso \ *\<^sub>C \\ by (simp add: bounded_clinear_scaleC_const) lemma vector_to_cblinfun_cblinfun_apply: "vector_to_cblinfun (A *\<^sub>V \) = A o\<^sub>C\<^sub>L (vector_to_cblinfun \)" apply transfer unfolding comp_def bounded_clinear_def clinear_def Vector_Spaces.linear_def module_hom_def module_hom_axioms_def by simp lemma vector_to_cblinfun_add: \vector_to_cblinfun (x + y) = vector_to_cblinfun x + vector_to_cblinfun y\ apply transfer by (simp add: scaleC_add_right) lemma norm_vector_to_cblinfun[simp]: "norm (vector_to_cblinfun x) = norm x" proof transfer have "bounded_clinear (one_dim_iso::'a \ complex)" by simp moreover have "onorm (one_dim_iso::'a \ complex) * norm x = norm x" for x :: 'b by simp ultimately show "onorm (\\. one_dim_iso (\::'a) *\<^sub>C x) = norm x" for x :: 'b by (subst onorm_scaleC_left) qed lemma bounded_clinear_vector_to_cblinfun[bounded_clinear]: "bounded_clinear vector_to_cblinfun" apply (rule bounded_clinearI[where K=1]) apply (transfer, simp add: scaleC_add_right) apply (transfer, simp add: mult.commute) by simp lemma vector_to_cblinfun_scaleC[simp]: "vector_to_cblinfun (a *\<^sub>C \) = a *\<^sub>C vector_to_cblinfun \" for a::complex proof (subst asm_rl [of "a *\<^sub>C \ = (a *\<^sub>C id_cblinfun) *\<^sub>V \"]) show "a *\<^sub>C \ = a *\<^sub>C id_cblinfun *\<^sub>V \" by (simp add: scaleC_cblinfun.rep_eq) show "vector_to_cblinfun (a *\<^sub>C id_cblinfun *\<^sub>V \) = a *\<^sub>C (vector_to_cblinfun \::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun_id_cblinfun_apply cblinfun_compose_scaleC_left vector_to_cblinfun_cblinfun_apply) qed lemma vector_to_cblinfun_apply_one_dim[simp]: shows "vector_to_cblinfun \ *\<^sub>V \ = one_dim_iso \ *\<^sub>C \" apply transfer by (rule refl) lemma vector_to_cblinfun_adj_apply[simp]: shows "vector_to_cblinfun \* *\<^sub>V \ = of_complex (cinner \ \)" by (simp add: cinner_adj_right one_dim_iso_def one_dim_iso_inj) lemma vector_to_cblinfun_comp_one[simp]: "(vector_to_cblinfun s :: 'a::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L 1 = (vector_to_cblinfun s :: 'b::one_dim \\<^sub>C\<^sub>L _)" apply (transfer fixing: s) by fastforce lemma vector_to_cblinfun_0[simp]: "vector_to_cblinfun 0 = 0" by (metis cblinfun.zero_left cblinfun_compose_zero_left vector_to_cblinfun_cblinfun_apply) lemma image_vector_to_cblinfun[simp]: "vector_to_cblinfun x *\<^sub>S top = ccspan {x}" proof transfer show "closure (range (\\::'b. one_dim_iso \ *\<^sub>C x)) = closure (cspan {x})" for x :: 'a proof (rule arg_cong [where f = closure]) have "k *\<^sub>C x \ range (\\. one_dim_iso \ *\<^sub>C x)" for k by (smt (z3) id_apply one_dim_iso_id one_dim_iso_idem range_eqI) thus "range (\\. one_dim_iso (\::'b) *\<^sub>C x) = cspan {x}" unfolding complex_vector.span_singleton by auto qed qed lemma vector_to_cblinfun_adj_comp_vector_to_cblinfun[simp]: shows "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \ = cinner \ \ *\<^sub>C id_cblinfun" proof - have "one_dim_iso \ *\<^sub>C one_dim_iso (of_complex \\, \\) = \\, \\ *\<^sub>C one_dim_iso \" for \ :: "'c::one_dim" by (metis complex_vector.scale_left_commute of_complex_def one_dim_iso_of_one one_dim_iso_scaleC one_dim_scaleC_1) hence "one_dim_iso ((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = one_dim_iso ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by simp hence "((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by (rule one_dim_iso_inj) thus ?thesis using cblinfun_eqI[where x = "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \" and y = "\\, \\ *\<^sub>C id_cblinfun"] by auto qed lemma isometry_vector_to_cblinfun[simp]: assumes "norm x = 1" shows "isometry (vector_to_cblinfun x)" using assms cnorm_eq_1 isometry_def by force subsection \Butterflies (rank-1 projectors)\ definition butterfly_def: "butterfly (s::'a::complex_normed_vector) (t::'b::chilbert_space) = vector_to_cblinfun s o\<^sub>C\<^sub>L (vector_to_cblinfun t :: complex \\<^sub>C\<^sub>L _)*" abbreviation "selfbutter s \ butterfly s s" lemma butterfly_add_left: \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (simp add: butterfly_def vector_to_cblinfun_add cbilinear_add_left bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma butterfly_add_right: \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (simp add: butterfly_def adj_plus vector_to_cblinfun_add cblinfun_compose_add_right) lemma butterfly_def_one_dim: "butterfly s t = (vector_to_cblinfun s :: 'c::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L (vector_to_cblinfun t :: 'c \\<^sub>C\<^sub>L _)*" (is "_ = ?rhs") for s :: "'a::complex_normed_vector" and t :: "'b::chilbert_space" proof - let ?isoAC = "1 :: 'c \\<^sub>C\<^sub>L complex" let ?isoCA = "1 :: complex \\<^sub>C\<^sub>L 'c" let ?vector = "vector_to_cblinfun :: _ \ ('c \\<^sub>C\<^sub>L _)" have "butterfly s t = (?vector s o\<^sub>C\<^sub>L ?isoCA) o\<^sub>C\<^sub>L (?vector t o\<^sub>C\<^sub>L ?isoCA)*" unfolding butterfly_def vector_to_cblinfun_comp_one by simp also have "\ = ?vector s o\<^sub>C\<^sub>L (?isoCA o\<^sub>C\<^sub>L ?isoCA*) o\<^sub>C\<^sub>L (?vector t)*" by (metis (no_types, lifting) cblinfun_compose_assoc adj_cblinfun_compose) also have "\ = ?rhs" by simp finally show ?thesis by simp qed lemma butterfly_comp_cblinfun: "butterfly \ \ o\<^sub>C\<^sub>L a = butterfly \ (a* *\<^sub>V \)" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma cblinfun_comp_butterfly: "a o\<^sub>C\<^sub>L butterfly \ \ = butterfly (a *\<^sub>V \) \" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma butterfly_apply[simp]: "butterfly \ \' *\<^sub>V \ = \\', \\ *\<^sub>C \" by (simp add: butterfly_def scaleC_cblinfun.rep_eq) lemma butterfly_scaleC_left[simp]: "butterfly (c *\<^sub>C \) \ = c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleC_right[simp]: "butterfly \ (c *\<^sub>C \) = cnj c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleR_left[simp]: "butterfly (r *\<^sub>R \) \ = r *\<^sub>C butterfly \ \" by (simp add: scaleR_scaleC) lemma butterfly_scaleR_right[simp]: "butterfly \ (r *\<^sub>R \) = r *\<^sub>C butterfly \ \" by (simp add: butterfly_scaleC_right scaleR_scaleC) lemma butterfly_adjoint[simp]: "(butterfly \ \)* = butterfly \ \" unfolding butterfly_def by auto lemma butterfly_comp_butterfly[simp]: "butterfly \1 \2 o\<^sub>C\<^sub>L butterfly \3 \4 = \\2, \3\ *\<^sub>C butterfly \1 \4" by (simp add: butterfly_comp_cblinfun) lemma butterfly_0_left[simp]: "butterfly 0 a = 0" by (simp add: butterfly_def) lemma butterfly_0_right[simp]: "butterfly a 0 = 0" by (simp add: butterfly_def) lemma norm_butterfly: "norm (butterfly \ \) = norm \ * norm \" proof (cases "\=0") case True then show ?thesis by simp next case False show ?thesis unfolding norm_cblinfun.rep_eq thm onormI[OF _ False] proof (rule onormI[OF _ False]) fix x have "cmod \\, x\ * norm \ \ norm \ * norm \ * norm x" by (metis ab_semigroup_mult_class.mult_ac(1) complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult_left_mono norm_ge_zero) thus "norm (butterfly \ \ *\<^sub>V x) \ norm \ * norm \ * norm x" by (simp add: power2_eq_square) show "norm (butterfly \ \ *\<^sub>V \) = norm \ * norm \ * norm \" by (smt (z3) ab_semigroup_mult_class.mult_ac(1) butterfly_apply mult.commute norm_eq_sqrt_cinner norm_ge_zero norm_scaleC power2_eq_square real_sqrt_abs real_sqrt_eq_iff) qed qed lemma bounded_sesquilinear_butterfly[bounded_sesquilinear]: \bounded_sesquilinear (\(b::'b::chilbert_space) (a::'a::chilbert_space). butterfly a b)\ proof standard fix a a' :: 'a and b b' :: 'b and r :: complex show \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (rule butterfly_add_left) show \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (rule butterfly_add_right) show \butterfly (r *\<^sub>C a) b = r *\<^sub>C butterfly a b\ by simp show \butterfly a (r *\<^sub>C b) = cnj r *\<^sub>C butterfly a b\ by simp show \\K. \b a. norm (butterfly a b) \ norm b * norm a * K \ apply (rule exI[of _ 1]) by (simp add: norm_butterfly) qed lemma inj_selfbutter_upto_phase: assumes "selfbutter x = selfbutter y" shows "\c. cmod c = 1 \ x = c *\<^sub>C y" proof (cases "x = 0") case True from assms have "y = 0" using norm_butterfly by (metis True butterfly_0_left divisors_zero norm_eq_zero) with True show ?thesis using norm_one by fastforce next case False define c where "c = \y, x\ / \x, x\" have "\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x" by (simp add: butterfly_apply) also have "\ = selfbutter y *\<^sub>V x" using assms by simp also have "\ = \y, x\ *\<^sub>C y" by (simp add: butterfly_apply) finally have xcy: "x = c *\<^sub>C y" by (simp add: c_def ceq_vector_fraction_iff) have "cmod c * norm x = cmod c * norm y" using assms norm_butterfly by (smt (verit, ccfv_SIG) \\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x\ \selfbutter y *\<^sub>V x = \y, x\ *\<^sub>C y\ cinner_scaleC_right complex_vector.scale_left_commute complex_vector.scale_right_imp_eq mult_cancel_left norm_eq_sqrt_cinner norm_eq_zero scaleC_scaleC xcy) also have "cmod c * norm y = norm (c *\<^sub>C y)" by simp also have "\ = norm x" unfolding xcy[symmetric] by simp finally have c: "cmod c = 1" by (simp add: False) from c xcy show ?thesis by auto qed lemma butterfly_eq_proj: assumes "norm x = 1" shows "selfbutter x = proj x" proof - define B and \ :: "complex \\<^sub>C\<^sub>L 'a" where "B = selfbutter x" and "\ = vector_to_cblinfun x" then have B: "B = \ o\<^sub>C\<^sub>L \*" unfolding butterfly_def by simp have \adj\: "\* o\<^sub>C\<^sub>L \ = id_cblinfun" using \_def assms isometry_def isometry_vector_to_cblinfun by blast have "B o\<^sub>C\<^sub>L B = \ o\<^sub>C\<^sub>L (\* o\<^sub>C\<^sub>L \) o\<^sub>C\<^sub>L \*" by (simp add: B cblinfun_assoc_left(1)) also have "\ = B" unfolding \adj\ by (simp add: B) finally have idem: "B o\<^sub>C\<^sub>L B = B". have herm: "B = B*" unfolding B by simp from idem herm have BProj: "B = Proj (B *\<^sub>S top)" by (rule Proj_on_own_range'[symmetric]) have "B *\<^sub>S top = ccspan {x}" by (simp add: B \_def assms cblinfun_compose_image range_adjoint_isometry) with BProj show "B = proj x" by simp qed lemma butterfly_sgn_eq_proj: shows "selfbutter (sgn x) = proj x" proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \selfbutter (sgn x) = proj (sgn x)\ by (simp add: butterfly_eq_proj norm_sgn) also have \ccspan {sgn x} = ccspan {x}\ by (metis ccspan_singleton_scaleC scaleC_eq_0_iff scaleR_scaleC sgn_div_norm sgn_zero_iff) finally show ?thesis by - qed lemma butterfly_is_Proj: \norm x = 1 \ is_Proj (selfbutter x)\ by (subst butterfly_eq_proj, simp_all) lemma cspan_butterfly_UNIV: assumes \cspan basisA = UNIV\ assumes \cspan basisB = UNIV\ assumes \is_ortho_set basisB\ assumes \\b. b \ basisB \ norm b = 1\ shows \cspan {butterfly a b| (a::'a::{complex_normed_vector}) (b::'b::{chilbert_space,cfinite_dim}). a \ basisA \ b \ basisB} = UNIV\ proof - have F: \\F\{butterfly a b |a b. a \ basisA \ b \ basisB}. \b'\basisB. F *\<^sub>V b' = (if b' = b then a else 0)\ if \a \ basisA\ and \b \ basisB\ for a b apply (rule bexI[where x=\butterfly a b\]) using assms that by (auto simp: is_ortho_set_def cnorm_eq_1) show ?thesis apply (rule cblinfun_cspan_UNIV[where basisA=basisB and basisB=basisA]) using assms apply auto[2] using F by (smt (verit, ccfv_SIG) image_iff) qed lemma cindependent_butterfly: fixes basisA :: \'a::chilbert_space set\ and basisB :: \'b::chilbert_space set\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes normA: \\a. a\basisA \ norm a = 1\ and normB: \\b. b\basisB \ norm b = 1\ shows \cindependent {butterfly a b| a b. a\basisA \ b\basisB}\ proof (unfold complex_vector.independent_explicit_module, intro allI impI, rename_tac T f g) fix T :: \('b \\<^sub>C\<^sub>L 'a) set\ and f :: \'b \\<^sub>C\<^sub>L 'a \ complex\ and g :: \'b \\<^sub>C\<^sub>L 'a\ assume \finite T\ assume T_subset: \T \ {butterfly a b |a b. a \ basisA \ b \ basisB}\ define lin where \lin = (\g\T. f g *\<^sub>C g)\ assume \lin = 0\ assume \g \ T\ (* To show: f g = 0 *) then obtain a b where g: \g = butterfly a b\ and [simp]: \a \ basisA\ \b \ basisB\ using T_subset by auto have *: "(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0" if \g \ T - {butterfly a b}\ for g proof - from that obtain a' b' where g: \g = butterfly a' b'\ and [simp]: \a' \ basisA\ \b' \ basisB\ using T_subset by auto from that have \g \ butterfly a b\ by auto with g consider (a) \a\a'\ | (b) \b\b'\ by auto then show \(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0\ proof cases case a then show ?thesis using \is_ortho_set basisA\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) next case b then show ?thesis using \is_ortho_set basisB\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) qed qed have \0 = (vector_to_cblinfun a)* *\<^sub>V lin *\<^sub>V b\ using \lin = 0\ by auto also have \\ = (\g\T. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ unfolding lin_def apply (rule complex_vector.linear_sum) by (smt (z3) cblinfun.scaleC_left cblinfun.scaleC_right cblinfun.add_right clinearI plus_cblinfun.rep_eq) also have \\ = (\g\{butterfly a b}. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ apply (rule sum.mono_neutral_right) using \finite T\ * \g \ T\ g by auto also have \\ = (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b\ by (simp add: g) also have \\ = f g\ unfolding g using normA normB by (auto simp: butterfly_def scaleC_cblinfun.rep_eq cnorm_eq_1) finally show \f g = 0\ by simp qed lemma clinear_eq_butterflyI: fixes F G :: \('a::{chilbert_space,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_inner) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes \cspan basisA = UNIV\ \cspan basisB = UNIV\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes "\a b. a\basisA \ b\basisB \ F (butterfly a b) = G (butterfly a b)" assumes \\b. b\basisB \ norm b = 1\ shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_UNIV) using assms by auto lemma sum_butterfly_is_Proj: assumes \is_ortho_set E\ assumes \\e. e\E \ norm e = 1\ shows \is_Proj (\e\E. butterfly e e)\ proof (cases \finite E\) case True show ?thesis proof (rule is_Proj_I) show \(\e\E. butterfly e e)* = (\e\E. butterfly e e)\ by (simp add: sum_adj) have ortho: \f \ e \ e \ E \ f \ E \ is_orthogonal f e\ for f e by (meson assms(1) is_ortho_set_def) have unit: \e \\<^sub>C e = 1\ if \e \ E\ for e using assms(2) cnorm_eq_1 that by blast have *: \(\f\E. (f \\<^sub>C e) *\<^sub>C butterfly f e) = butterfly e e\ if \e \ E\ for e apply (subst sum_single[where i=e]) by (auto intro!: simp: that ortho unit True) show \(\e\E. butterfly e e) o\<^sub>C\<^sub>L (\e\E. butterfly e e) = (\e\E. butterfly e e)\ by (auto simp: * cblinfun_compose_sum_right cblinfun_compose_sum_left) qed next case False then show ?thesis by simp qed subsection \Bifunctionals\ lift_definition bifunctional :: \'a::complex_normed_vector \\<^sub>C\<^sub>L (('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex)\ is \\x f. f *\<^sub>V x\ by (simp add: cblinfun.flip) lemma bifunctional_apply[simp]: \(bifunctional *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ by (transfer fixing: x f, simp) lemma bifunctional_isometric[simp]: \norm (bifunctional *\<^sub>V x) = norm x\ for x :: \'a::complex_inner\ proof - define f :: \'a \\<^sub>C\<^sub>L complex\ where \f = CBlinfun (\y. cinner x y)\ then have [simp]: \f *\<^sub>V y = cinner x y\ for y by (simp add: bounded_clinear_CBlinfun_apply bounded_clinear_cinner_right) then have [simp]: \norm f = norm x\ apply (auto intro!: norm_cblinfun_eqI[where x=x] simp: power2_norm_eq_cinner[symmetric]) apply (smt (verit, best) norm_eq_sqrt_cinner norm_ge_zero power2_norm_eq_cinner real_div_sqrt) using Cauchy_Schwarz_ineq2 by blast show ?thesis apply (auto intro!: norm_cblinfun_eqI[where x=f]) apply (metis norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) by (metis norm_cblinfun ordered_field_class.sign_simps(33)) qed lemma norm_bifunctional[simp]: \norm (bifunctional :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ proof - obtain x :: 'a where [simp]: \norm x = 1\ by (meson UNIV_not_singleton ex_norm1) show ?thesis by (auto intro!: norm_cblinfun_eqI[where x=x]) qed subsection \Banach-Steinhaus\ theorem cbanach_steinhaus: fixes F :: \'c \ 'a::cbanach \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\x. \M. \n. norm ((F n) *\<^sub>V x) \ M\ shows \\M. \ n. norm (F n) \ M\ using cblinfun_blinfun_transfer[transfer_rule] apply (rule TrueI)? (* Deletes current facts *) proof (use assms in transfer) fix F :: \'c \ 'a \\<^sub>L 'b\ assume \(\x. \M. \n. norm (F n *\<^sub>v x) \ M)\ hence \\x. bounded (range (\n. blinfun_apply (F n) x))\ by (metis (no_types, lifting) boundedI rangeE) hence \bounded (range F)\ by (simp add: banach_steinhaus) thus \\M. \n. norm (F n) \ M\ by (simp add: bounded_iff) qed subsection \Riesz-representation theorem\ theorem riesz_frechet_representation_cblinfun_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ shows \\t. \x. f *\<^sub>V x = \t, x\\ apply transfer by (rule riesz_frechet_representation_existence) lemma riesz_frechet_representation_cblinfun_unique: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::complex_inner \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ assumes \\x. f *\<^sub>V x = \u, x\\ shows \t = u\ using assms by (rule riesz_frechet_representation_unique) theorem riesz_frechet_representation_cblinfun_norm: includes notation_norm fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ shows \\f\ = \t\\ using assms proof transfer fix f::\'a \ complex\ and t assume \bounded_clinear f\ and \\x. f x = \t, x\\ from \\x. f x = \t, x\\ have \(norm (f x)) / (norm x) \ norm t\ for x proof(cases \norm x = 0\) case True thus ?thesis by simp next case False have \norm (f x) = norm (\t, x\)\ using \\x. f x = \t, x\\ by simp also have \norm \t, x\ \ norm t * norm x\ by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) finally have \norm (f x) \ norm t * norm x\ by blast thus ?thesis by (metis False linordered_field_class.divide_right_mono nonzero_mult_div_cancel_right norm_ge_zero) qed moreover have \(norm (f t)) / (norm t) = norm t\ proof(cases \norm t = 0\) case True thus ?thesis by simp next case False have \f t = \t, t\\ using \\x. f x = \t, x\\ by blast also have \\ = (norm t)^2\ by (meson cnorm_eq_square) also have \\ = (norm t)*(norm t)\ by (simp add: power2_eq_square) finally have \f t = (norm t)*(norm t)\ by blast thus ?thesis by (metis False Re_complex_of_real \\x. f x = cinner t x\ cinner_ge_zero complex_of_real_cmod nonzero_divide_eq_eq) qed ultimately have \Sup {(norm (f x)) / (norm x)| x. True} = norm t\ by (smt cSup_eq_maximum mem_Collect_eq) moreover have \Sup {(norm (f x)) / (norm x)| x. True} = (SUP x. (norm (f x)) / (norm x))\ by (simp add: full_SetCompr_eq) ultimately show \onorm f = norm t\ by (simp add: onorm_def) qed definition the_riesz_rep :: \('a::chilbert_space \\<^sub>C\<^sub>L complex) \ 'a\ where \the_riesz_rep f = (SOME t. \x. f x = \t, x\)\ lemma the_riesz_rep[simp]: \the_riesz_rep f \\<^sub>C x = f *\<^sub>V x\ unfolding the_riesz_rep_def apply (rule someI2_ex) by (simp_all add: riesz_frechet_representation_cblinfun_existence) lemma the_riesz_rep_unique: assumes \\x. f *\<^sub>V x = t \\<^sub>C x\ shows \t = the_riesz_rep f\ using assms riesz_frechet_representation_cblinfun_unique the_riesz_rep by metis lemma the_riesz_rep_scaleC: \the_riesz_rep (c *\<^sub>C f) = cnj c *\<^sub>C the_riesz_rep f\ apply (rule the_riesz_rep_unique[symmetric]) by auto lemma the_riesz_rep_add: \the_riesz_rep (f + g) = the_riesz_rep f + the_riesz_rep g\ apply (rule the_riesz_rep_unique[symmetric]) by (auto simp: cinner_add_left cblinfun.add_left) lemma the_riesz_rep_norm[simp]: \norm (the_riesz_rep f) = norm f\ apply (rule riesz_frechet_representation_cblinfun_norm[symmetric]) by simp lemma bounded_antilinear_the_riesz_rep[bounded_antilinear]: \bounded_antilinear the_riesz_rep\ by (metis (no_types, opaque_lifting) bounded_antilinear_intro dual_order.refl mult.commute mult_1 the_riesz_rep_add the_riesz_rep_norm the_riesz_rep_scaleC) definition the_riesz_rep_bilinear' :: \('a::complex_normed_vector \ 'b::chilbert_space \ complex) \ ('a \ 'b)\ where \the_riesz_rep_bilinear' p x = the_riesz_rep (CBlinfun (p x))\ lemma the_riesz_rep_bilinear'_correct: assumes \bounded_clinear (p x)\ shows \(the_riesz_rep_bilinear' p x) \\<^sub>C y = p x y\ by (auto simp add: the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply) lemma the_riesz_rep_bilinear'_plus1: assumes \\x. bounded_clinear (p x)\ and \\x. bounded_clinear (q x)\ shows \the_riesz_rep_bilinear' (p + q) = the_riesz_rep_bilinear' p + the_riesz_rep_bilinear' q\ by (auto intro!: ext simp add: the_riesz_rep_add CBlinfun_plus the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply) lemma the_riesz_rep_bilinear'_plus2: assumes \bounded_sesquilinear p\ shows \the_riesz_rep_bilinear' p (x + y) = the_riesz_rep_bilinear' p x + the_riesz_rep_bilinear' p y\ proof - have [simp]: \p (x + y) = p x + p y\ using assms bounded_sesquilinear.add_left by fastforce have [simp]: \bounded_clinear (p x)\ for x by (simp add: assms bounded_sesquilinear.bounded_clinear_right) show ?thesis by (auto intro!: ext simp add: the_riesz_rep_add CBlinfun_plus the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply) qed lemma the_riesz_rep_bilinear'_scaleC1: assumes \\x. bounded_clinear (p x)\ shows \the_riesz_rep_bilinear' (\x y. c *\<^sub>C p x y) = (\x. cnj c *\<^sub>C the_riesz_rep_bilinear' p x)\ by (auto intro!: ext simp add: the_riesz_rep_scaleC CBlinfun_scaleC the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply simp del: complex_scaleC_def scaleC_conv_of_complex) lemma the_riesz_rep_bilinear'_scaleC2: assumes \bounded_sesquilinear p\ shows \the_riesz_rep_bilinear' p (c *\<^sub>C x) = c *\<^sub>C the_riesz_rep_bilinear' p x\ proof - have [simp]: \p (c *\<^sub>C x) = (\y. cnj c *\<^sub>C p x y)\ using assms bounded_sesquilinear.scaleC_left by blast have [simp]: \bounded_clinear (p x)\ for x by (simp add: assms bounded_sesquilinear.bounded_clinear_right) show ?thesis by (auto intro!: ext simp add: the_riesz_rep_scaleC CBlinfun_scaleC the_riesz_rep_bilinear'_def assms bounded_clinear_CBlinfun_apply simp del: complex_scaleC_def scaleC_conv_of_complex) qed lemma bounded_clinear_the_riesz_rep_bilinear'2: assumes \bounded_sesquilinear p\ shows \bounded_clinear (the_riesz_rep_bilinear' p)\ proof - obtain K0 where K0: \cmod (p x y) \ norm x * norm y * K0\ for x y using assms bounded_sesquilinear.bounded by blast define K where \K = max K0 0\ with K0 have K: \cmod (p x y) \ norm x * norm y * K\ for x y by (smt (verit, del_insts) mult_nonneg_nonneg mult_nonneg_nonpos norm_ge_zero) have [simp]: \K \ 0\ by (simp add: K_def) have [simp]: \bounded_clinear (p x)\ for x by (simp add: assms bounded_sesquilinear.bounded_clinear_right) have "norm (the_riesz_rep_bilinear' p x) \ norm x * K" for x unfolding the_riesz_rep_bilinear'_def using K by (auto intro!: norm_cblinfun_bound simp: bounded_clinear_CBlinfun_apply mult.commute mult.left_commute) then show ?thesis apply (rule_tac bounded_clinearI) by (auto simp: assms the_riesz_rep_bilinear'_plus2 the_riesz_rep_bilinear'_scaleC2) qed lift_definition the_riesz_rep_bilinear:: \('a::complex_normed_vector \ 'b::chilbert_space \ complex) \ ('a \\<^sub>C\<^sub>L 'b)\ is \\p. if bounded_sesquilinear p then the_riesz_rep_bilinear' p else 0\ by (auto simp: bounded_clinear_the_riesz_rep_bilinear'2 zero_fun_def) lemma the_riesz_rep_bilinear_correct: assumes \bounded_sesquilinear p\ shows \(the_riesz_rep_bilinear p x) \\<^sub>C y = p x y\ apply (transfer fixing: p) by (simp add: assms bounded_sesquilinear.bounded_clinear_right the_riesz_rep_bilinear'_correct) subsection \Extension of complex bounded operators\ definition cblinfun_extension where "cblinfun_extension S \ = (SOME B. \x\S. B *\<^sub>V x = \ x)" definition cblinfun_extension_exists where "cblinfun_extension_exists S \ = (\B. \x\S. B *\<^sub>V x = \ x)" lemma cblinfun_extension_existsI: assumes "\x. x\S \ B *\<^sub>V x = \ x" shows "cblinfun_extension_exists S \" using assms cblinfun_extension_exists_def by blast lemma cblinfun_extension_exists_finite_dim: fixes \::"'a::{complex_normed_vector,cfinite_dim} \ 'b::complex_normed_vector" assumes "cindependent S" and "cspan S = UNIV" shows "cblinfun_extension_exists S \" proof- define f::"'a \ 'b" where "f = complex_vector.construct S \" have "clinear f" by (simp add: complex_vector.linear_construct assms linear_construct f_def) have "bounded_clinear f" using \clinear f\ assms by auto then obtain B::"'a \\<^sub>C\<^sub>L 'b" where "B *\<^sub>V x = f x" for x using cblinfun_apply_cases by blast have "B *\<^sub>V x = \ x" if c1: "x\S" for x proof- have "B *\<^sub>V x = f x" by (simp add: \\x. B *\<^sub>V x = f x\) also have "\ = \ x" using assms complex_vector.construct_basis f_def that by (simp add: complex_vector.construct_basis) finally show?thesis by blast qed thus ?thesis unfolding cblinfun_extension_exists_def by blast qed lemma cblinfun_extension_apply: assumes "cblinfun_extension_exists S f" and "v \ S" shows "(cblinfun_extension S f) *\<^sub>V v = f v" by (smt assms cblinfun_extension_def cblinfun_extension_exists_def tfl_some) lemma fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes \closure S = UNIV\ assumes f_add: \\x y. x \ S \ y \ S \ f (x + y) = f x + f y\ assumes f_scale: \\c x y. x \ S \ f (c *\<^sub>C x) = c *\<^sub>C f x\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows cblinfun_extension_exists_bounded_dense: \cblinfun_extension_exists S f\ and cblinfun_extension_exists_norm: \B \ 0 \ norm (cblinfun_extension S f) \ B\ proof - define B' where \B' = (if B\0 then 1 else B)\ then have bounded': \\x. x \ S \ norm (f x) \ B' * norm x\ using bounded by (metis mult_1 mult_le_0_iff norm_ge_zero order_trans) have \B' > 0\ by (simp add: B'_def) have \\xi. (xi \ x) \ (\i. xi i \ S)\ for x using assms(2) closure_sequential by blast then obtain seq :: \'a \ nat \ 'a\ where seq_lim: \seq x \ x\ and seq_S: \seq x i \ S\ for x i apply (atomize_elim, subst all_conj_distrib[symmetric]) apply (rule choice) by auto define g where \g x = lim (\i. f (seq x i))\ for x have \Cauchy (\i. f (seq x i))\ for x proof (rule CauchyI) fix e :: real assume \e > 0\ have \Cauchy (seq x)\ using LIMSEQ_imp_Cauchy seq_lim by blast then obtain M where less_eB: \norm (seq x m - seq x n) < e/B'\ if \n \ M\ and \m \ M\ for n m apply atomize_elim by (meson CauchyD \0 < B'\ \0 < e\ linordered_field_class.divide_pos_pos) have \norm (f (seq x m) - f (seq x n)) < e\ if \n \ M\ and \m \ M\ for n m proof - have \norm (f (seq x m) - f (seq x n)) = norm (f (seq x m - seq x n))\ using f_add f_scale seq_S by (metis add_diff_cancel assms(1) complex_vector.subspace_diff diff_add_cancel) also have \\ \ B' * norm (seq x m - seq x n)\ apply (rule bounded') by (simp add: assms(1) complex_vector.subspace_diff seq_S) also from less_eB have \\ < B' * (e/B')\ by (meson \0 < B'\ linordered_semiring_strict_class.mult_strict_left_mono that) also have \\ \ e\ using \0 < B'\ by auto finally show ?thesis by - qed then show \\M. \m\M. \n\M. norm (f (seq x m) - f (seq x n)) < e\ by auto qed then have f_seq_lim: \(\i. f (seq x i)) \ g x\ for x by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff g_def) have f_xi_lim: \(\i. f (xi i)) \ g x\ if \xi \ x\ and \\i. xi i \ S\ for xi x proof - from seq_lim that have \(\i. B' * norm (xi i - seq x i)) \ 0\ by (metis (no_types) \0 < B'\ cancel_comm_monoid_add_class.diff_cancel norm_not_less_zero norm_zero tendsto_diff tendsto_norm_zero_iff tendsto_zero_mult_left_iff) then have \(\i. f (xi i + (-1) *\<^sub>C seq x i)) \ 0\ apply (rule Lim_null_comparison[rotated]) using bounded' by (simp add: assms(1) complex_vector.subspace_diff seq_S that(2)) then have \(\i. f (xi i) - f (seq x i)) \ 0\ apply (subst (asm) f_add) apply (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S)[2] apply (subst (asm) f_scale) by (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S) then show \(\i. f (xi i)) \ g x\ using Lim_transform f_seq_lim by fastforce qed have g_add: \g (x + y) = g x + g y\ for x y proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto obtain yi :: \nat \ 'a\ where \yi \ y\ and \yi i \ S\ for i using seq_S seq_lim by auto have \(\i. xi i + yi i) \ x + y\ using \xi \ x\ \yi \ y\ tendsto_add by blast then have lim1: \(\i. f (xi i + yi i)) \ g (x + y)\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ assms(1) complex_vector.subspace_add f_xi_lim) have \(\i. f (xi i + yi i)) = (\i. f (xi i) + f (yi i))\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ f_add) also have \\ \ g x + g y\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ \xi \ x\ \yi \ y\ f_xi_lim tendsto_add) finally show ?thesis using lim1 LIMSEQ_unique by blast qed have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ for c x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto have \(\i. c *\<^sub>C xi i) \ c *\<^sub>C x\ using \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at isCont_tendsto_compose by blast then have lim1: \(\i. f (c *\<^sub>C xi i)) \ g (c *\<^sub>C x)\ by (simp add: \\i. xi i \ S\ assms(1) complex_vector.subspace_scale f_xi_lim) have \(\i. f (c *\<^sub>C xi i)) = (\i. c *\<^sub>C f (xi i))\ by (simp add: \\i. xi i \ S\ f_scale) also have \\ \ c *\<^sub>C g x\ using \\i. xi i \ S\ \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at f_xi_lim isCont_tendsto_compose by blast finally show ?thesis using lim1 LIMSEQ_unique by blast qed have [simp]: \f x = g x\ if \x \ S\ for x proof - have \(\_. x) \ x\ by auto then have \(\_. f x) \ g x\ using that by (rule f_xi_lim) then show \f x = g x\ by (simp add: LIMSEQ_const_iff) qed have g_bounded: \norm (g x) \ B' * norm x\ for x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto then have \(\i. f (xi i)) \ g x\ using f_xi_lim by presburger then have \(\i. norm (f (xi i))) \ norm (g x)\ by (metis tendsto_norm) moreover have \(\i. B' * norm (xi i)) \ B' * norm x\ by (simp add: \xi \ x\ tendsto_mult_left tendsto_norm) ultimately show \norm (g x) \ B' * norm x\ apply (rule lim_mono[rotated]) using bounded' using \xi _ \ S\ by blast qed have \bounded_clinear g\ using g_add g_scale apply (rule bounded_clinearI[where K=B']) using g_bounded by (simp add: ordered_field_class.sign_simps(5)) then have [simp]: \CBlinfun g *\<^sub>V x = g x\ for x by (subst CBlinfun_inverse, auto) show \cblinfun_extension_exists S f\ apply (rule cblinfun_extension_existsI[where B=\CBlinfun g\]) by auto then have \cblinfun_extension S f *\<^sub>V \ = CBlinfun g *\<^sub>V \\ if \\ \ S\ for \ by (simp add: cblinfun_extension_apply that) then have ext_is_g: \(*\<^sub>V) (cblinfun_extension S f) = g\ apply (rule_tac ext) apply (rule on_closure_eqI[where S=S]) using \closure S = UNIV\ \bounded_clinear g\ by (auto simp add: continuous_at_imp_continuous_on clinear_continuous_within) show \norm (cblinfun_extension S f) \ B\ if \B \ 0\ proof (cases \B > 0\) case True then have \B = B'\ unfolding B'_def by auto moreover have *: \norm (cblinfun_extension S f) \ B'\ by (metis ext_is_g \0 < B'\ g_bounded norm_cblinfun_bound order_le_less) ultimately show ?thesis by simp next case False with bounded have \f x = 0\ if \x \ S\ for x by (smt (verit) mult_nonpos_nonneg norm_ge_zero norm_le_zero_iff that) then have \g x = (\_. 0) x\ if \x \ S\ for x using that by simp then have \g x = 0\ for x apply (rule on_closure_eqI[where S=S]) using \closure S = UNIV\ \bounded_clinear g\ by (auto simp add: continuous_at_imp_continuous_on clinear_continuous_within) with ext_is_g have \cblinfun_extension S f = 0\ by (simp add: cblinfun_eqI) then show ?thesis using that by simp qed qed lemma cblinfun_extension_cong: assumes \cspan A = cspan B\ assumes \B \ A\ assumes fg: \\x. x\B \ f x = g x\ assumes \cblinfun_extension_exists A f\ shows \cblinfun_extension A f = cblinfun_extension B g\ proof - from \cblinfun_extension_exists A f\ fg \B \ A\ have \cblinfun_extension_exists B g\ by (metis assms(2) cblinfun_extension_exists_def subset_eq) have \(\x\A. C *\<^sub>V x = f x) \ (\x\B. C *\<^sub>V x = f x)\ for C by (smt (verit, ccfv_SIG) assms(1) assms(2) assms(4) cblinfun_eq_on_span cblinfun_extension_exists_def complex_vector.span_eq subset_iff) also from fg have \\ C \ (\x\B. C *\<^sub>V x = g x)\ for C by auto finally show \cblinfun_extension A f = cblinfun_extension B g\ unfolding cblinfun_extension_def by auto qed lemma fixes f :: \'a::complex_inner \ 'b::chilbert_space\ and S assumes \is_ortho_set S\ and \closure (cspan S) = UNIV\ assumes ortho_f: \\x y. x\S \ y\S \ x\y \ is_orthogonal (f x) (f y)\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows cblinfun_extension_exists_ortho: \cblinfun_extension_exists S f\ and cblinfun_extension_exists_ortho_norm: \B \ 0 \ norm (cblinfun_extension S f) \ B\ proof - define g where \g = cconstruct S f\ have \cindependent S\ using assms(1) is_ortho_set_cindependent by blast have g_f: \g x = f x\ if \x\S\ for x unfolding g_def using \cindependent S\ that by (rule complex_vector.construct_basis) have [simp]: \clinear g\ unfolding g_def using \cindependent S\ by (rule complex_vector.linear_construct) then have g_add: \g (x + y) = g x + g y\ if \x \ cspan S\ and \y \ cspan S\ for x y using clinear_iff by blast from \clinear g\ have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ if \x \ cspan S\ for x c by (simp add: complex_vector.linear_scale) moreover have g_bounded: \norm (g x) \ abs B * norm x\ if \x \ cspan S\ for x proof - from that obtain t r where x_sum: \x = (\a\t. r a *\<^sub>C a)\ and \finite t\ and \t \ S\ unfolding complex_vector.span_explicit by auto have \(norm (g x))\<^sup>2 = (norm (\a\t. r a *\<^sub>C g a))\<^sup>2\ by (simp add: x_sum complex_vector.linear_sum clinear.scaleC) also have \\ = (norm (\a\t. r a *\<^sub>C f a))\<^sup>2\ by (smt (verit) \t \ S\ g_f in_mono sum.cong) also have \\ = (\a\t. (norm (r a *\<^sub>C f a))\<^sup>2)\ using _ \finite t\ apply (rule pythagorean_theorem_sum) using \t \ S\ ortho_f in_mono by fastforce also have \\ = (\a\t. (cmod (r a) * norm (f a))\<^sup>2)\ by simp also have \\ \ (\a\t. (cmod (r a) * B * norm a)\<^sup>2)\ apply (rule sum_mono) by (metis \t \ S\ assms(4) in_mono mult_left_mono mult_nonneg_nonneg norm_ge_zero power_mono vector_space_over_itself.scale_scale) also have \\ = B\<^sup>2 * (\a\t. (norm (r a *\<^sub>C a))\<^sup>2)\ by (simp add: sum_distrib_left mult.commute vector_space_over_itself.scale_left_commute flip: power_mult_distrib) also have \\ = B\<^sup>2 * (norm (\a\t. (r a *\<^sub>C a)))\<^sup>2\ apply (subst pythagorean_theorem_sum) using \finite t\ apply auto by (meson \t \ S\ assms(1) is_ortho_set_def subsetD) also have \\ = (abs B * norm x)\<^sup>2\ by (simp add: power_mult_distrib x_sum) finally show \norm (g x) \ abs B * norm x\ by auto qed from g_add g_scale g_bounded have extg_exists: \cblinfun_extension_exists (cspan S) g\ apply (rule_tac cblinfun_extension_exists_bounded_dense[where B=\abs B\]) using \closure (cspan S) = UNIV\ by auto then show \cblinfun_extension_exists S f\ by (metis (mono_tags, opaque_lifting) g_f cblinfun_extension_apply cblinfun_extension_existsI complex_vector.span_base) have norm_extg: \norm (cblinfun_extension (cspan S) g) \ B\ if \B \ 0\ apply (rule cblinfun_extension_exists_norm) using g_add g_scale g_bounded \closure (cspan S) = UNIV\ that by auto have extg_extf: \cblinfun_extension (cspan S) g = cblinfun_extension S f\ apply (rule cblinfun_extension_cong) by (auto simp add: complex_vector.span_base g_f extg_exists) from norm_extg extg_extf show \norm (cblinfun_extension S f) \ B\ if \B \ 0\ using that by simp qed lemma cblinfun_extension_exists_proj: fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes \\P. is_projection_on P (closure S) \ bounded_clinear P\ (* Maybe can be replaced by is_Proj if the latter's type class is widened *) assumes 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 (cases \B \ 0\) case True note True[simp] obtain P where P_proj: \is_projection_on P (closure S)\ and P_blin[simp]: \bounded_clinear P\ using assms(2) by blast have P_lin[simp]: \clinear P\ by (simp add: bounded_clinear.clinear) define f' S' where \f' \ = f (P \)\ and \S' = S + (P -` {0})\ for \ have \csubspace S'\ by (simp add: S'_def assms(1) csubspace_set_plus) moreover have \closure S' = UNIV\ proof auto fix \ have \\ = P \ + (id - P) \\ by simp also have \\ \ closure S + (P -` {0})\ apply (rule set_plus_intro) using P_proj is_projection_on_in_image by (auto simp: complex_vector.linear_diff is_projection_on_fixes_image is_projection_on_in_image) also have \\ \ closure (closure S + (P -` {0}))\ using closure_subset by blast also have \\ = closure (S + (P -` {0}))\ using closed_sum_closure_left closed_sum_def by blast also have \\ = closure S'\ using S'_def by fastforce finally show \\ \ closure S'\ by - qed moreover have \f' (x + y) = f' x + f' y\ if \x \ S'\ and \y \ S'\ for x y by (smt (z3) P_blin P_proj S'_def f'_def add.right_neutral bounded_clinear_CBlinfun_apply cblinfun.add_right closure_subset f_add is_projection_on_fixes_image set_plus_elim singletonD subset_eq that(1) that(2) vimageE) moreover have \f' (c *\<^sub>C x) = c *\<^sub>C f' x\ if \x \ S'\ for c x by (smt (verit, ccfv_SIG) P_blin P_proj S'_def f'_def add.right_neutral bounded_clinear_CBlinfun_apply cblinfun.add_right cblinfun.scaleC_right closure_subset f_scale is_projection_on_fixes_image set_plus_elim singletonD subset_eq that vimageE) moreover from P_blin obtain B' where B': \norm (P x) \ B' * norm x\ for x by (metis bounded_clinear.bounded mult.commute) have \norm (f' x) \ (B * B') * norm x\ if \x \ S'\ for x proof - have \norm (f' x) \ B* norm (P x)\ apply (auto simp: f'_def) by (smt (verit) P_blin P_proj S'_def add.right_neutral bounded bounded_clinear_CBlinfun_apply cblinfun.add_right closure_subset is_projection_on_fixes_image set_plus_elim singletonD subset_eq that vimageE) also have \\ \ B * B' * norm x\ by (simp add: B' mult.assoc mult_mono) finally show ?thesis by auto qed ultimately have F_ex: \cblinfun_extension_exists S' f'\ by (rule cblinfun_extension_exists_bounded_dense) define F where \F = cblinfun_extension S' f'\ from F_ex have *: \F \ = f' \\ if \\ \ S'\ for \ by (simp add: F_def cblinfun_extension_apply that) then have \F \ = f \\ if \\ \ S\ for \ apply (auto simp: S'_def f'_def) by (metis (no_types, lifting) P_lin P_proj add.right_neutral closure_subset complex_vector.linear_subspace_vimage complex_vector.subspace_0 complex_vector.subspace_single_0 is_projection_on_fixes_image set_plus_intro subset_eq that) then show \cblinfun_extension_exists S f\ using cblinfun_extension_exists_def by blast next case False then have \S \ {0}\ using bounded apply auto by (meson norm_ge_zero norm_le_zero_iff order_trans zero_le_mult_iff) then show \cblinfun_extension_exists S f\ apply (rule_tac cblinfun_extension_existsI[where B=0]) apply auto using bounded by fastforce qed subsection \Bijections between different ONBs\ text \Some of the theorems here logically belong into \<^theory>\Complex_Bounded_Operators.Complex_Inner_Product\ but the proof uses some concepts from the present theory.\ (* TODO mention: follows conway functional, Prop 4.14 *) lemma all_ortho_bases_same_card: fixes E F :: \'a::chilbert_space set\ assumes \is_ortho_set E\ \is_ortho_set F\ \ccspan E = top\ \ccspan F = top\ shows \\f. bij_betw f E F\ proof - have \|F| \o |E|\ if \infinite E\ and \is_ortho_set E\ \is_ortho_set F\ \ccspan E = top\ \ccspan F = top\ for E F :: \'a set\ proof - define F' where \F' e = {f\F. \ is_orthogonal f e}\ for e have \\e\E. cinner f e \ 0\ if \f \ F\ for f proof (rule ccontr, simp) assume \\e\E. is_orthogonal f e\ then have \f \ orthogonal_complement E\ by (simp add: orthogonal_complementI) also have \\ = orthogonal_complement (closure (cspan E))\ using orthogonal_complement_of_closure orthogonal_complement_of_cspan by blast also have \\ = space_as_set (- ccspan E)\ apply transfer by simp also have \\ = space_as_set (- top)\ by (simp add: \ccspan E = top\) also have \\ = {0}\ by (auto simp add: top_ccsubspace.rep_eq uminus_ccsubspace.rep_eq) finally have \f = 0\ by simp with \f \ F\ \is_ortho_set F\ show False by (simp add: is_onb_def is_ortho_set_def) qed then have F'_union: \F = (\e\E. F' e)\ unfolding F'_def by auto have \countable (F' e)\ for e proof - have \(\f\M. (cmod (cinner (sgn f) e))\<^sup>2) \ (norm e)\<^sup>2\ if \finite M\ and \M \ F\ for M proof - have [simp]: \is_ortho_set M\ using \is_ortho_set F\ is_ortho_set_antimono that(2) by blast have [simp]: \norm (sgn f) = 1\ if \f \ M\ for f by (metis \is_ortho_set M\ is_ortho_set_def norm_sgn that) have \(\f\M. (cmod (cinner (sgn f) e))\<^sup>2) = (\f\M. (norm ((cinner (sgn f) e) *\<^sub>C sgn f))\<^sup>2)\ apply (rule sum.cong[OF refl]) by simp also have \\ = (norm (\f\M. ((cinner (sgn f) e) *\<^sub>C sgn f)))\<^sup>2\ apply (rule pythagorean_theorem_sum[symmetric]) using that apply auto by (meson \is_ortho_set M\ is_ortho_set_def) also have \\ = (norm (\f\M. proj f e))\<^sup>2\ by (metis butterfly_apply butterfly_sgn_eq_proj) also have \\ = (norm (Proj (ccspan M) e))\<^sup>2\ apply (rule arg_cong[where f=\\x. (norm x)\<^sup>2\]) using \finite M\ \is_ortho_set M\ apply induction apply simp by (smt (verit, ccfv_threshold) Proj_orthog_ccspan_insert insertCI is_ortho_set_def plus_cblinfun.rep_eq sum.insert) also have \\ \ (norm (Proj (ccspan M)) * norm e)\<^sup>2\ by (auto simp: norm_cblinfun intro!: power_mono) also have \\ \ (norm e)\<^sup>2\ apply (rule power_mono) apply (metis norm_Proj_leq1 mult_left_le_one_le norm_ge_zero) by simp finally show ?thesis by - qed then have \(\f. (cmod (cinner (sgn f) e))\<^sup>2) abs_summable_on F\ apply (intro nonneg_bdd_above_summable_on bdd_aboveI) by auto then have \countable {f \ F. (cmod (sgn f \\<^sub>C e))\<^sup>2 \ 0}\ by (rule abs_summable_countable) then have \countable {f \ F. \ is_orthogonal (sgn f) e}\ by force then have \countable {f \ F. \ is_orthogonal f e}\ by force then show ?thesis unfolding F'_def by simp qed then have F'_leq: \|F' e| \o natLeq\ for e using countable_leq_natLeq by auto from F'_union have \|F| \o |Sigma E F'|\ (is \_ \o \\) using card_of_UNION_Sigma by blast also have \\ \o |E \ (UNIV::nat set)|\ (is \_ \o \\) apply (rule card_of_Sigma_mono1) using F'_leq using card_of_nat ordIso_symmetric ordLeq_ordIso_trans by blast also have \\ =o |E| *c natLeq\ (is \ordIso2 _ \\) by (metis Field_card_of Field_natLeq card_of_ordIso_subst cprod_def) also have \\ =o |E|\ apply (rule card_prod_omega) using that by (simp add: cinfinite_def) finally show \|F| \o |E|\ by - qed then have infinite: \|E| =o |F|\ if \infinite E\ and \infinite F\ by (simp add: assms ordIso_iff_ordLeq that(1) that(2)) have \|E| =o |F|\ if \finite E\ and \is_ortho_set E\ \is_ortho_set F\ \ccspan E = top\ \ccspan F = top\ for E F :: \'a set\ proof - have \card E = card F\ using that by (metis bij_betw_same_card ccspan.rep_eq closure_finite_cspan complex_vector.bij_if_span_eq_span_bases complex_vector.independent_span_bound is_ortho_set_cindependent top_ccsubspace.rep_eq top_greatest) with \finite E\ have \finite F\ by (metis ccspan.rep_eq closure_finite_cspan complex_vector.independent_span_bound is_ortho_set_cindependent that(3) that(4) top_ccsubspace.rep_eq top_greatest) with \card E = card F\ that show ?thesis apply (rule_tac finite_card_of_iff_card[THEN iffD2]) by auto qed with infinite assms have \|E| =o |F|\ by (meson ordIso_symmetric) then show ?thesis by (simp add: card_of_ordIso) qed lemma all_onbs_same_card: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows \\f. bij_betw f E F\ apply (rule all_ortho_bases_same_card) using assms by (auto simp: is_onb_def) definition bij_between_bases where \bij_between_bases E F = (SOME f. bij_betw f E F)\ for E F :: \'a::chilbert_space set\ lemma fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows bij_between_bases_bij: \bij_betw (bij_between_bases E F) E F\ using all_onbs_same_card by (metis assms(1) assms(2) bij_between_bases_def someI) definition unitary_between where \unitary_between E F = cblinfun_extension E (bij_between_bases E F)\ lemma unitary_between_apply: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ \e \ E\ shows \unitary_between E F *\<^sub>V e = bij_between_bases E F e\ proof - from \is_onb E\ \is_onb F\ have EF: \bij_between_bases E F e \ F\ if \e \ E\ for e by (meson bij_betwE bij_between_bases_bij that) have ortho: \is_orthogonal (bij_between_bases E F x) (bij_between_bases E F y)\ if \x \ y\ and \x \ E\ and \y \ E\ for x y by (smt (verit, del_insts) assms(1) assms(2) bij_betw_iff_bijections bij_between_bases_bij is_onb_def is_ortho_set_def that(1) that(2) that(3)) have spanE: \closure (cspan E) = UNIV\ by (metis assms(1) ccspan.rep_eq is_onb_def top_ccsubspace.rep_eq) show ?thesis unfolding unitary_between_def apply (rule cblinfun_extension_apply) apply (rule cblinfun_extension_exists_ortho[where B=1]) using assms EF ortho spanE by (auto simp: is_onb_def) qed lemma unitary_between_unitary: fixes E F :: \'a::chilbert_space set\ assumes \is_onb E\ \is_onb F\ shows \unitary (unitary_between E F)\ proof - have \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V c) = b \\<^sub>C c\ if \b \ E\ and \c \ E\ for b c proof (cases \b = c\) case True from \is_onb E\ that have 1: \b \\<^sub>C b = 1\ using cnorm_eq_1 is_onb_def by blast from that have \unitary_between E F *\<^sub>V b \ F\ by (metis assms(1) assms(2) bij_betw_apply bij_between_bases_bij unitary_between_apply) with \is_onb F\ have 2: \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V b) = 1\ by (simp add: cnorm_eq_1 is_onb_def) from 1 2 True show ?thesis by simp next case False from \is_onb E\ that have 1: \b \\<^sub>C c = 0\ by (simp add: False is_onb_def is_ortho_set_def) from that have inF: \unitary_between E F *\<^sub>V b \ F\ \unitary_between E F *\<^sub>V c \ F\ by (metis assms(1) assms(2) bij_betw_apply bij_between_bases_bij unitary_between_apply)+ have neq: \unitary_between E F *\<^sub>V b \ unitary_between E F *\<^sub>V c\ by (metis (no_types, lifting) False assms(1) assms(2) bij_betw_iff_bijections bij_between_bases_bij that(1) that(2) unitary_between_apply) from inF neq \is_onb F\ have 2: \(unitary_between E F *\<^sub>V b) \\<^sub>C (unitary_between E F *\<^sub>V c) = 0\ by (simp add: is_onb_def is_ortho_set_def) from 1 2 show ?thesis by simp qed then have iso: \isometry (unitary_between E F)\ apply (rule_tac orthogonal_on_basis_is_isometry[where B=E]) using assms(1) is_onb_def by auto have \unitary_between E F *\<^sub>S top = unitary_between E F *\<^sub>S ccspan E\ by (metis assms(1) is_onb_def) also have \\ \ ccspan (unitary_between E F ` E)\ (is \_ \ \\) by (simp add: cblinfun_image_ccspan) also have \\ = ccspan (bij_between_bases E F ` E)\ by (metis assms(1) assms(2) image_cong unitary_between_apply) also have \\ = ccspan F\ by (metis assms(1) assms(2) bij_betw_imp_surj_on bij_between_bases_bij) also have \\ = top\ using assms(2) is_onb_def by blast finally have surj: \unitary_between E F *\<^sub>S top = top\ by (simp add: top.extremum_unique) from iso surj show ?thesis by (rule surj_isometry_is_unitary) qed subsection \Notation\ bundle cblinfun_notation begin notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 67) notation cblinfun_apply (infixr "*\<^sub>V" 70) notation cblinfun_image (infixr "*\<^sub>S" 70) notation adj ("_*" [99] 100) end bundle no_cblinfun_notation begin no_notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 67) no_notation cblinfun_apply (infixr "*\<^sub>V" 70) no_notation cblinfun_image (infixr "*\<^sub>S" 70) no_notation adj ("_*" [99] 100) end bundle blinfun_notation begin notation blinfun_apply (infixr "*\<^sub>V" 70) end bundle no_blinfun_notation begin no_notation blinfun_apply (infixr "*\<^sub>V" 70) end unbundle no_cblinfun_notation unbundle no_lattice_syntax end diff --git a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy --- a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy +++ b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy @@ -1,2494 +1,2523 @@ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) section \\Complex_Inner_Product\ -- Complex Inner Product Spaces\ theory Complex_Inner_Product imports Complex_Inner_Product0 begin subsection \Complex inner product spaces\ 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 (simp add: cdot_square_norm) lemmas cinner_commute' [simp] = cinner_commute[symmetric] lemma (in complex_inner) cinner_eq_flip: \(cinner x y = cinner z w) \ (cinner y x = cinner w z)\ by (metis cinner_commute) lemma Im_cinner_x_x[simp]: "Im \x , x\ = 0" using comp_Im_same[OF cinner_ge_zero] by simp lemma of_complex_inner_1' [simp]: "cinner (1 :: 'a :: {complex_inner, complex_normed_algebra_1}) (of_complex x) = x" by (metis cinner_commute complex_cnj_cnj of_complex_inner_1) class chilbert_space = complex_inner + complete_space begin subclass cbanach by standard end instantiation complex :: "chilbert_space" begin instance .. end subsection \Misc facts\ lemma cinner_scaleR_left [simp]: "cinner (scaleR r x) y = of_real r * (cinner x y)" by (simp add: scaleR_scaleC) lemma cinner_scaleR_right [simp]: "cinner x (scaleR r y) = of_real r * (cinner x y)" by (simp add: scaleR_scaleC) text \This is a useful rule for establishing the equality of vectors\ lemma cinner_extensionality: assumes \\\. \\, \\ = \\, \\\ 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) 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 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 by blast qed proposition parallelogram_law: includes notation_norm fixes x y :: "'a::complex_inner" shows \\x+y\^2 + \x-y\^2 = 2*( \x\^2 + \y\^2 )\ \ \Shown in the proof of Theorem 2.3 in @{cite conway2013course}\ by (simp add: polar_identity_minus polar_identity) theorem pythagorean_theorem: includes notation_norm shows \\x , 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 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 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 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) have \bounded (range y)\ using a2 by (simp add: Elementary_Metric_Spaces.cauchy_imp_bounded) hence b2: \\ M. \ n. norm (y n) < M\ by (meson bounded_pos_less rangeI) have \\M. \n. norm (x n) < M \ norm (y n) < M\ using b1 b2 by (metis dual_order.strict_trans linorder_neqE_linordered_idom) then obtain M where M1: \\n. norm (x n) < M\ and M2: \\n. norm (y n) < M\ by blast have M3: \M > 0\ by (smt M2 norm_not_less_zero) have \\N. \n \ N. \m \ N. norm ( (\ i. \ x i, 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) then obtain N1 where N1_def: \\n m. n\N1 \ m\N1 \ norm (x n - x m) < e / (2*M)\ by blast have x1: \\N. \ n\N. \ m\N. norm (y n - y m) < e / (2*M)\ using a2 \e / (2*M) > 0\ by (simp add: Cauchy_iff) obtain N2 where N2_def: \\n m. n\N2 \ m\N2 \ norm (y n - y m) < e / (2*M)\ using x1 by blast define N where N_def: \N = N1 + N2\ hence \N \ N1\ by auto have \N \ N2\ using N_def by auto have \norm ( \ x n, 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) 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 hence y2: \norm (\ x n, y n \ - \ x m, y n \) < e/2\ 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 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: fixes A :: \'a::{real_normed_vector,not_singleton} \ 'b::complex_inner\ assumes \bounded_linear A\ shows \onorm A = (SUP (\,\). cmod (cinner \ (A \)) / (norm \ * norm \))\ proof (unfold onorm_def, rule cSup_eq_cSup) show \bdd_above (range (\x. norm (A x) / norm x))\ by (meson assms bdd_aboveI2 le_onorm) next fix a assume \a \ range (\\. norm (A \) / norm \)\ then obtain \ where \a = norm (A \) / norm \\ by auto then have \a \ cmod (cinner (A \) (A \)) / (norm (A \) * norm \)\ apply auto by (smt (verit) divide_divide_eq_left norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) then show \\b\range (\(\, \). cmod (cinner \ (A \)) / (norm \ * norm \)). a \ b\ by force next fix b assume \b \ range (\(\, \). cmod (cinner \ (A \)) / (norm \ * norm \))\ then obtain \ \ where b: \b = cmod (cinner \ (A \)) / (norm \ * norm \)\ by auto then have \b \ norm (A \) / norm \\ apply auto by (smt (verit, ccfv_threshold) complex_inner_class.Cauchy_Schwarz_ineq2 division_ring_divide_zero linordered_field_class.divide_right_mono mult_cancel_left1 nonzero_mult_divide_mult_cancel_left2 norm_imp_pos_and_ge ordered_field_class.sign_simps(33) zero_le_divide_iff) then show \\a\range (\x. norm (A x) / norm x). b \ a\ by auto qed lemma sum_cinner: fixes f :: "'a \ 'b::complex_inner" shows "cinner (sum f A) (sum g B) = (\i\A. \j\B. cinner (f i) (g j))" by (simp add: cinner_sum_right cinner_sum_left) (rule sum.swap) lemma Cauchy_cinner_product_summable': fixes a b :: "nat \ 'a::complex_inner" shows \(\(x, y). cinner (a x) (b y)) summable_on UNIV \ (\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ proof - have img: \(\(k::nat, i). (i, k - i)) ` {(k, i). i \ k} = UNIV\ apply (auto simp: image_def) by (metis add.commute add_diff_cancel_right' diff_le_self) have inj: \inj_on (\(k::nat, i). (i, k - i)) {(k, i). i \ k}\ by (smt (verit, del_insts) Pair_inject case_prodE case_prod_conv eq_diff_iff inj_onI mem_Collect_eq) have \(\(x, y). cinner (a x) (b y)) summable_on UNIV \ (\(k, l). cinner (a k) (b l)) summable_on (\(k, i). (i, k - i)) ` {(k, i). i \ k}\ by (simp only: img) also have \\ \ ((\(k, l). cinner (a k) (b l)) \ (\(k, i). (i, k - i))) summable_on {(k, i). i \ k}\ using inj by (rule summable_on_reindex) also have \\ \ (\(x, y). cinner (a y) (b (x - y))) summable_on {(k, i). i \ k}\ by (simp add: o_def case_prod_unfold) finally show ?thesis by - qed instantiation prod :: (complex_inner, complex_inner) complex_inner begin definition cinner_prod_def: "cinner x y = cinner (fst x) (fst y) + cinner (snd x) (snd y)" instance proof fix r :: complex fix x y z :: "'a::complex_inner \ 'b::complex_inner" show "cinner x y = cnj (cinner y x)" unfolding cinner_prod_def by simp show "cinner (x + y) z = cinner x z + cinner y z" unfolding cinner_prod_def by (simp add: cinner_add_left) show "cinner (scaleC r x) y = cnj r * cinner x y" unfolding cinner_prod_def by (simp add: distrib_left) show "0 \ cinner x x" unfolding cinner_prod_def by (intro add_nonneg_nonneg cinner_ge_zero) show "cinner x x = 0 \ x = 0" unfolding cinner_prod_def prod_eq_iff by (metis antisym cinner_eq_zero_iff cinner_ge_zero fst_zero le_add_same_cancel2 snd_zero verit_sum_simplify) show "norm x = sqrt (cmod (cinner x x))" unfolding norm_prod_def cinner_prod_def by (metis (no_types, lifting) Re_complex_of_real add_nonneg_nonneg cinner_ge_zero complex_of_real_cmod plus_complex.simps(1) power2_norm_eq_cinner') qed end instance prod :: (chilbert_space, chilbert_space) chilbert_space.. subsection \Orthogonality\ definition "orthogonal_complement S = {x| x. \y\S. cinner x y = 0}" lemma orthogonal_complement_orthoI: \x \ orthogonal_complement M \ y \ M \ \ x, 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 is_orthogonal_sgn_right[simp]: \is_orthogonal e (sgn f) \ is_orthogonal e f\ proof (cases \f = 0\) case True then show ?thesis by simp next case False have \cinner e (sgn f) = cinner e f / norm f\ by (simp add: sgn_div_norm divide_inverse scaleR_scaleC) moreover have \norm f \ 0\ by (simp add: False) ultimately show ?thesis by force qed lemma is_orthogonal_sgn_left[simp]: \is_orthogonal (sgn e) f \ is_orthogonal e f\ by (simp add: is_orthogonal_sym) lemma orthogonal_complement_closed_subspace[simp]: "closed_csubspace (orthogonal_complement A)" for A :: \('a::complex_inner) set\ proof (intro closed_csubspace.intro complex_vector.subspaceI) fix x y and c show \0 \ orthogonal_complement A\ by (rule orthogonal_complementI, simp) show \x + y \ orthogonal_complement A\ if \x \ orthogonal_complement A\ and \y \ orthogonal_complement A\ using that by (auto intro!: orthogonal_complementI dest!: orthogonal_complement_orthoI simp add: cinner_add_left) show \c *\<^sub>C x \ orthogonal_complement A\ if \x \ orthogonal_complement A\ using that by (auto intro!: orthogonal_complementI dest!: orthogonal_complement_orthoI) show "closed (orthogonal_complement A)" proof (auto simp add: closed_sequential_limits, rename_tac an a) fix an a assume ortho: \\n::nat. an n \ orthogonal_complement A\ assume lim: \an \ a\ have \\ y \ A. \ n. \ 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\ by fastforce 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 definition is_ortho_set :: "'a::complex_inner set \ bool" where \ \Orthogonal set\ \is_ortho_set S = ((\x\S. \y\S. x \ y \ \x, y\ = 0) \ 0 \ S)\ definition is_onb where \is_onb E \ is_ortho_set E \ (\b\E. norm b = 1) \ ccspan E = top\ lemma is_ortho_set_empty[simp]: "is_ortho_set {}" unfolding is_ortho_set_def by auto lemma is_ortho_set_antimono: \A \ B \ is_ortho_set B \ is_ortho_set A\ unfolding is_ortho_set_def by auto lemma orthogonal_complement_of_closure: fixes A ::"('a::complex_inner) set" shows "orthogonal_complement A = orthogonal_complement (closure A)" proof- have s1: \\ y, x \ = 0\ if a1: "x \ (orthogonal_complement 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 using limI by force qed hence "x \ orthogonal_complement (closure A)" if a1: "x \ (orthogonal_complement A)" for x using that by (meson orthogonal_complementI is_orthogonal_sym) moreover have \x \ (orthogonal_complement A)\ if "x \ (orthogonal_complement (closure A))" for x using that by (meson closure_subset orthogonal_complement_orthoI orthogonal_complementI subset_eq) ultimately show ?thesis by blast qed lemma is_orthogonal_closure: assumes \\s. s \ S \ is_orthogonal a s\ assumes \x \ closure S\ shows \is_orthogonal a x\ by (metis assms(1) assms(2) orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI) lemma is_orthogonal_cspan: assumes a1: "\s. s \ S \ is_orthogonal a s" and a3: "x \ cspan S" shows "\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) 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) 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" 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) 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 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) 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 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) 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\) 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 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 using \x = (\ a\T. r a *\<^sub>C a)\ by fastforce qed lemma is_ortho_set_singleton[simp]: \is_ortho_set {x} \ x \ 0\ by (simp add: is_ortho_set_def) subsection \Projections\ lemma smallest_norm_exists: \ \Theorem 2.5 in @{cite conway2013course} (inside the proof)\ includes notation_norm fixes M :: \'a::chilbert_space set\ assumes q1: \convex M\ and q2: \closed M\ and q3: \M \ {}\ shows \\k. is_arg_min (\ x. \x\) (\ t. t \ M) k\ proof- define d where \d = Inf { \x\^2 | x. x \ M }\ have w4: \{ \x\^2 | x. x \ M } \ {}\ by (simp add: assms(3)) have \\ x. \x\^2 \ 0\ by simp hence bdd_below1: \bdd_below { \x\^2 | x. x \ M }\ by fastforce have \d \ \x\^2\ if a1: "x \ M" for x proof- have "\v. (\w. Re (\v , 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 hence \\ \ > 0. \ x \ M. \x\^2 < d + \\ by (simp add: \\x. x \ M \ d \ \x\\<^sup>2\) hence w1: \\ n::nat. \ x \ M. \x\^2 < d + 1/(n+1)\ by auto then obtain r::\nat \ 'a\ where w2: \\ n. r n \ M \ \ r n \^2 < d + 1/(n+1)\ by metis have w3: \\ n. r n \ M\ by (simp add: w2) have \\ n. \ r n \^2 < d + 1/(n+1)\ by (simp add: w2) have w5: \\ (r n) - (r m) \^2 < 2*(1/(n+1) + 1/(m+1))\ for m n proof- have w6: \\ r n \^2 < d + 1/(n+1)\ by (metis w2 of_nat_1 of_nat_add) have \ \ r m \^2 < d + 1/(m+1)\ by (metis w2 of_nat_1 of_nat_add) have \(r n) \ M\ by (simp add: \\n. r n \ M\) moreover have \(r m) \ M\ by (simp add: \\n. r n \ M\) ultimately have \(1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \ M\ using \convex M\ by (simp add: convexD) hence \\ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2 \ d\ by (simp add: \\x. x \ M \ d \ \x\\<^sup>2\) have \\ (1/2) *\<^sub>R (r n) - (1/2) *\<^sub>R (r m) \^2 = (1/2)*( \ r n \^2 + \ r m \^2 ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ by (smt (z3) div_by_1 field_sum_of_halves nonzero_mult_div_cancel_left parallelogram_law polar_identity power2_norm_eq_cinner' scaleR_collapse times_divide_eq_left) also have \... < (1/2)*( d + 1/(n+1) + \ r m \^2 ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ using \\r n\\<^sup>2 < d + 1 / real (n + 1)\ by auto also have \... < (1/2)*( d + 1/(n+1) + d + 1/(m+1) ) - \ (1/2) *\<^sub>R (r n) + (1/2) *\<^sub>R (r m) \^2\ using \\r m\\<^sup>2 < d + 1 / real (m + 1)\ by auto also have \... \ (1/2)*( d + 1/(n+1) + d + 1/(m+1) ) - d\ by (simp add: \d \ \(1 / 2) *\<^sub>R r n + (1 / 2) *\<^sub>R r m\\<^sup>2\) also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) + 2*d ) - d\ by simp also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + (1/2)*(2*d) - d\ by (simp add: distrib_left) also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) ) + d - d\ by simp also have \... \ (1/2)*( 1/(n+1) + 1/(m+1) )\ by simp finally have \ \(1 / 2) *\<^sub>R r n - (1 / 2) *\<^sub>R r m\\<^sup>2 < 1 / 2 * (1 / real (n + 1) + 1 / real (m + 1)) \ by blast hence \ \(1 / 2) *\<^sub>R (r n - r m) \\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (simp add: real_vector.scale_right_diff_distrib) hence \ ((1 / 2)*\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by simp hence \ (1 / 2)^2*(\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (metis power_mult_distrib) hence \ (1 / 4) *(\ (r n - r m) \)\<^sup>2 < (1 / 2) * (1 / real (n + 1) + 1 / real (m + 1)) \ by (simp add: power_divide) hence \ \ (r n - r m) \\<^sup>2 < 2 * (1 / real (n + 1) + 1 / real (m + 1)) \ by simp thus ?thesis by (metis of_nat_1 of_nat_add) qed hence "\ N. \ n m. n \ N \ m \ N \ \ (r n) - (r m) \^2 < \^2" if "\ > 0" for \ proof- obtain N::nat where \1/(N + 1) < \^2/4\ using LIMSEQ_ignore_initial_segment[OF lim_inverse_n', where k=1] by (metis Suc_eq_plus1 \0 < \\ nat_approx_posE zero_less_divide_iff zero_less_numeral zero_less_power ) hence \4/(N + 1) < \^2\ by simp have "2*(1/(n+1) + 1/(m+1)) < \^2" if f1: "n \ N" and f2: "m \ N" for m n::nat proof- have \1/(n+1) \ 1/(N+1)\ by (simp add: f1 linordered_field_class.frac_le) moreover have \1/(m+1) \ 1/(N+1)\ by (simp add: f2 linordered_field_class.frac_le) ultimately have \2*(1/(n+1) + 1/(m+1)) \ 4/(N+1)\ by simp thus ?thesis using \4/(N + 1) < \^2\ by linarith qed hence "\ (r n) - (r m) \^2 < \^2" if y1: "n \ N" and y2: "m \ N" for m n::nat using that by (smt \\n m. \r n - r m\\<^sup>2 < 2 * (1 / (real n + 1) + 1 / (real m + 1))\ of_nat_1 of_nat_add) thus ?thesis by blast qed hence \\ \ > 0. \ N::nat. \ n m::nat. n \ N \ m \ N \ \ (r n) - (r m) \^2 < \^2\ by blast hence \\ \ > 0. \ N::nat. \ n m::nat. n \ N \ m \ N \ \ (r n) - (r m) \ < \\ by (meson less_eq_real_def power_less_imp_less_base) hence \Cauchy r\ using CauchyI by fastforce then obtain k where \r \ k\ using convergent_eq_Cauchy by auto have \k \ M\ using \closed M\ using \\n. r n \ M\ \r \ k\ closed_sequentially by auto have \(\ n. \ r n \^2) \ \ k \^2\ by (simp add: \r \ k\ tendsto_norm tendsto_power) moreover have \(\ n. \ r n \^2) \ d\ proof- have \\\ r n \^2 - d\ < 1/(n+1)\ for n :: nat using \\x. x \ M \ d \ \x\\<^sup>2\ \\n. r n \ M \ \r n\\<^sup>2 < d + 1 / (real n + 1)\ of_nat_1 of_nat_add by smt moreover have \(\n. 1 / real (n + 1)) \ 0\ using LIMSEQ_ignore_initial_segment[OF lim_inverse_n', where k=1] by blast ultimately have \(\ n. \\ r n \^2 - d\ ) \ 0\ by (simp add: LIMSEQ_norm_0) hence \(\ n. \ r n \^2 - d ) \ 0\ by (simp add: tendsto_rabs_zero_iff) moreover have \(\ n. d ) \ d\ by simp ultimately have \(\ n. (\ r n \^2 - d)+d ) \ 0+d\ using tendsto_add by fastforce thus ?thesis by simp qed ultimately have \d = \ k \^2\ using LIMSEQ_unique by auto hence \t \ M \ \ k \^2 \ \ t \^2\ for t using \\x. x \ M \ d \ \x\\<^sup>2\ by auto hence q1: \\ k. is_arg_min (\ x. \x\^2) (\ t. t \ M) k\ using \k \ M\ is_arg_min_def \d = \k\\<^sup>2\ by smt thus \\ k. is_arg_min (\ x. \x\) (\ t. t \ M) k\ by (smt is_arg_min_def norm_ge_zero power2_eq_square power2_le_imp_le) qed lemma smallest_norm_unique: \ \Theorem 2.5 in @{cite conway2013course} (inside the proof)\ includes notation_norm fixes M :: \'a::complex_inner set\ assumes q1: \convex M\ assumes r: \is_arg_min (\ x. \x\) (\ t. t \ M) r\ assumes s: \is_arg_min (\ x. \x\) (\ t. t \ M) s\ shows \r = s\ proof - have \r \ M\ using \is_arg_min (\x. \x\) (\ t. t \ M) r\ by (simp add: is_arg_min_def) moreover have \s \ M\ using \is_arg_min (\x. \x\) (\ t. t \ M) s\ by (simp add: is_arg_min_def) ultimately have \((1/2) *\<^sub>R r + (1/2) *\<^sub>R s) \ M\ using \convex M\ by (simp add: convexD) hence \\r\ \ \ (1/2) *\<^sub>R r + (1/2) *\<^sub>R s \\ by (metis is_arg_min_linorder r) hence u2: \\r\^2 \ \ (1/2) *\<^sub>R r + (1/2) *\<^sub>R s \^2\ using norm_ge_zero power_mono by blast have \\r\ \ \s\\ using r s is_arg_min_def by (metis is_arg_min_linorder) moreover have \\s\ \ \r\\ using r s is_arg_min_def by (metis is_arg_min_linorder) ultimately have u3: \\r\ = \s\\ by simp have \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \^2 \ 0\ using u2 u3 parallelogram_law by (smt (verit, ccfv_SIG) polar_identity_minus power2_norm_eq_cinner' scaleR_add_right scaleR_half_double) hence \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \^2 = 0\ by simp hence \\ (1/2) *\<^sub>R r - (1/2) *\<^sub>R s \ = 0\ by auto hence \(1/2) *\<^sub>R r - (1/2) *\<^sub>R s = 0\ using norm_eq_zero by blast thus ?thesis by simp qed theorem smallest_dist_exists: \ \Theorem 2.5 in @{cite conway2013course}\ fixes M::\'a::chilbert_space set\ and h assumes a1: \convex M\ and a2: \closed M\ and a3: \M \ {}\ shows \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ proof- have *: "is_arg_min (\x. dist x h) (\x. x\M) (k+h) \ is_arg_min (\x. norm x) (\x. x\(\x. x-h) ` M) k" for k unfolding dist_norm is_arg_min_def apply auto using add_implies_diff by blast have \\k. is_arg_min (\x. dist x h) (\x. x\M) (k+h)\ apply (subst *) apply (rule smallest_norm_exists) using assms by (auto simp: closed_translation_subtract) then show \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ by metis qed theorem smallest_dist_unique: \ \Theorem 2.5 in @{cite conway2013course}\ fixes M::\'a::complex_inner set\ and h assumes a1: \convex M\ assumes \is_arg_min (\ x. dist x h) (\ x. x \ M) r\ assumes \is_arg_min (\ x. dist x h) (\ x. x \ M) s\ shows \r = s\ proof- have *: "is_arg_min (\x. dist x h) (\x. x\M) k \ is_arg_min (\x. norm x) (\x. x\(\x. x-h) ` M) (k-h)" for k unfolding dist_norm is_arg_min_def by auto have \r - h = s - h\ using _ assms(2,3)[unfolded *] apply (rule smallest_norm_unique) by (simp add: a1) thus \r = s\ by auto qed \ \Theorem 2.6 in @{cite conway2013course}\ theorem smallest_dist_is_ortho: fixes M::\'a::complex_inner set\ and h k::'a assumes b1: \closed_csubspace M\ shows \(is_arg_min (\ x. dist x h) (\ x. x \ M) k) \ h - k \ (orthogonal_complement M) \ k \ M\ proof- include notation_norm have \csubspace M\ using \closed_csubspace M\ unfolding closed_csubspace_def by blast have r1: \2 * Re (\ h - k , f \) \ \ f \^2\ if "f \ M" and \k \ M\ and \is_arg_min (\x. dist x h) (\ x. x \ M) k\ for f proof- have \k + f \ M\ using \csubspace M\ by (simp add:complex_vector.subspace_add that) have "\f A a b. \ is_arg_min f (\ x. x \ A) (a::'a) \ (f a::real) \ f b \ b \ A" by (metis (no_types) is_arg_min_linorder) hence "dist k h \ dist (f + k) h" by (metis \is_arg_min (\x. dist x h) (\ x. x \ M) k\ \k + f \ M\ add.commute) hence \dist h k \ dist h (k + f)\ by (simp add: add.commute dist_commute) hence \\ h - k \ \ \ h - (k + f) \\ by (simp add: dist_norm) hence \\ h - k \^2 \ \ h - (k + f) \^2\ by (simp add: power_mono) also have \... \ \ (h - k) - f \^2\ by (simp add: diff_diff_add) also have \... \ \ (h - k) \^2 + \ f \^2 - 2 * Re (\ h - k , 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 using True by auto next case False hence \\ f \^2 = 0\ by simp thus ?thesis by auto qed have q3: \\ c::real. c > 0 \ 2 * Re (\ h - k , 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 using q4 by smt qed have w2: "h - k \ orthogonal_complement M \ k \ M" if a1: "is_arg_min (\ x. dist x h) (\ x. x \ M) k" proof- have \k \ M\ using is_arg_min_def that by fastforce hence \\ f. f \ M \ 2 * Re (\ h - k , f \) \ \ f \^2\ using r1 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 \ (\ c::real. c * (2 * Re (\ h - k , f \)) \ c^2*\ f \^2)\ by auto hence \\ f. f \ M \ (\ c::real. c > 0 \ c * (2 * Re (\ h - k , f \)) \ c^2*\ f \^2)\ by simp 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 \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ c*\ f \^2)\ 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 \ (\ 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 \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) \ 0)\ by simp hence \\ f. f \ M \ (\ c::real. c > 0 \ 2 * Re (\ h - k , f \) = 0)\ using \\ f. f \ M \ (\ c::real. c > 0 \ (2 * Re (\ h - k , f \)) \ 0)\ by fastforce 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 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 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) 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\ if "f \ M" for f using \h - k \ orthogonal_complement M \ k \ M\ orthogonal_complement_orthoI that by auto hence \h - k \ orthogonal_complement M\ by (simp add: orthogonal_complement_def) thus ?thesis using \k \ M\ by auto qed have q1: \dist h k \ dist h f \ if "f \ M" and \h - k \ orthogonal_complement M \ k \ M\ for f proof- have \\ h - k, k - f \ = 0\ by (metis (no_types, lifting) that cinner_diff_right diff_0_right orthogonal_complement_orthoI that) have \\ h - f \^2 = \ (h - k) + (k - f) \^2\ by simp also have \... = \ h - k \^2 + \ k - f \^2\ using \\ h - k, k - f \ = 0\ pythagorean_theorem by blast also have \... \ \ h - k \^2\ by simp finally have \\h - k\\<^sup>2 \ \h - f\\<^sup>2 \ by blast hence \\h - k\ \ \h - f\\ using norm_ge_zero power2_le_imp_le by blast thus ?thesis by (simp add: dist_norm) qed have w1: "is_arg_min (\ x. dist x h) (\ x. x \ M) k" if "h - k \ orthogonal_complement M \ k \ M" proof- have \h - k \ orthogonal_complement M\ using that by blast have \k \ M\ using \h - k \ orthogonal_complement M \ k \ M\ by blast thus ?thesis by (metis (no_types, lifting) dist_commute is_arg_min_linorder q1 that) qed show ?thesis using w1 w2 by blast qed corollary orthog_proj_exists: fixes M :: \'a::chilbert_space set\ assumes \closed_csubspace M\ shows \\k. h - k \ orthogonal_complement M \ k \ M\ proof- from \closed_csubspace M\ have \M \ {}\ using closed_csubspace.subspace complex_vector.subspace_0 by blast have \closed M\ using \closed_csubspace M\ by (simp add: closed_csubspace.closed) have \convex M\ using \closed_csubspace M\ by (simp) have \\k. is_arg_min (\ x. dist x h) (\ x. x \ M) k\ by (simp add: smallest_dist_exists \closed M\ \convex M\ \M \ {}\) thus ?thesis by (simp add: assms smallest_dist_is_ortho) qed corollary orthog_proj_unique: fixes M :: \'a::complex_inner set\ assumes \closed_csubspace M\ assumes \h - r \ orthogonal_complement M \ r \ M\ assumes \h - s \ orthogonal_complement M \ s \ M\ shows \r = s\ using _ assms(2,3) unfolding smallest_dist_is_ortho[OF assms(1), symmetric] apply (rule smallest_dist_unique) using assms(1) by (simp) definition is_projection_on::\('a \ 'a) \ ('a::metric_space) set \ bool\ where \is_projection_on \ M \ (\h. is_arg_min (\ x. dist x h) (\ x. x \ M) (\ h))\ lemma is_projection_on_iff_orthog: \closed_csubspace M \ is_projection_on \ M \ (\h. h - \ h \ orthogonal_complement M \ \ h \ M)\ by (simp add: is_projection_on_def smallest_dist_is_ortho) lemma is_projection_on_exists: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows "\\. is_projection_on \ M" unfolding is_projection_on_def apply (rule choice) using smallest_dist_exists[OF assms] by auto lemma is_projection_on_unique: fixes M :: \'a::complex_inner set\ assumes \convex M\ assumes "is_projection_on \\<^sub>1 M" assumes "is_projection_on \\<^sub>2 M" shows "\\<^sub>1 = \\<^sub>2" using smallest_dist_unique[OF assms(1)] using assms(2,3) unfolding is_projection_on_def by blast definition projection :: \'a::metric_space set \ ('a \ 'a)\ where \projection M \ SOME \. is_projection_on \ M\ lemma projection_is_projection_on: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows "is_projection_on (projection M) M" by (metis assms(1) assms(2) assms(3) is_projection_on_exists projection_def someI) lemma projection_is_projection_on'[simp]: \ \Common special case of @{thm projection_is_projection_on}\ fixes M :: \'a::chilbert_space set\ assumes \closed_csubspace M\ shows "is_projection_on (projection M) M" apply (rule projection_is_projection_on) apply (auto simp add: assms closed_csubspace.closed) using assms closed_csubspace.subspace complex_vector.subspace_0 by blast lemma projection_orthogonal: fixes M :: \'a::chilbert_space set\ assumes "closed_csubspace M" and \m \ M\ shows \is_orthogonal (h - projection M h) m\ by (metis assms(1) assms(2) closed_csubspace.closed closed_csubspace.subspace csubspace_is_convex empty_iff is_projection_on_iff_orthog orthogonal_complement_orthoI projection_is_projection_on) lemma is_projection_on_in_image: assumes "is_projection_on \ M" shows "\ h \ M" using assms by (simp add: is_arg_min_def is_projection_on_def) lemma is_projection_on_image: assumes "is_projection_on \ M" shows "range \ = M" using assms apply (auto simp: is_projection_on_in_image) by (smt (verit, ccfv_threshold) dist_pos_lt dist_self is_arg_min_def is_projection_on_def rangeI) lemma projection_in_image[simp]: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows \projection M h \ M\ by (simp add: assms(1) assms(2) assms(3) is_projection_on_in_image projection_is_projection_on) lemma projection_image[simp]: fixes M :: \'a::chilbert_space set\ assumes \convex M\ and \closed M\ and \M \ {}\ shows \range (projection M) = M\ by (simp add: assms(1) assms(2) assms(3) is_projection_on_image projection_is_projection_on) lemma projection_eqI': fixes M :: \'a::complex_inner set\ assumes \convex M\ assumes \is_projection_on f M\ shows \projection M = f\ by (metis assms(1) assms(2) is_projection_on_unique projection_def someI_ex) lemma is_projection_on_eqI: fixes M :: \'a::complex_inner set\ assumes a1: \closed_csubspace M\ and a2: \h - x \ orthogonal_complement M\ and a3: \x \ M\ and a4: \is_projection_on \ M\ shows \\ h = x\ by (meson a1 a2 a3 a4 closed_csubspace.subspace csubspace_is_convex is_projection_on_def smallest_dist_is_ortho smallest_dist_unique) lemma projection_eqI: fixes M :: \('a::chilbert_space) set\ assumes \closed_csubspace M\ and \h - x \ orthogonal_complement M\ and \x \ M\ shows \projection M h = x\ by (metis assms(1) assms(2) assms(3) is_projection_on_iff_orthog orthog_proj_exists projection_def is_projection_on_eqI tfl_some) lemma is_projection_on_fixes_image: fixes M :: \'a::metric_space set\ assumes a1: "is_projection_on \ M" and a3: "x \ M" shows "\ x = x" by (metis a1 a3 dist_pos_lt dist_self is_arg_min_def is_projection_on_def) lemma projection_fixes_image: fixes M :: \('a::chilbert_space) set\ assumes 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 hence \\ k \ M. \ h - \ h , k \ = 0\ 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 using norm_ge_zero power2_le_imp_le by blast qed proposition projection_reduces_norm: includes notation_norm fixes M :: \'a::chilbert_space set\ assumes a1: "closed_csubspace M" shows \\ projection M h \ \ \ h \\ using assms is_projection_on_iff_orthog orthog_proj_exists is_projection_on_reduces_norm projection_eqI by blast \ \Theorem 2.7 (version) in @{cite conway2013course}\ theorem is_projection_on_bounded_clinear: fixes M :: \'a::complex_inner set\ assumes a1: "is_projection_on \ M" and a2: "closed_csubspace M" shows "bounded_clinear \" proof have b1: \csubspace (orthogonal_complement M)\ by (simp add: a2) have f1: "\a. a - \ a \ orthogonal_complement M \ \ a \ M" using a1 a2 is_projection_on_iff_orthog by blast hence "c *\<^sub>C x - c *\<^sub>C \ x \ orthogonal_complement M" for c x by (metis (no_types) b1 add_diff_cancel_right' complex_vector.subspace_def diff_add_cancel scaleC_add_right) thus r1: \\ (c *\<^sub>C x) = c *\<^sub>C (\ x)\ for x c using f1 by (meson a2 a1 closed_csubspace.subspace complex_vector.subspace_def is_projection_on_eqI) show r2: \\ (x + y) = (\ x) + (\ y)\ for x y proof- have "\A. \ closed_csubspace (A::'a set) \ csubspace A" by (metis closed_csubspace.subspace) hence "csubspace M" using a2 by auto hence \\ (x + y) - ( (\ x) + (\ y) ) \ M\ by (simp add: complex_vector.subspace_add complex_vector.subspace_diff f1) have \closed_csubspace (orthogonal_complement M)\ using a2 by simp have f1: "\a b. (b::'a) + (a - b) = a" by (metis add.commute diff_add_cancel) have f2: "\a b. (b::'a) - b = a - a" by auto hence f3: "\a. a - a \ orthogonal_complement M" by (simp add: complex_vector.subspace_0) have "\a b. (a \ orthogonal_complement M \ a + b \ orthogonal_complement M) \ b \ orthogonal_complement M" using add_diff_cancel_right' b1 complex_vector.subspace_diff by metis hence "\a b c. (a \ orthogonal_complement M \ c - (b + a) \ orthogonal_complement M) \ c - b \ orthogonal_complement M" using f1 by (metis diff_diff_add) hence f4: "\a b f. (f a - b \ orthogonal_complement M \ a - b \ orthogonal_complement M) \ \ is_projection_on f M" using f1 by (metis a2 is_projection_on_iff_orthog) have f5: "\a b c d. (d::'a) - (c + (b - a)) = d + (a - (b + c))" by auto have "x - \ x \ orthogonal_complement M" using a1 a2 is_projection_on_iff_orthog by blast hence q1: \\ (x + y) - ( (\ x) + (\ y) ) \ orthogonal_complement M\ using f5 f4 f3 by (metis \csubspace (orthogonal_complement M)\ \is_projection_on \ M\ add_diff_eq complex_vector.subspace_diff diff_diff_add diff_diff_eq2) hence \\ (x + y) - ( (\ x) + (\ y) ) \ M \ (orthogonal_complement M)\ by (simp add: \\ (x + y) - (\ x + \ y) \ M\) moreover have \M \ (orthogonal_complement M) = {0}\ by (simp add: \closed_csubspace M\ complex_vector.subspace_0 orthogonal_complement_zero_intersection) ultimately have \\ (x + y) - ( (\ x) + (\ y) ) = 0\ by auto thus ?thesis by simp qed from is_projection_on_reduces_norm show t1: \\ K. \ x. norm (\ x) \ norm x * K\ by (metis a1 a2 mult.left_neutral ordered_field_class.sign_simps(5)) qed theorem projection_bounded_clinear: fixes M :: \('a::chilbert_space) set\ assumes a1: "closed_csubspace M" shows \bounded_clinear (projection M)\ \ \Theorem 2.7 in @{cite conway2013course}\ using assms is_projection_on_iff_orthog orthog_proj_exists is_projection_on_bounded_clinear projection_eqI by blast proposition is_projection_on_idem: fixes M :: \('a::complex_inner) set\ assumes "is_projection_on \ M" shows "\ (\ x) = \ x" using is_projection_on_fixes_image is_projection_on_in_image assms by blast proposition projection_idem: fixes M :: "'a::chilbert_space set" assumes a1: "closed_csubspace M" shows "projection M (projection M x) = projection M x" by (metis assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex equals0D projection_fixes_image projection_in_image) proposition is_projection_on_kernel_is_orthogonal_complement: fixes M :: \'a::complex_inner set\ assumes a1: "is_projection_on \ M" and a2: "closed_csubspace M" shows "\ -` {0} = orthogonal_complement M" proof- have "x \ (\ -` {0})" if "x \ orthogonal_complement M" for x by (smt (verit, ccfv_SIG) a1 a2 closed_csubspace_def complex_vector.subspace_def complex_vector.subspace_diff is_projection_on_eqI orthogonal_complement_closed_subspace that vimage_singleton_eq) moreover have "x \ orthogonal_complement M" if s1: "x \ \ -` {0}" for x by (metis a1 a2 diff_zero is_projection_on_iff_orthog that vimage_singleton_eq) ultimately show ?thesis by blast qed \ \Theorem 2.7 in @{cite conway2013course}\ proposition projection_kernel_is_orthogonal_complement: fixes M :: \'a::chilbert_space set\ assumes "closed_csubspace M" shows "(projection M) -` {0} = (orthogonal_complement M)" by (metis assms closed_csubspace_def complex_vector.subspace_def csubspace_is_convex insert_absorb insert_not_empty is_projection_on_kernel_is_orthogonal_complement projection_is_projection_on) lemma is_projection_on_id_minus: fixes M :: \'a::complex_inner set\ assumes is_proj: "is_projection_on \ M" and cc: "closed_csubspace M" shows "is_projection_on (id - \) (orthogonal_complement M)" using is_proj apply (simp add: cc is_projection_on_iff_orthog) using double_orthogonal_complement_increasing by blast text \Exercise 2 (section 2, chapter I) in @{cite conway2013course}\ lemma projection_on_orthogonal_complement[simp]: fixes M :: "'a::chilbert_space set" assumes a1: "closed_csubspace M" shows "projection (orthogonal_complement M) = id - projection M" apply (auto intro!: ext) by (smt (verit, ccfv_SIG) add_diff_cancel_left' assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex diff_add_cancel double_orthogonal_complement_increasing insert_absorb insert_not_empty is_projection_on_iff_orthog orthogonal_complement_closed_subspace projection_eqI projection_is_projection_on subset_eq) lemma is_projection_on_zero: "is_projection_on (\_. 0) {0}" by (simp add: is_projection_on_def is_arg_min_def) lemma projection_zero[simp]: "projection {0} = (\_. 0)" using is_projection_on_zero by (metis (full_types) is_projection_on_in_image projection_def singletonD someI_ex) lemma is_projection_on_rank1: fixes t :: \'a::complex_inner\ shows \is_projection_on (\x. (\t , 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 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" proof- have b2: "x \ (id - projection M) -` {0}" if c1: "x \ M" for x by (simp add: assms projection_fixes_image that) have b3: \x \ M\ if c1: \x \ (id - projection M) -` {0}\ for x by (metis assms closed_csubspace.closed closed_csubspace.subspace complex_vector.subspace_0 csubspace_is_convex eq_id_iff equals0D fun_diff_def projection_in_image right_minus_eq that vimage_singleton_eq) have \x \ M \ x \ (id - projection M) -` {0}\ for x using b2 b3 by blast hence b4: \( id - (projection M) ) -` {0} = M\ by blast have b1: "orthogonal_complement (orthogonal_complement M) = (projection (orthogonal_complement M)) -` {0}" by (simp add: a1 projection_kernel_is_orthogonal_complement del: projection_on_orthogonal_complement) also have \... = ( id - (projection M) ) -` {0}\ by (simp add: a1) also have \... = M\ by (simp add: b4) finally show ?thesis by blast qed lemma orthogonal_complement_antimono[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]: "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: fixes A B::"('a::complex_inner) set" assumes \0 \ A\ and \0 \ B\ shows \orthogonal_complement (A +\<^sub>M B) = (orthogonal_complement A) \ (orthogonal_complement B)\ proof- have "x \ (orthogonal_complement A) \ (orthogonal_complement B)" if "x \ orthogonal_complement (A +\<^sub>M B)" for x proof- have \orthogonal_complement (A +\<^sub>M B) = orthogonal_complement (A + B)\ unfolding closed_sum_def by (subst orthogonal_complement_of_closure[symmetric], simp) hence \x \ orthogonal_complement (A + B)\ using that by blast hence t1: \\z \ (A + B). \ z , x \ = 0\ by (simp add: orthogonal_complement_orthoI') have \A \ A + B\ using subset_iff add.commute set_zero_plus2 \0 \ B\ by fastforce hence \\z \ A. \ z , x \ = 0\ using t1 by auto hence w1: \x \ (orthogonal_complement A)\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) have \B \ A + B\ using \0 \ A\ subset_iff set_zero_plus2 by blast hence \\ z \ B. \ z , x \ = 0\ using t1 by auto hence \x \ (orthogonal_complement B)\ by (smt mem_Collect_eq is_orthogonal_sym orthogonal_complement_def) thus ?thesis using w1 by auto qed moreover have "x \ (orthogonal_complement (A +\<^sub>M B))" if v1: "x \ (orthogonal_complement A) \ (orthogonal_complement B)" for x proof- have \x \ (orthogonal_complement A)\ using v1 by blast hence \\y\ A. \ y , x \ = 0\ by (simp add: orthogonal_complement_orthoI') 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 lemma orthogonal_complement_of_cspan: \orthogonal_complement A = orthogonal_complement (cspan A)\ by (metis (no_types, opaque_lifting) closed_csubspace.subspace complex_vector.span_minimal complex_vector.span_superset double_orthogonal_complement_increasing orthogonal_complement_antimono orthogonal_complement_closed_subspace subset_antisym) lemma orthogonal_complement_orthogonal_complement_closure_cspan: \orthogonal_complement (orthogonal_complement S) = closure (cspan S)\ for S :: \'a::chilbert_space set\ proof - have \orthogonal_complement (orthogonal_complement S) = orthogonal_complement (orthogonal_complement (closure (cspan S)))\ by (simp flip: orthogonal_complement_of_closure orthogonal_complement_of_cspan) also have \\ = closure (cspan S)\ by simp finally show \orthogonal_complement (orthogonal_complement S) = closure (cspan S)\ by - qed +subsection \Orthogonal spaces\ +(* TODO: Add to report overview *) + +definition \orthogonal_spaces S T \ (\x\space_as_set S. \y\space_as_set T. is_orthogonal x y)\ + +lemma orthogonal_spaces_leq_compl: \orthogonal_spaces S T \ S \ -T\ + unfolding orthogonal_spaces_def apply transfer + by (auto simp: orthogonal_complement_def) + +lemma orthogonal_bot[simp]: \orthogonal_spaces S bot\ + by (simp add: orthogonal_spaces_def) + +lemma orthogonal_spaces_sym: \orthogonal_spaces S T \ orthogonal_spaces T S\ + unfolding orthogonal_spaces_def + using is_orthogonal_sym by blast + +lemma orthogonal_sup: \orthogonal_spaces S T1 \ orthogonal_spaces S T2 \ orthogonal_spaces S (sup T1 T2)\ + apply (rule orthogonal_spaces_sym) + apply (simp add: orthogonal_spaces_leq_compl) + using orthogonal_spaces_leq_compl orthogonal_spaces_sym by blast + +lemma orthogonal_sum: + assumes \finite F\ and \\x. x\F \ orthogonal_spaces S (T x)\ + shows \orthogonal_spaces S (sum T F)\ + using assms + apply induction + by (auto intro!: orthogonal_sup) + + subsection \Orthonormal bases\ lemma ortho_basis_exists: fixes S :: \'a::chilbert_space set\ assumes \is_ortho_set S\ shows \\B. B \ S \ is_ortho_set B \ closure (cspan B) = UNIV\ proof - define on where \on B \ B \ S \ is_ortho_set B\ for B :: \'a set\ have \\B\Collect on. \B'\Collect on. B \ B' \ B' = B\ proof (rule subset_Zorn_nonempty; simp) show \\S. on S\ apply (rule exI[of _ S]) using assms on_def by fastforce next fix C :: \'a set set\ assume \C \ {}\ assume \subset.chain (Collect on) C\ then have C_on: \B \ C \ on B\ and C_order: \B \ C \ B' \ C \ B \ B' \ B' \ B\ for B B' by (auto simp: subset.chain_def) have \is_orthogonal x y\ if \x\\C\ \y\\C\ \x \ y\ for x y by (smt (verit) UnionE C_order C_on on_def is_ortho_set_def subsetD that(1) that(2) that(3)) moreover have \0 \ \ C\ by (meson UnionE C_on is_ortho_set_def on_def) moreover have \\C \ S\ using C_on \C \ {}\ on_def by blast ultimately show \on (\ C)\ unfolding on_def is_ortho_set_def by simp qed then obtain B where \on B\ and B_max: \B' \ B \ on B' \ B=B'\ for B' by auto have \\ = 0\ if \ortho: \\b\B. is_orthogonal \ b\ for \ :: 'a proof (rule ccontr) assume \\ \ 0\ define \ B' where \\ = \ /\<^sub>R norm \\ and \B' = B \ {\}\ have [simp]: \norm \ = 1\ using \\ \ 0\ by (auto simp: \_def) have \ortho: \is_orthogonal \ b\ if \b \ B\ for b using \ortho that \_def by auto have orthoB': \is_orthogonal x y\ if \x\B'\ \y\B'\ \x \ y\ for x y using that \on B\ \ortho \ortho[THEN is_orthogonal_sym[THEN iffD1]] by (auto simp: B'_def on_def is_ortho_set_def) have B'0: \0 \ B'\ using B'_def \norm \ = 1\ \on B\ is_ortho_set_def on_def by fastforce have \S \ B'\ using B'_def \on B\ on_def by auto from orthoB' B'0 \S \ B'\ have \on B'\ by (simp add: on_def is_ortho_set_def) with B_max have \B = B'\ by (metis B'_def Un_upper1) then have \\ \ B\ using B'_def by blast then have \is_orthogonal \ \\ using \ortho by blast then show False using B'0 \B = B'\ \\ \ B\ by fastforce qed then have \orthogonal_complement B = {0}\ by (auto simp: orthogonal_complement_def) then have \UNIV = orthogonal_complement (orthogonal_complement B)\ by simp also have \\ = orthogonal_complement (orthogonal_complement (closure (cspan B)))\ by (metis (mono_tags, opaque_lifting) \orthogonal_complement B = {0}\ cinner_zero_left complex_vector.span_superset empty_iff insert_iff orthogonal_complementI orthogonal_complement_antimono orthogonal_complement_of_closure subsetI subset_antisym) also have \\ = closure (cspan B)\ apply (rule double_orthogonal_complement_id) by simp finally have \closure (cspan B) = UNIV\ by simp with \on B\ show ?thesis by (auto simp: on_def) qed lemma orthonormal_basis_exists: fixes S :: \'a::chilbert_space set\ assumes \is_ortho_set S\ and \\x. x\S \ norm x = 1\ shows \\B. B \ S \ is_onb B\ proof - from \is_ortho_set S\ obtain B where \is_ortho_set B\ and \B \ S\ and \closure (cspan B) = UNIV\ using ortho_basis_exists by blast define B' where \B' = (\x. x /\<^sub>R norm x) ` B\ have \S = (\x. x /\<^sub>R norm x) ` S\ by (simp add: assms(2)) then have \B' \ S\ using B'_def \S \ B\ by blast moreover have \ccspan B' = top\ apply (transfer fixing: B') apply (simp add: B'_def scaleR_scaleC) apply (subst complex_vector.span_image_scale') using \is_ortho_set B\ \closure (cspan B) = UNIV\ is_ortho_set_def by auto moreover have \is_ortho_set B'\ using \is_ortho_set B\ by (auto simp: B'_def is_ortho_set_def) moreover have \\b\B'. norm b = 1\ using \is_ortho_set B\ apply (auto simp: B'_def is_ortho_set_def) by (metis field_class.field_inverse norm_eq_zero) ultimately show ?thesis by (auto simp: is_onb_def) qed definition some_chilbert_basis :: \'a::chilbert_space set\ where \some_chilbert_basis = (SOME B::'a set. is_onb B)\ lemma is_onb_some_chilbert_basis[simp]: \is_onb (some_chilbert_basis :: 'a::chilbert_space set)\ using orthonormal_basis_exists[OF is_ortho_set_empty] by (auto simp add: some_chilbert_basis_def intro: someI2) lemma is_ortho_set_some_chilbert_basis[simp]: \is_ortho_set some_chilbert_basis\ using is_onb_def is_onb_some_chilbert_basis by blast lemma is_normal_some_chilbert_basis: \\x. x \ some_chilbert_basis \ norm x = 1\ using is_onb_def is_onb_some_chilbert_basis by blast lemma ccspan_some_chilbert_basis[simp]: \ccspan some_chilbert_basis = top\ using is_onb_def is_onb_some_chilbert_basis by blast lemma span_some_chilbert_basis[simp]: \closure (cspan some_chilbert_basis) = UNIV\ by (metis ccspan.rep_eq ccspan_some_chilbert_basis top_ccsubspace.rep_eq) lemma cindependent_some_chilbert_basis[simp]: \cindependent some_chilbert_basis\ using is_ortho_set_cindependent is_ortho_set_some_chilbert_basis by blast lemma finite_some_chilbert_basis[simp]: \finite (some_chilbert_basis :: 'a :: {chilbert_space, cfinite_dim} set)\ apply (rule cindependent_cfinite_dim_finite) by simp lemma some_chilbert_basis_nonempty: \(some_chilbert_basis :: 'a::{chilbert_space, not_singleton} set) \ {}\ proof (rule ccontr, simp) define B :: \'a set\ where \B = some_chilbert_basis\ assume [simp]: \B = {}\ have \UNIV = closure (cspan B)\ using B_def span_some_chilbert_basis by blast also have \\ = {0}\ by simp also have \\ \ UNIV\ using Extra_General.UNIV_not_singleton by blast finally show False by simp qed subsection \Riesz-representation theorem\ lemma orthogonal_complement_kernel_functional: fixes f :: \'a::complex_inner \ complex\ assumes \bounded_clinear f\ shows \\x. orthogonal_complement (f -` {0}) = cspan {x}\ proof (cases \orthogonal_complement (f -` {0}) = {0}\) case True then show ?thesis apply (rule_tac x=0 in exI) by auto next case False then obtain x where xortho: \x \ orthogonal_complement (f -` {0})\ and xnon0: \x \ 0\ using complex_vector.subspace_def by fastforce from xnon0 xortho have r1: \f x \ 0\ by (metis cinner_eq_zero_iff orthogonal_complement_orthoI vimage_singleton_eq) have \\ k. y = k *\<^sub>C x\ if \y \ orthogonal_complement (f -` {0})\ for y proof (cases \y = 0\) case True then show ?thesis by auto next case False with that have \f y \ 0\ by (metis cinner_eq_zero_iff orthogonal_complement_orthoI vimage_singleton_eq) then obtain k where k_def: \f x = k * f y\ by (metis add.inverse_inverse minus_divide_eq_eq) with assms have \f x = f (k *\<^sub>C y)\ by (simp add: bounded_clinear.axioms(1) clinear.scaleC) hence \f x - f (k *\<^sub>C y) = 0\ by simp with assms have s1: \f (x - k *\<^sub>C y) = 0\ by (simp add: bounded_clinear.axioms(1) complex_vector.linear_diff) from that have \k *\<^sub>C y \ orthogonal_complement (f -` {0})\ by (simp add: complex_vector.subspace_scale) with xortho have s2: \x - (k *\<^sub>C y) \ orthogonal_complement (f -` {0})\ by (simp add: complex_vector.subspace_diff) have s3: \(x - (k *\<^sub>C y)) \ f -` {0}\ using s1 by simp moreover have \(f -` {0}) \ (orthogonal_complement (f -` {0})) = {0}\ by (meson assms closed_csubspace_def complex_vector.subspace_def kernel_is_closed_csubspace orthogonal_complement_zero_intersection) ultimately have \x - (k *\<^sub>C y) = 0\ using s2 by blast thus ?thesis by (metis ceq_vector_fraction_iff eq_iff_diff_eq_0 k_def r1 scaleC_scaleC) qed then have \orthogonal_complement (f -` {0}) \ cspan {x}\ using complex_vector.span_superset complex_vector.subspace_scale by blast moreover from xortho have \orthogonal_complement (f -` {0}) \ cspan {x}\ by (simp add: complex_vector.span_minimal) ultimately show ?thesis by auto qed lemma riesz_frechet_representation_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \ complex\ assumes a1: \bounded_clinear f\ shows \\t. \x. f x = \t, x\\ proof(cases \\ x. f x = 0\) case True thus ?thesis by (metis cinner_zero_left) next case False obtain t where spant: \orthogonal_complement (f -` {0}) = cspan {t}\ using orthogonal_complement_kernel_functional using assms by blast have \projection (orthogonal_complement (f -` {0})) x = (\t , 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) 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\ 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 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 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) 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 less_eq_complex_def) ultimately have b5: \\ (A\<^sup>\) x \^2 \ \x\ * \A ((A\<^sup>\) x)\\ for x by (metis complex_of_real_mono_iff) have \\M. M \ 0 \ (\ x. \A ((A\<^sup>\) x)\ \ M * \(A\<^sup>\) x\)\ using a1 by (metis (mono_tags, opaque_lifting) bounded_clinear.bounded linear mult_nonneg_nonpos mult_zero_right norm_ge_zero order.trans semiring_normalization_rules(7)) then obtain M where q1: \M \ 0\ and q2: \\ x. \A ((A\<^sup>\) x)\ \ M * \(A\<^sup>\) x\\ by blast have \\ x::'b. \x\ \ 0\ by simp hence b6: \\x\ * \A ((A\<^sup>\) x)\ \ \x\ * M * \(A\<^sup>\) x\\ for x using q2 by (smt ordered_comm_semiring_class.comm_mult_left_mono vector_space_over_itself.scale_scale) have z3: \\ (A\<^sup>\) x \ \ \x\ * M\ for x proof(cases \\(A\<^sup>\) x\ = 0\) case True thus ?thesis by (simp add: \0 \ M\) next case False have \\ (A\<^sup>\) x \^2 \ \x\ * M * \(A\<^sup>\) x\\ by (smt b5 b6) thus ?thesis by (smt False mult_right_cancel mult_right_mono norm_ge_zero semiring_normalization_rules(29)) qed thus \\K. \x. \(A\<^sup>\) x\ \ \x\ * K\ by auto qed proposition double_cadjoint: fixes U :: \'a::chilbert_space \ 'b::complex_inner\ assumes a1: "bounded_clinear U" shows "U\<^sup>\\<^sup>\ = U" by (metis assms cadjoint_def cadjoint_is_cadjoint is_adjoint_sym is_cadjoint_unique someI_ex) lemma cadjoint_id: \(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) 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) 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\ \projection (orthogonal_complement x) t \ orthogonal_complement x\ \projection x t \ x\ \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 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 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) 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) 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 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 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) 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) 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)\ by (smt is_generator_set is_normal is_orthonormal) 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 hence x2: \(\ n. (\t\set canonical_basis. \ t, X n \ *\<^sub>C t)) \ l\ using l_def by blast have \X \ l\ using x1 x2 by simp 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 subsection \Misc (ctd.)\ lemma separating_dense_span: assumes \\F G :: 'a::chilbert_space \ 'b::{complex_normed_vector,not_singleton}. bounded_clinear F \ bounded_clinear G \ (\x\S. F x = G x) \ F = G\ shows \closure (cspan S) = UNIV\ proof - have \\ = 0\ if \\ \ orthogonal_complement S\ for \ proof - obtain \ :: 'b where \\ \ 0\ by fastforce have \(\x. cinner \ x *\<^sub>C \) = (\_. 0)\ apply (rule assms[rule_format]) using orthogonal_complement_orthoI that by (auto simp add: bounded_clinear_cinner_right bounded_clinear_scaleC_const) then have \cinner \ \ = 0\ by (meson \\ \ 0\ scaleC_eq_0_iff) then show \\ = 0\ by auto qed then have \orthogonal_complement (orthogonal_complement S) = UNIV\ by (metis UNIV_eq_I cinner_zero_right orthogonal_complementI) then show \closure (cspan S) = UNIV\ by (simp add: orthogonal_complement_orthogonal_complement_closure_cspan) qed instance conjugate_space :: (chilbert_space) chilbert_space.. end diff --git a/thys/Complex_Bounded_Operators/Complex_L2.thy b/thys/Complex_Bounded_Operators/Complex_L2.thy --- a/thys/Complex_Bounded_Operators/Complex_L2.thy +++ b/thys/Complex_Bounded_Operators/Complex_L2.thy @@ -1,1518 +1,1563 @@ section \\Complex_L2\ -- Hilbert space of square-summable functions\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_L2 imports Complex_Bounded_Linear_Function "HOL-Analysis.L2_Norm" "HOL-Library.Rewrite" "HOL-Analysis.Infinite_Sum" begin unbundle lattice_syntax unbundle cblinfun_notation unbundle no_notation_blinfun_apply subsection \l2 norm of functions\ definition "has_ell2_norm (x::_\complex) \ (\i. (x i)\<^sup>2) abs_summable_on UNIV" lemma has_ell2_norm_bdd_above: \has_ell2_norm x \ bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (simp add: has_ell2_norm_def abs_summable_iff_bdd_above) lemma has_ell2_norm_L2_set: "has_ell2_norm x = bdd_above (L2_set (norm o x) ` Collect finite)" proof (rule iffI) have \mono sqrt\ using monoI real_sqrt_le_mono by blast assume \has_ell2_norm x\ then have *: \bdd_above (sum (\xa. norm ((x xa)\<^sup>2)) ` Collect finite)\ by (subst (asm) has_ell2_norm_bdd_above) have \bdd_above ((\F. sqrt (sum (\xa. norm ((x xa)\<^sup>2)) F)) ` Collect finite)\ using bdd_above_image_mono[OF \mono sqrt\ *] by (auto simp: image_image) then show \bdd_above (L2_set (norm o x) ` Collect finite)\ by (auto simp: L2_set_def norm_power) next define p2 where \p2 x = (if x < 0 then 0 else x^2)\ for x :: real have \mono p2\ by (simp add: monoI p2_def) have [simp]: \p2 (L2_set f F) = (\i\F. (f i)\<^sup>2)\ for f and F :: \'a set\ by (smt (verit) L2_set_def L2_set_nonneg p2_def power2_less_0 real_sqrt_pow2 sum.cong sum_nonneg) assume *: \bdd_above (L2_set (norm o x) ` Collect finite)\ have \bdd_above (p2 ` L2_set (norm o x) ` Collect finite)\ using bdd_above_image_mono[OF \mono p2\ *] by auto then show \has_ell2_norm x\ apply (simp add: image_image has_ell2_norm_def abs_summable_iff_bdd_above) by (simp add: norm_power) qed definition ell2_norm :: \('a \ complex) \ real\ where \ell2_norm x = sqrt (\\<^sub>\i. norm (x i)^2)\ lemma ell2_norm_SUP: assumes \has_ell2_norm x\ shows "ell2_norm x = sqrt (SUP F\{F. finite F}. sum (\i. norm (x i)^2) F)" using assms apply (auto simp add: ell2_norm_def has_ell2_norm_def) apply (subst infsum_nonneg_is_SUPREMUM_real) by (auto simp: norm_power) lemma ell2_norm_L2_set: assumes "has_ell2_norm x" shows "ell2_norm x = (SUP F\{F. finite F}. L2_set (norm o x) F)" proof- have "sqrt (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)) = (SUP F\{F. finite F}. sqrt (\i\F. (cmod (x i))\<^sup>2))" proof (subst continuous_at_Sup_mono) show "mono sqrt" by (simp add: mono_def) show "continuous (at_left (\ (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite))) sqrt" using continuous_at_split isCont_real_sqrt by blast show "sum (\i. (cmod (x i))\<^sup>2) ` Collect finite \ {}" by auto show "bdd_above (sum (\i. (cmod (x i))\<^sup>2) ` Collect finite)" using has_ell2_norm_bdd_above[THEN iffD1, OF assms] by (auto simp: norm_power) show "\ (sqrt ` sum (\i. (cmod (x i))\<^sup>2) ` Collect finite) = (SUP F\Collect finite. sqrt (\i\F. (cmod (x i))\<^sup>2))" by (metis image_image) qed thus ?thesis using assms by (auto simp: ell2_norm_SUP L2_set_def) qed lemma has_ell2_norm_finite[simp]: "has_ell2_norm (x::'a::finite\_)" unfolding has_ell2_norm_def by simp lemma ell2_norm_finite: "ell2_norm (x::'a::finite\complex) = sqrt (sum (\i. (norm(x i))^2) UNIV)" by (simp add: ell2_norm_def) lemma ell2_norm_finite_L2_set: "ell2_norm (x::'a::finite\complex) = L2_set (norm o x) UNIV" by (simp add: ell2_norm_finite L2_set_def) lemma ell2_norm_square: \(ell2_norm x)\<^sup>2 = (\\<^sub>\i. (cmod (x i))\<^sup>2)\ unfolding ell2_norm_def apply (subst real_sqrt_pow2) by (simp_all add: infsum_nonneg) lemma ell2_ket: fixes a defines \f \ (\i. if a = i then 1 else 0)\ shows has_ell2_norm_ket: \has_ell2_norm f\ and ell2_norm_ket: \ell2_norm f = 1\ proof - have \(\x. (f x)\<^sup>2) abs_summable_on {a}\ apply (rule summable_on_finite) by simp then show \has_ell2_norm f\ unfolding has_ell2_norm_def apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1]) unfolding f_def by auto have \(\\<^sub>\x\{a}. (f x)\<^sup>2) = 1\ apply (subst infsum_finite) by (auto simp: f_def) then show \ell2_norm f = 1\ unfolding ell2_norm_def apply (subst infsum_cong_neutral[where T=\{a}\ and g=\\x. (cmod (f x))\<^sup>2\]) by (auto simp: f_def) qed lemma ell2_norm_geq0: \ell2_norm x \ 0\ by (auto simp: ell2_norm_def intro!: infsum_nonneg) lemma ell2_norm_point_bound: assumes \has_ell2_norm x\ shows \ell2_norm x \ cmod (x i)\ proof - have \(cmod (x i))\<^sup>2 = norm ((x i)\<^sup>2)\ by (simp add: norm_power) also have \norm ((x i)\<^sup>2) = sum (\i. (norm ((x i)\<^sup>2))) {i}\ by auto also have \\ = infsum (\i. (norm ((x i)\<^sup>2))) {i}\ by (rule infsum_finite[symmetric], simp) also have \\ \ infsum (\i. (norm ((x i)\<^sup>2))) UNIV\ apply (rule infsum_mono_neutral) using assms by (auto simp: has_ell2_norm_def) also have \\ = (ell2_norm x)\<^sup>2\ by (metis (no_types, lifting) ell2_norm_def ell2_norm_geq0 infsum_cong norm_power real_sqrt_eq_iff real_sqrt_unique) finally show ?thesis using ell2_norm_geq0 power2_le_imp_le by blast qed lemma ell2_norm_0: assumes "has_ell2_norm x" shows "(ell2_norm x = 0) = (x = (\_. 0))" proof assume u1: "x = (\_. 0)" have u2: "(SUP x::'a set\Collect finite. (0::real)) = 0" if "x = (\_. 0)" by (metis cSUP_const empty_Collect_eq finite.emptyI) show "ell2_norm x = 0" unfolding ell2_norm_def using u1 u2 by auto next assume norm0: "ell2_norm x = 0" show "x = (\_. 0)" proof fix i have \cmod (x i) \ ell2_norm x\ using assms by (rule ell2_norm_point_bound) also have \\ = 0\ by (fact norm0) finally show "x i = 0" by auto qed qed lemma ell2_norm_smult: assumes "has_ell2_norm x" shows "has_ell2_norm (\i. c * x i)" and "ell2_norm (\i. c * x i) = cmod c * ell2_norm x" proof - have L2_set_mul: "L2_set (cmod \ (\i. c * x i)) F = cmod c * L2_set (cmod \ x) F" for F proof- have "L2_set (cmod \ (\i. c * x i)) F = L2_set (\i. (cmod c * (cmod o x) i)) F" by (metis comp_def norm_mult) also have "\ = cmod c * L2_set (cmod o x) F" by (metis norm_ge_zero L2_set_right_distrib) finally show ?thesis . qed from assms obtain M where M: "M \ L2_set (cmod o x) F" if "finite F" for F unfolding has_ell2_norm_L2_set bdd_above_def by auto hence "cmod c * M \ L2_set (cmod o (\i. c * x i)) F" if "finite F" for F unfolding L2_set_mul by (simp add: ordered_comm_semiring_class.comm_mult_left_mono that) thus has: "has_ell2_norm (\i. c * x i)" unfolding has_ell2_norm_L2_set bdd_above_def using L2_set_mul[symmetric] by auto have "ell2_norm (\i. c * x i) = (SUP F \ Collect finite. (L2_set (cmod \ (\i. c * x i)) F))" by (simp add: ell2_norm_L2_set has) also have "\ = (SUP F \ Collect finite. (cmod c * L2_set (cmod \ x) F))" using L2_set_mul by auto also have "\ = cmod c * ell2_norm x" proof (subst ell2_norm_L2_set) show "has_ell2_norm x" by (simp add: assms) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = cmod c * \ (L2_set (cmod \ x) ` Collect finite)" proof (subst continuous_at_Sup_mono [where f = "\x. cmod c * x"]) show "mono ((*) (cmod c))" by (simp add: mono_def ordered_comm_semiring_class.comm_mult_left_mono) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) ((*) (cmod c))" proof (rule continuous_mult) show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. cmod c)" by simp show "continuous (at_left (\ (L2_set (cmod \ x) ` Collect finite))) (\x. x)" by simp qed show "L2_set (cmod \ x) ` Collect finite \ {}" by auto show "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms has_ell2_norm_L2_set) show "(SUP F\Collect finite. cmod c * L2_set (cmod \ x) F) = \ ((*) (cmod c) ` L2_set (cmod \ x) ` Collect finite)" by (metis image_image) qed qed finally show "ell2_norm (\i. c * x i) = cmod c * ell2_norm x". qed lemma ell2_norm_triangle: assumes "has_ell2_norm x" and "has_ell2_norm y" shows "has_ell2_norm (\i. x i + y i)" and "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" proof - have triangle: "L2_set (cmod \ (\i. x i + y i)) F \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" (is "?lhs\?rhs") if "finite F" for F proof - have "?lhs \ L2_set (\i. (cmod o x) i + (cmod o y) i) F" proof (rule L2_set_mono) show "(cmod \ (\i. x i + y i)) i \ (cmod \ x) i + (cmod \ y) i" if "i \ F" for i :: 'a using that norm_triangle_ineq by auto show "0 \ (cmod \ (\i. x i + y i)) i" if "i \ F" for i :: 'a using that by simp qed also have "\ \ ?rhs" by (rule L2_set_triangle_ineq) finally show ?thesis . qed obtain Mx My where Mx: "Mx \ L2_set (cmod o x) F" and My: "My \ L2_set (cmod o y) F" if "finite F" for F using assms unfolding has_ell2_norm_L2_set bdd_above_def by auto hence MxMy: "Mx + My \ L2_set (cmod \ x) F + L2_set (cmod \ y) F" if "finite F" for F using that by fastforce hence bdd_plus: "bdd_above ((\xa. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa) ` Collect finite)" unfolding bdd_above_def by auto from MxMy have MxMy': "Mx + My \ L2_set (cmod \ (\i. x i + y i)) F" if "finite F" for F using triangle that by fastforce thus has: "has_ell2_norm (\i. x i + y i)" unfolding has_ell2_norm_L2_set bdd_above_def by auto have SUP_plus: "(SUP x\A. f x + g x) \ (SUP x\A. f x) + (SUP x\A. g x)" if notempty: "A\{}" and bddf: "bdd_above (f`A)"and bddg: "bdd_above (g`A)" for f g :: "'a set \ real" and A proof- have xleq: "x \ (SUP x\A. f x) + (SUP x\A. g x)" if x: "x \ (\x. f x + g x) ` A" for x proof - obtain a where aA: "a:A" and ax: "x = f a + g a" using x by blast have fa: "f a \ (SUP x\A. f x)" by (simp add: bddf aA cSUP_upper) moreover have "g a \ (SUP x\A. g x)" by (simp add: bddg aA cSUP_upper) ultimately have "f a + g a \ (SUP x\A. f x) + (SUP x\A. g x)" by simp with ax show ?thesis by simp qed have "(\x. f x + g x) ` A \ {}" using notempty by auto moreover have "x \ \ (f ` A) + \ (g ` A)" if "x \ (\x. f x + g x) ` A" for x :: real using that by (simp add: xleq) ultimately show ?thesis by (meson bdd_above_def cSup_le_iff) qed have a2: "bdd_above (L2_set (cmod \ x) ` Collect finite)" by (meson assms(1) has_ell2_norm_L2_set) have a3: "bdd_above (L2_set (cmod \ y) ` Collect finite)" by (meson assms(2) has_ell2_norm_L2_set) have a1: "Collect finite \ {}" by auto have a4: "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ (SUP xa\Collect finite. L2_set (cmod \ x) xa + L2_set (cmod \ y) xa)" by (metis (mono_tags, lifting) a1 bdd_plus cSUP_mono mem_Collect_eq triangle) have "\r. \ (L2_set (cmod \ (\a. x a + y a)) ` Collect finite) \ r \ \ (SUP A\Collect finite. L2_set (cmod \ x) A + L2_set (cmod \ y) A) \ r" using a4 by linarith hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ \ (L2_set (cmod \ x) ` Collect finite) + \ (L2_set (cmod \ y) ` Collect finite)" by (metis (no_types) SUP_plus a1 a2 a3) hence "\ (L2_set (cmod \ (\i. x i + y i)) ` Collect finite) \ ell2_norm x + ell2_norm y" by (simp add: assms(1) assms(2) ell2_norm_L2_set) thus "ell2_norm (\i. x i + y i) \ ell2_norm x + ell2_norm y" by (simp add: ell2_norm_L2_set has) qed lemma ell2_norm_uminus: assumes "has_ell2_norm x" shows \has_ell2_norm (\i. - x i)\ and \ell2_norm (\i. - x i) = ell2_norm x\ using assms by (auto simp: has_ell2_norm_def ell2_norm_def) subsection \The type \ell2\ of square-summable functions\ typedef 'a ell2 = "{x::'a\complex. has_ell2_norm x}" unfolding has_ell2_norm_def by (rule exI[of _ "\_.0"], auto) setup_lifting type_definition_ell2 instantiation ell2 :: (type)complex_vector begin lift_definition zero_ell2 :: "'a ell2" is "\_. 0" by (auto simp: has_ell2_norm_def) lift_definition uminus_ell2 :: "'a ell2 \ 'a ell2" is uminus by (simp add: has_ell2_norm_def) lift_definition plus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x + g x" by (rule ell2_norm_triangle) lift_definition minus_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\f g x. f x - g x" apply (subst add_uminus_conv_diff[symmetric]) apply (rule ell2_norm_triangle) by (auto simp add: ell2_norm_uminus) lift_definition scaleR_ell2 :: "real \ 'a ell2 \ 'a ell2" is "\r f x. complex_of_real r * f x" by (rule ell2_norm_smult) lift_definition scaleC_ell2 :: "complex \ 'a ell2 \ 'a ell2" is "\c f x. c * f x" by (rule ell2_norm_smult) instance proof fix a b c :: "'a ell2" show "((*\<^sub>R) r::'a ell2 \ _) = (*\<^sub>C) (complex_of_real r)" for r apply (rule ext) apply transfer by auto show "a + b + c = a + (b + c)" by (transfer; rule ext; simp) show "a + b = b + a" by (transfer; rule ext; simp) show "0 + a = a" by (transfer; rule ext; simp) show "- a + a = 0" by (transfer; rule ext; simp) show "a - b = a + - b" by (transfer; rule ext; simp) show "r *\<^sub>C (a + b) = r *\<^sub>C a + r *\<^sub>C b" for r apply (transfer; rule ext) by (simp add: vector_space_over_itself.scale_right_distrib) show "(r + r') *\<^sub>C a = r *\<^sub>C a + r' *\<^sub>C a" for r r' apply (transfer; rule ext) by (simp add: ring_class.ring_distribs(2)) show "r *\<^sub>C r' *\<^sub>C a = (r * r') *\<^sub>C a" for r r' by (transfer; rule ext; simp) show "1 *\<^sub>C a = a" by (transfer; rule ext; simp) qed end instantiation ell2 :: (type)complex_normed_vector begin lift_definition norm_ell2 :: "'a ell2 \ real" is ell2_norm . declare norm_ell2_def[code del] definition "dist x y = norm (x - y)" for x y::"'a ell2" definition "sgn x = x /\<^sub>R norm x" for x::"'a ell2" definition [code del]: "uniformity = (INF e\{0<..}. principal {(x::'a ell2, y). norm (x - y) < e})" definition [code del]: "open U = (\x\U. \\<^sub>F (x', y) in INF e\{0<..}. principal {(x, y). norm (x - y) < e}. x' = x \ y \ U)" for U :: "'a ell2 set" instance proof fix a b :: "'a ell2" show "dist a b = norm (a - b)" by (simp add: dist_ell2_def) show "sgn a = a /\<^sub>R norm a" by (simp add: sgn_ell2_def) show "uniformity = (INF e\{0<..}. principal {(x, y). dist (x::'a ell2) y < e})" unfolding dist_ell2_def uniformity_ell2_def by simp show "open U = (\x\U. \\<^sub>F (x', y) in uniformity. (x'::'a ell2) = x \ y \ U)" for U :: "'a ell2 set" unfolding uniformity_ell2_def open_ell2_def by simp_all show "(norm a = 0) = (a = 0)" apply transfer by (fact ell2_norm_0) show "norm (a + b) \ norm a + norm b" apply transfer by (fact ell2_norm_triangle) show "norm (r *\<^sub>R (a::'a ell2)) = \r\ * norm a" for r and a :: "'a ell2" apply transfer by (simp add: ell2_norm_smult(2)) show "norm (r *\<^sub>C a) = cmod r * norm a" for r apply transfer by (simp add: ell2_norm_smult(2)) qed end lemma norm_point_bound_ell2: "norm (Rep_ell2 x i) \ norm x" apply transfer by (simp add: ell2_norm_point_bound) lemma ell2_norm_finite_support: assumes \finite S\ \\ i. i \ S \ Rep_ell2 x i = 0\ shows \norm x = sqrt ((sum (\i. (cmod (Rep_ell2 x i))\<^sup>2)) S)\ proof (insert assms(2), transfer fixing: S) fix x :: \'a \ complex\ assume zero: \\i. i \ S \ x i = 0\ have \ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)\ by (auto simp: ell2_norm_def) also have \\ = sqrt (\\<^sub>\i\S. (cmod (x i))\<^sup>2)\ apply (subst infsum_cong_neutral[where g=\\i. (cmod (x i))\<^sup>2\ and S=UNIV and T=S]) using zero by auto also have \\ = sqrt (\i\S. (cmod (x i))\<^sup>2)\ using \finite S\ by simp finally show \ell2_norm x = sqrt (\i\S. (cmod (x i))\<^sup>2)\ by - qed instantiation ell2 :: (type) complex_inner begin lift_definition cinner_ell2 :: "'a ell2 \ 'a ell2 \ complex" is "\x y. infsum (\i. (cnj (x i) * y i)) UNIV" . declare cinner_ell2_def[code del] instance proof standard fix x y z :: "'a ell2" fix c :: complex show "cinner x y = cnj (cinner y x)" proof transfer fix x y :: "'a\complex" assume "has_ell2_norm x" and "has_ell2_norm y" have "(\\<^sub>\i. cnj (x i) * y i) = (\\<^sub>\i. cnj (cnj (y i) * x i))" by (metis complex_cnj_cnj complex_cnj_mult mult.commute) also have "\ = cnj (\\<^sub>\i. cnj (y i) * x i)" by (metis infsum_cnj) finally show "(\\<^sub>\i. cnj (x i) * y i) = cnj (\\<^sub>\i. cnj (y i) * x i)" . qed show "cinner (x + y) z = cinner x z + cinner y z" proof transfer fix x y z :: "'a \ complex" assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence cnj_y: "(\i. cnj (y i) * cnj (y i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm z" hence z: "(\i. z i * z i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_z:"(\i. cnj (x i) * z i) abs_summable_on UNIV" using cnj_x z by (rule abs_summable_product) have cnj_y_z:"(\i. cnj (y i) * z i) abs_summable_on UNIV" using cnj_y z by (rule abs_summable_product) show "(\\<^sub>\i. cnj (x i + y i) * z i) = (\\<^sub>\i. cnj (x i) * z i) + (\\<^sub>\i. cnj (y i) * z i)" apply (subst infsum_add [symmetric]) using cnj_x_z cnj_y_z by (auto simp add: summable_on_iff_abs_summable_on_complex distrib_left mult.commute) qed show "cinner (c *\<^sub>C x) y = cnj c * cinner x y" proof transfer fix x y :: "'a \ complex" and c :: complex assume "has_ell2_norm x" hence cnj_x: "(\i. cnj (x i) * cnj (x i)) abs_summable_on UNIV" by (simp del: complex_cnj_mult add: norm_mult[symmetric] complex_cnj_mult[symmetric] has_ell2_norm_def power2_eq_square) assume "has_ell2_norm y" hence y: "(\i. y i * y i) abs_summable_on UNIV" by (simp add: norm_mult[symmetric] has_ell2_norm_def power2_eq_square) have cnj_x_y:"(\i. cnj (x i) * y i) abs_summable_on UNIV" using cnj_x y by (rule abs_summable_product) thus "(\\<^sub>\i. cnj (c * x i) * y i) = cnj c * (\\<^sub>\i. cnj (x i) * y i)" by (auto simp flip: infsum_cmult_right simp add: abs_summable_summable mult.commute vector_space_over_itself.scale_left_commute) qed show "0 \ cinner x x" proof transfer fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence "(\i. cnj (x i) * x i) abs_summable_on UNIV" by auto hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square. have "0 = (\\<^sub>\i::'a. 0)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_complex) by (auto simp add: abs_summable_summable sum) finally show "0 \ (\\<^sub>\i. cnj (x i) * x i)" by assumption qed show "(cinner x x = 0) = (x = 0)" proof (transfer, auto) fix x :: "'a \ complex" assume "has_ell2_norm x" hence "(\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (smt (verit, del_insts) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_ge_zero norm_power real_norm_def summable_on_cong) hence cmod_x2: "(\i. cnj (x i) * x i) abs_summable_on UNIV" unfolding has_ell2_norm_def power2_eq_square by simp assume eq0: "(\\<^sub>\i. cnj (x i) * x i) = 0" show "x = (\_. 0)" proof (rule ccontr) assume "x \ (\_. 0)" then obtain i where "x i \ 0" by auto hence "0 < cnj (x i) * x i" by (metis le_less cnj_x_x_geq0 complex_cnj_zero_iff vector_space_over_itself.scale_eq_0_iff) also have "\ = (\\<^sub>\i\{i}. cnj (x i) * x i)" by auto also have "\ \ (\\<^sub>\i. cnj (x i) * x i)" apply (rule infsum_mono_neutral_complex) by (auto simp add: abs_summable_summable cmod_x2) also from eq0 have "\ = 0" by assumption finally show False by simp qed qed show "norm x = sqrt (cmod (cinner x x))" proof transfer fix x :: "'a \ complex" assume x: "has_ell2_norm x" have "(\i::'a. cmod (x i) * cmod (x i)) abs_summable_on UNIV \ (\i::'a. cmod (cnj (x i) * x i)) abs_summable_on UNIV" by (simp add: norm_mult has_ell2_norm_def power2_eq_square) hence sum: "(\i. cnj (x i) * x i) abs_summable_on UNIV" by (metis (no_types, lifting) complex_mod_mult_cnj has_ell2_norm_def mult.commute norm_power summable_on_cong x) from x have "ell2_norm x = sqrt (\\<^sub>\i. (cmod (x i))\<^sup>2)" unfolding ell2_norm_def by simp also have "\ = sqrt (\\<^sub>\i. cmod (cnj (x i) * x i))" unfolding norm_complex_def power2_eq_square by auto also have "\ = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by (auto simp: infsum_cmod abs_summable_summable sum) finally show "ell2_norm x = sqrt (cmod (\\<^sub>\i. cnj (x i) * x i))" by assumption qed qed end instance ell2 :: (type) chilbert_space proof fix X :: \nat \ 'a ell2\ define x where \x n a = Rep_ell2 (X n) a\ for n a have [simp]: \has_ell2_norm (x n)\ for n using Rep_ell2 x_def[abs_def] by simp assume \Cauchy X\ moreover have "dist (x n a) (x m a) \ dist (X n) (X m)" for n m a by (metis Rep_ell2 x_def dist_norm ell2_norm_point_bound mem_Collect_eq minus_ell2.rep_eq norm_ell2.rep_eq) ultimately have \Cauchy (\n. x n a)\ for a by (meson Cauchy_def le_less_trans) then obtain l where x_lim: \(\n. x n a) \ l a\ for a apply atomize_elim apply (rule choice) by (simp add: convergent_eq_Cauchy) define L where \L = Abs_ell2 l\ define normF where \normF F x = L2_set (cmod \ x) F\ for F :: \'a set\ and x have normF_triangle: \normF F (\a. x a + y a) \ normF F x + normF F y\ if \finite F\ for F x y proof - have \normF F (\a. x a + y a) = L2_set (\a. cmod (x a + y a)) F\ by (metis (mono_tags, lifting) L2_set_cong comp_apply normF_def) also have \\ \ L2_set (\a. cmod (x a) + cmod (y a)) F\ by (meson L2_set_mono norm_ge_zero norm_triangle_ineq) also have \\ \ L2_set (\a. cmod (x a)) F + L2_set (\a. cmod (y a)) F\ by (simp add: L2_set_triangle_ineq) also have \\ \ normF F x + normF F y\ by (smt (verit, best) L2_set_cong normF_def comp_apply) finally show ?thesis by - qed have normF_negate: \normF F (\a. - x a) = normF F x\ if \finite F\ for F x unfolding normF_def o_def by simp have normF_ell2norm: \normF F x \ ell2_norm x\ if \finite F\ and \has_ell2_norm x\ for F x apply (auto intro!: cSUP_upper2[where x=F] simp: that normF_def ell2_norm_L2_set) by (meson has_ell2_norm_L2_set that(2)) note Lim_bounded2[rotated, rule_format, trans] from \Cauchy X\ obtain I where cauchyX: \norm (X n - X m) \ \\ if \\>0\ \n\I \\ \m\I \\ for \ n m by (metis Cauchy_def dist_norm less_eq_real_def) have normF_xx: \normF F (\a. x n a - x m a) \ \\ if \finite F\ \\>0\ \n\I \\ \m\I \\ for \ n m F apply (subst asm_rl[of \(\a. x n a - x m a) = Rep_ell2 (X n - X m)\]) apply (simp add: x_def minus_ell2.rep_eq) using that cauchyX by (metis Rep_ell2 mem_Collect_eq normF_ell2norm norm_ell2.rep_eq order_trans) have normF_xl_lim: \(\m. normF F (\a. x m a - l a)) \ 0\ if \finite F\ for F proof - have \(\xa. cmod (x xa m - l m)) \ 0\ for m using x_lim by (simp add: LIM_zero_iff tendsto_norm_zero) then have \(\m. \i\F. ((cmod \ (\a. x m a - l a)) i)\<^sup>2) \ 0\ by (auto intro: tendsto_null_sum) then show ?thesis unfolding normF_def L2_set_def using tendsto_real_sqrt by force qed have normF_xl: \normF F (\a. x n a - l a) \ \\ if \n \ I \\ and \\ > 0\ and \finite F\ for n \ F proof - have \normF F (\a. x n a - l a) - \ \ normF F (\a. x n a - x m a) + normF F (\a. x m a - l a) - \\ for m using normF_triangle[OF \finite F\, where x=\(\a. x n a - x m a)\ and y=\(\a. x m a - l a)\] by auto also have \\ m \ normF F (\a. x m a - l a)\ if \m \ I \\ for m using normF_xx[OF \finite F\ \\>0\ \n \ I \\ \m \ I \\] by auto also have \(\m. \ m) \ 0\ using \finite F\ by (rule normF_xl_lim) finally show ?thesis by auto qed have \normF F l \ 1 + normF F (x (I 1))\ if [simp]: \finite F\ for F using normF_xl[where F=F and \=1 and n=\I 1\] using normF_triangle[where F=F and x=\x (I 1)\ and y=\\a. l a - x (I 1) a\] using normF_negate[where F=F and x=\(\a. x (I 1) a - l a)\] by auto also have \\ F \ 1 + ell2_norm (x (I 1))\ if \finite F\ for F using normF_ell2norm that by simp finally have [simp]: \has_ell2_norm l\ unfolding has_ell2_norm_L2_set by (auto intro!: bdd_aboveI simp flip: normF_def) then have \l = Rep_ell2 L\ by (simp add: Abs_ell2_inverse L_def) have [simp]: \has_ell2_norm (\a. x n a - l a)\ for n apply (subst diff_conv_add_uminus) apply (rule ell2_norm_triangle) by (auto intro!: ell2_norm_uminus) from normF_xl have ell2norm_xl: \ell2_norm (\a. x n a - l a) \ \\ if \n \ I \\ and \\ > 0\ for n \ apply (subst ell2_norm_L2_set) using that by (auto intro!: cSUP_least simp: normF_def) have \norm (X n - L) \ \\ if \n \ I \\ and \\ > 0\ for n \ using ell2norm_xl[OF that] by (simp add: x_def norm_ell2.rep_eq \l = Rep_ell2 L\ minus_ell2.rep_eq) then have \X \ L\ unfolding tendsto_iff apply (auto simp: dist_norm eventually_sequentially) by (meson field_lbound_gt_zero le_less_trans) then show \convergent X\ by (rule convergentI) qed instantiation ell2 :: (CARD_1) complex_algebra_1 begin lift_definition one_ell2 :: "'a ell2" is "\_. 1" by simp lift_definition times_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x * b x" by simp instance proof fix a b c :: "'a ell2" and r :: complex show "a * b * c = a * (b * c)" by (transfer, auto) show "(a + b) * c = a * c + b * c" apply (transfer, rule ext) by (simp add: distrib_left mult.commute) show "a * (b + c) = a * b + a * c" apply transfer by (simp add: ring_class.ring_distribs(1)) show "r *\<^sub>C a * b = r *\<^sub>C (a * b)" by (transfer, auto) show "(a::'a ell2) * r *\<^sub>C b = r *\<^sub>C (a * b)" by (transfer, auto) show "1 * a = a" by (transfer, rule ext, auto) show "a * 1 = a" by (transfer, rule ext, auto) show "(0::'a ell2) \ 1" apply transfer by (meson zero_neq_one) qed end instantiation ell2 :: (CARD_1) field begin lift_definition divide_ell2 :: "'a ell2 \ 'a ell2 \ 'a ell2" is "\a b x. a x / b x" by simp lift_definition inverse_ell2 :: "'a ell2 \ 'a ell2" is "\a x. inverse (a x)" by simp instance proof (intro_classes; transfer) fix a :: "'a \ complex" assume "a \ (\_. 0)" then obtain y where ay: "a y \ 0" by auto show "(\x. inverse (a x) * a x) = (\_. 1)" proof (rule ext) fix x have "x = y" by auto with ay have "a x \ 0" by metis then show "inverse (a x) * a x = 1" by auto qed qed (auto simp add: divide_complex_def mult.commute ring_class.ring_distribs) end +lemma sum_ell2_transfer[transfer_rule]: + includes lifting_syntax + shows \(((=) ===> pcr_ell2 (=)) ===> rel_set (=) ===> pcr_ell2 (=)) + (\f X x. sum (\y. f y x) X) sum\ +proof (intro rel_funI, rename_tac f f' X X') + fix f and f' :: \'a \ 'b ell2\ + assume [transfer_rule]: \((=) ===> pcr_ell2 (=)) f f'\ + fix X X' :: \'a set\ + assume \rel_set (=) X X'\ + then have [simp]: \X' = X\ + by (simp add: rel_set_eq) + show \pcr_ell2 (=) (\x. \y\X. f y x) (sum f' X')\ + unfolding \X' = X\ + proof (induction X rule: infinite_finite_induct) + case (infinite X) + show ?case + apply (simp add: infinite) + by transfer_prover + next + case empty + show ?case + apply (simp add: empty) + by transfer_prover + next + case (insert x F) + note [transfer_rule] = insert.IH + show ?case + apply (simp add: insert) + by transfer_prover + qed +qed + subsection \Orthogonality\ lemma ell2_pointwise_ortho: assumes \\ i. Rep_ell2 x i = 0 \ Rep_ell2 y i = 0\ shows \is_orthogonal x y\ using assms apply transfer by (simp add: infsum_0) subsection \Truncated vectors\ lift_definition trunc_ell2:: \'a set \ 'a ell2 \ 'a ell2\ is \\ S x. (\ i. (if i \ S then x i else 0))\ proof (rename_tac S x) fix x :: \'a \ complex\ and S :: \'a set\ assume \has_ell2_norm x\ then have \(\i. (x i)\<^sup>2) abs_summable_on UNIV\ unfolding has_ell2_norm_def by - then have \(\i. (x i)\<^sup>2) abs_summable_on S\ using summable_on_subset_banach by blast then have \(\xa. (if xa \ S then x xa else 0)\<^sup>2) abs_summable_on UNIV\ apply (rule summable_on_cong_neutral[THEN iffD1, rotated -1]) by auto then show \has_ell2_norm (\i. if i \ S then x i else 0)\ unfolding has_ell2_norm_def by - qed lemma trunc_ell2_empty[simp]: \trunc_ell2 {} x = 0\ apply transfer by simp lemma trunc_ell2_UNIV[simp]: \trunc_ell2 UNIV \ = \\ apply transfer by simp lemma norm_id_minus_trunc_ell2: \(norm (x - trunc_ell2 S x))^2 = (norm x)^2 - (norm (trunc_ell2 S x))^2\ proof- have \Rep_ell2 (trunc_ell2 S x) i = 0 \ Rep_ell2 (x - trunc_ell2 S x) i = 0\ for i apply transfer by auto hence \\ (trunc_ell2 S x), (x - trunc_ell2 S x) \ = 0\ using ell2_pointwise_ortho by blast hence \(norm x)^2 = (norm (trunc_ell2 S x))^2 + (norm (x - trunc_ell2 S x))^2\ using pythagorean_theorem by fastforce thus ?thesis by simp qed lemma norm_trunc_ell2_finite: \finite S \ (norm (trunc_ell2 S x)) = sqrt ((sum (\i. (cmod (Rep_ell2 x i))\<^sup>2)) S)\ proof- assume \finite S\ moreover have \\ i. i \ S \ Rep_ell2 ((trunc_ell2 S x)) i = 0\ by (simp add: trunc_ell2.rep_eq) ultimately have \(norm (trunc_ell2 S x)) = sqrt ((sum (\i. (cmod (Rep_ell2 ((trunc_ell2 S x)) i))\<^sup>2)) S)\ using ell2_norm_finite_support by blast moreover have \\ i. i \ S \ Rep_ell2 ((trunc_ell2 S x)) i = Rep_ell2 x i\ by (simp add: trunc_ell2.rep_eq) ultimately show ?thesis by simp qed lemma trunc_ell2_lim_at_UNIV: \((\S. trunc_ell2 S \) \ \) (finite_subsets_at_top UNIV)\ proof - define f where \f i = (cmod (Rep_ell2 \ i))\<^sup>2\ for i have has: \has_ell2_norm (Rep_ell2 \)\ using Rep_ell2 by blast then have summable: "f abs_summable_on UNIV" by (smt (verit, del_insts) f_def has_ell2_norm_def norm_ge_zero norm_power real_norm_def summable_on_cong) have \norm \ = (ell2_norm (Rep_ell2 \))\ apply transfer by simp also have \\ = sqrt (infsum f UNIV)\ by (simp add: ell2_norm_def f_def[symmetric]) finally have norm\: \norm \ = sqrt (infsum f UNIV)\ by - have norm_trunc: \norm (trunc_ell2 S \) = sqrt (sum f S)\ if \finite S\ for S using f_def that norm_trunc_ell2_finite by fastforce have \(sum f \ infsum f UNIV) (finite_subsets_at_top UNIV)\ using f_def[abs_def] infsum_tendsto local.summable by fastforce then have \((\S. sqrt (sum f S)) \ sqrt (infsum f UNIV)) (finite_subsets_at_top UNIV)\ using tendsto_real_sqrt by blast then have \((\S. norm (trunc_ell2 S \)) \ norm \) (finite_subsets_at_top UNIV)\ apply (subst tendsto_cong[where g=\\S. sqrt (sum f S)\]) by (auto simp add: eventually_finite_subsets_at_top_weakI norm_trunc norm\) then have \((\S. (norm (trunc_ell2 S \))\<^sup>2) \ (norm \)\<^sup>2) (finite_subsets_at_top UNIV)\ by (simp add: tendsto_power) then have \((\S. (norm \)\<^sup>2 - (norm (trunc_ell2 S \))\<^sup>2) \ 0) (finite_subsets_at_top UNIV)\ apply (rule tendsto_diff[where a=\(norm \)^2\ and b=\(norm \)^2\, simplified, rotated]) by auto then have \((\S. (norm (\ - trunc_ell2 S \))\<^sup>2) \ 0) (finite_subsets_at_top UNIV)\ unfolding norm_id_minus_trunc_ell2 by simp then have \((\S. norm (\ - trunc_ell2 S \)) \ 0) (finite_subsets_at_top UNIV)\ by auto then have \((\S. \ - trunc_ell2 S \) \ 0) (finite_subsets_at_top UNIV)\ by (rule tendsto_norm_zero_cancel) then show ?thesis apply (rule Lim_transform2[where f=\\_. \\, rotated]) by simp qed lemma trunc_ell2_norm_mono: \M \ N \ norm (trunc_ell2 M \) \ norm (trunc_ell2 N \)\ proof (rule power2_le_imp_le[rotated], force, transfer) fix M N :: \'a set\ and \ :: \'a \ complex\ assume \M \ N\ and \has_ell2_norm \\ have \(ell2_norm (\i. if i \ M then \ i else 0))\<^sup>2 = (\\<^sub>\i\M. (cmod (\ i))\<^sup>2)\ unfolding ell2_norm_square apply (rule infsum_cong_neutral) by auto also have \\ \ (\\<^sub>\i\N. (cmod (\ i))\<^sup>2)\ apply (rule infsum_mono2) using \has_ell2_norm \\ \M \ N\ by (auto simp add: ell2_norm_square has_ell2_norm_def simp flip: norm_power intro: summable_on_subset_banach) also have \\ = (ell2_norm (\i. if i \ N then \ i else 0))\<^sup>2\ unfolding ell2_norm_square apply (rule infsum_cong_neutral) by auto finally show \(ell2_norm (\i. if i \ M then \ i else 0))\<^sup>2 \ (ell2_norm (\i. if i \ N then \ i else 0))\<^sup>2\ by - qed lemma trunc_ell2_twice[simp]: \trunc_ell2 M (trunc_ell2 N \) = trunc_ell2 (M\N) \\ apply transfer by auto lemma trunc_ell2_union: \trunc_ell2 (M \ N) \ = trunc_ell2 M \ + trunc_ell2 N \ - trunc_ell2 (M\N) \\ apply transfer by auto lemma trunc_ell2_union_disjoint: \M\N = {} \ trunc_ell2 (M \ N) \ = trunc_ell2 M \ + trunc_ell2 N \\ by (simp add: trunc_ell2_union) lemma trunc_ell2_union_Diff: \M \ N \ trunc_ell2 (N-M) \ = trunc_ell2 N \ - trunc_ell2 M \\ using trunc_ell2_union_disjoint[where M=\N-M\ and N=M and \=\] by (simp add: Un_commute inf.commute le_iff_sup) lemma trunc_ell2_lim: \((\S. trunc_ell2 S \) \ trunc_ell2 M \) (finite_subsets_at_top M)\ proof - have \((\S. trunc_ell2 S (trunc_ell2 M \)) \ trunc_ell2 M \) (finite_subsets_at_top UNIV)\ using trunc_ell2_lim_at_UNIV by blast then have \((\S. trunc_ell2 (S\M) \) \ trunc_ell2 M \) (finite_subsets_at_top UNIV)\ by simp then show \((\S. trunc_ell2 S \) \ trunc_ell2 M \) (finite_subsets_at_top M)\ unfolding filterlim_def apply (subst (asm) filtermap_filtermap[where g=\\S. S\M\, symmetric]) apply (subst (asm) finite_subsets_at_top_inter[where A=M and B=UNIV]) by auto qed lemma trunc_ell2_lim_general: assumes big: \\G. finite G \ G \ M \ (\\<^sub>F H in F. H \ G)\ assumes small: \\\<^sub>F H in F. H \ M\ shows \((\S. trunc_ell2 S \) \ trunc_ell2 M \) F\ proof (rule tendstoI) fix e :: real assume \e > 0\ from trunc_ell2_lim[THEN tendsto_iff[THEN iffD1], rule_format, OF \e > 0\, where M=M and \=\] obtain G where \finite G\ and \G \ M\ and close: \dist (trunc_ell2 G \) (trunc_ell2 M \) < e\ apply atomize_elim unfolding eventually_finite_subsets_at_top by blast from \finite G\ \G \ M\ and big have \\\<^sub>F H in F. H \ G\ by - with small have \\\<^sub>F H in F. H \ M \ H \ G\ by (simp add: eventually_conj_iff) then show \\\<^sub>F H in F. dist (trunc_ell2 H \) (trunc_ell2 M \) < e\ proof (rule eventually_mono) fix H assume GHM: \H \ M \ H \ G\ have \dist (trunc_ell2 H \) (trunc_ell2 M \) = norm (trunc_ell2 (M-H) \)\ by (simp add: GHM dist_ell2_def norm_minus_commute trunc_ell2_union_Diff) also have \\ \ norm (trunc_ell2 (M-G) \)\ by (simp add: Diff_mono GHM trunc_ell2_norm_mono) also have \\ = dist (trunc_ell2 G \) (trunc_ell2 M \)\ by (simp add: \G \ M\ dist_ell2_def norm_minus_commute trunc_ell2_union_Diff) also have \\ < e\ using close by simp finally show \dist (trunc_ell2 H \) (trunc_ell2 M \) < e\ by - qed qed subsection \Kets and bras\ lift_definition ket :: "'a \ 'a ell2" is "\x y. if x=y then 1 else 0" by (rule has_ell2_norm_ket) abbreviation bra :: "'a \ (_,complex) cblinfun" where "bra i \ vector_to_cblinfun (ket i)*" for i instance ell2 :: (type) not_singleton proof standard have "ket undefined \ (0::'a ell2)" proof transfer show "(\y. if (undefined::'a) = y then 1::complex else 0) \ (\_. 0)" by (meson one_neq_zero) qed thus \\x y::'a ell2. x \ y\ by blast qed lemma cinner_ket_left: \\ket i, \\ = Rep_ell2 \ i\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto lemma cinner_ket_right: \\\, ket i\ = cnj (Rep_ell2 \ i)\ apply (transfer fixing: i) apply (subst infsum_cong_neutral[where T=\{i}\]) by auto lemma cinner_ket_eqI: assumes \\i. cinner (ket i) \ = cinner (ket i) \\ shows \\ = \\ by (metis Rep_ell2_inject assms cinner_ket_left ext) lemma norm_ket[simp]: "norm (ket i) = 1" apply transfer by (rule ell2_norm_ket) lemma cinner_ket_same[simp]: \\ket i, ket i\ = 1\ proof- have \norm (ket i) = 1\ by simp hence \sqrt (cmod \ket i, ket i\) = 1\ by (metis norm_eq_sqrt_cinner) hence \cmod \ket i, ket i\ = 1\ using real_sqrt_eq_1_iff by blast moreover have \\ket i, ket i\ = cmod \ket i, ket i\\ proof- have \\ket i, ket i\ \ \\ by (simp add: cinner_real) thus ?thesis by (metis cinner_ge_zero complex_of_real_cmod) qed ultimately show ?thesis by simp qed lemma orthogonal_ket[simp]: \is_orthogonal (ket i) (ket j) \ i \ j\ by (simp add: cinner_ket_left ket.rep_eq) lemma cinner_ket: \\ket i, ket j\ = (if i=j then 1 else 0)\ by (simp add: cinner_ket_left ket.rep_eq) lemma ket_injective[simp]: \ket i = ket j \ i = j\ by (metis cinner_ket one_neq_zero) lemma inj_ket[simp]: \inj ket\ by (simp add: inj_on_def) lemma trunc_ell2_ket_cspan: \trunc_ell2 S x \ (cspan (range ket))\ if \finite S\ proof (use that in induction) case empty then show ?case by (auto intro: complex_vector.span_zero) next case (insert a F) from insert.hyps have \trunc_ell2 (insert a F) x = trunc_ell2 F x + Rep_ell2 x a *\<^sub>C ket a\ apply (transfer fixing: F a) by auto with insert.IH show ?case by (simp add: complex_vector.span_add_eq complex_vector.span_base complex_vector.span_scale) qed lemma closed_cspan_range_ket[simp]: \closure (cspan (range ket)) = UNIV\ proof (intro set_eqI iffI UNIV_I closure_approachable[THEN iffD2] allI impI) fix \ :: \'a ell2\ fix e :: real assume \e > 0\ have \((\S. trunc_ell2 S \) \ \) (finite_subsets_at_top UNIV)\ by (rule trunc_ell2_lim_at_UNIV) then obtain F where \finite F\ and \dist (trunc_ell2 F \) \ < e\ apply (drule_tac tendstoD[OF _ \e > 0\]) by (auto dest: simp: eventually_finite_subsets_at_top) moreover have \trunc_ell2 F \ \ cspan (range ket)\ using \finite F\ trunc_ell2_ket_cspan by blast ultimately show \\\\cspan (range ket). dist \ \ < e\ by auto qed lemma ccspan_range_ket[simp]: "ccspan (range ket) = (top::('a ell2 ccsubspace))" proof- have \closure (complex_vector.span (range ket)) = (UNIV::'a ell2 set)\ using Complex_L2.closed_cspan_range_ket by blast thus ?thesis by (simp add: ccspan.abs_eq top_ccsubspace.abs_eq) qed lemma cspan_range_ket_finite[simp]: "cspan (range ket :: 'a::finite ell2 set) = UNIV" by (metis closed_cspan_range_ket closure_finite_cspan finite_class.finite_UNIV finite_imageI) instance ell2 :: (finite) cfinite_dim proof define basis :: \'a ell2 set\ where \basis = range ket\ have \finite basis\ unfolding basis_def by simp moreover have \cspan basis = UNIV\ by (simp add: basis_def) ultimately show \\basis::'a ell2 set. finite basis \ cspan basis = UNIV\ by auto qed instantiation ell2 :: (enum) onb_enum begin definition "canonical_basis_ell2 = map ket Enum.enum" instance proof show "distinct (canonical_basis::'a ell2 list)" proof- have \finite (UNIV::'a set)\ by simp have \distinct (enum_class.enum::'a list)\ using enum_distinct by blast moreover have \inj_on ket (set enum_class.enum)\ by (meson inj_onI ket_injective) ultimately show ?thesis unfolding canonical_basis_ell2_def using distinct_map by blast qed show "is_ortho_set (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (z3) norm_ket f_inv_into_f is_ortho_set_def orthogonal_ket norm_zero) show "cindependent (set (canonical_basis::'a ell2 list))" apply (auto simp: canonical_basis_ell2_def enum_UNIV) by (smt (verit, best) norm_ket f_inv_into_f is_ortho_set_def is_ortho_set_cindependent orthogonal_ket norm_zero) show "cspan (set (canonical_basis::'a ell2 list)) = UNIV" by (auto simp: canonical_basis_ell2_def enum_UNIV) show "norm (x::'a ell2) = 1" if "(x::'a ell2) \ set canonical_basis" for x :: "'a ell2" using that unfolding canonical_basis_ell2_def by auto qed end lemma canonical_basis_length_ell2[code_unfold, simp]: "length (canonical_basis ::'a::enum ell2 list) = CARD('a)" unfolding canonical_basis_ell2_def apply simp using card_UNIV_length_enum by metis lemma ket_canonical_basis: "ket x = canonical_basis ! enum_idx x" proof- have "x = (enum_class.enum::'a list) ! enum_idx x" using enum_idx_correct[where i = x] by simp hence p1: "ket x = ket ((enum_class.enum::'a list) ! enum_idx x)" by simp have "enum_idx x < length (enum_class.enum::'a list)" using enum_idx_bound[where x = x]. hence "(map ket (enum_class.enum::'a list)) ! enum_idx x = ket ((enum_class.enum::'a list) ! enum_idx x)" by auto thus ?thesis unfolding canonical_basis_ell2_def using p1 by auto qed lemma clinear_equal_ket: fixes f g :: \'a::finite ell2 \ _\ assumes \clinear f\ assumes \clinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule complex_vector.linear_eq_on_span[where f=f and g=g and B=\range ket\]) using assms by auto lemma equal_ket: fixes A B :: \('a ell2, 'b::complex_normed_vector) cblinfun\ assumes \\ x. cblinfun_apply A (ket x) = cblinfun_apply B (ket x)\ shows \A = B\ apply (rule cblinfun_eq_gen_eqI[where G=\range ket\]) using assms by auto lemma antilinear_equal_ket: fixes f g :: \'a::finite ell2 \ _\ assumes \antilinear f\ assumes \antilinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ proof - have [simp]: \clinear (f \ from_conjugate_space)\ apply (rule antilinear_o_antilinear) using assms by (simp_all add: antilinear_from_conjugate_space) have [simp]: \clinear (g \ from_conjugate_space)\ apply (rule antilinear_o_antilinear) using assms by (simp_all add: antilinear_from_conjugate_space) have [simp]: \cspan (to_conjugate_space ` (range ket :: 'a ell2 set)) = UNIV\ by simp have "f o from_conjugate_space = g o from_conjugate_space" apply (rule ext) apply (rule complex_vector.linear_eq_on_span[where f="f o from_conjugate_space" and g="g o from_conjugate_space" and B=\to_conjugate_space ` range ket\]) apply (simp, simp) using assms(3) by (auto simp: to_conjugate_space_inverse) then show "f = g" by (smt (verit) UNIV_I from_conjugate_space_inverse surj_def surj_fun_eq to_conjugate_space_inject) qed lemma cinner_ket_adjointI: fixes F::"'a ell2 \\<^sub>C\<^sub>L _" and G::"'b ell2 \\<^sub>C\<^sub>L_" assumes "\ i j. \F *\<^sub>V ket i, ket j\ = \ket i, G *\<^sub>V ket j\" shows "F = G*" proof - from assms have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ if \x \ range ket\ and \y \ range ket\ for x y using that by auto then have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ if \x \ range ket\ for x y apply (rule bounded_clinear_eq_on[where G=\range ket\ and t=y, rotated 2]) using that by (auto intro!: bounded_linear_intros) then have \(F *\<^sub>V x) \\<^sub>C y = x \\<^sub>C (G *\<^sub>V y)\ for x y apply (rule bounded_antilinear_eq_on[where G=\range ket\ and t=x, rotated 2]) by (auto intro!: bounded_linear_intros) then show ?thesis by (rule adjoint_eqI) qed lemma ket_nonzero[simp]: "ket i \ 0" using norm_ket[of i] by force lemma cindependent_ket: "cindependent (range (ket::'a\_))" proof- define S where "S = range (ket::'a\_)" have "is_ortho_set S" unfolding S_def is_ortho_set_def by auto moreover have "0 \ S" unfolding S_def using ket_nonzero by (simp add: image_iff) ultimately show ?thesis using is_ortho_set_cindependent[where A = S] unfolding S_def by blast qed lemma cdim_UNIV_ell2[simp]: \cdim (UNIV::'a::finite ell2 set) = CARD('a)\ apply (subst cspan_range_ket_finite[symmetric]) by (metis card_image cindependent_ket complex_vector.dim_span_eq_card_independent inj_ket) lemma is_ortho_set_ket[simp]: \is_ortho_set (range ket)\ using is_ortho_set_def by fastforce lemma bounded_clinear_equal_ket: fixes f g :: \'a ell2 \ _\ assumes \bounded_clinear f\ assumes \bounded_clinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule bounded_clinear_eq_on[of f g \range ket\]) using assms by auto lemma bounded_antilinear_equal_ket: fixes f g :: \'a ell2 \ _\ assumes \bounded_antilinear f\ assumes \bounded_antilinear g\ assumes \\i. f (ket i) = g (ket i)\ shows \f = g\ apply (rule ext) apply (rule bounded_antilinear_eq_on[of f g \range ket\]) using assms by auto lemma ket_CARD_1_is_1: \ket x = 1\ for x :: \'a::CARD_1\ apply transfer by simp lemma is_onb_ket[simp]: \is_onb (range ket)\ by (auto simp: is_onb_def) +lemma ell2_sum_ket: \\ = (\i\UNIV. Rep_ell2 \ i *\<^sub>C ket i)\ for \ :: \_::finite ell2\ + apply transfer apply (rule ext) + apply (subst sum_single) + by auto + subsection \Butterflies\ lemma cspan_butterfly_ket: \cspan {butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = UNIV\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b::finite) (j::'a::finite). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cspan_butterfly_UNIV) by auto qed lemma cindependent_butterfly_ket: \cindependent {butterfly (ket i) (ket j)| (i::'b) (j::'a). True}\ proof - have *: \{butterfly (ket i) (ket j)| (i::'b) (j::'a). True} = {butterfly a b |a b. a \ range ket \ b \ range ket}\ by auto show ?thesis apply (subst *) apply (rule cindependent_butterfly) by auto qed lemma clinear_eq_butterfly_ketI: fixes F G :: \('a::finite ell2 \\<^sub>C\<^sub>L 'b::finite ell2) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes "\i j. F (butterfly (ket i) (ket j)) = G (butterfly (ket i) (ket j))" shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_ket) using assms by auto lemma sum_butterfly_ket[simp]: \(\(i::'a::finite)\UNIV. butterfly (ket i) (ket i)) = id_cblinfun\ apply (rule equal_ket) apply (subst complex_vector.linear_sum[where f=\\y. y *\<^sub>V ket _\]) apply (auto simp add: scaleC_cblinfun.rep_eq cblinfun.add_left clinearI butterfly_def cblinfun_compose_image cinner_ket) apply (subst sum.mono_neutral_cong_right[where S=\{_}\]) by auto subsection \One-dimensional spaces\ instantiation ell2 :: ("{enum,CARD_1}") one_dim begin text \Note: enum is not needed logically, but without it this instantiation clashes with \instantiation ell2 :: (enum) onb_enum\\ instance proof show "canonical_basis = [1::'a ell2]" unfolding canonical_basis_ell2_def apply transfer by (simp add: enum_CARD_1[of undefined]) show "a *\<^sub>C 1 * b *\<^sub>C 1 = (a * b) *\<^sub>C (1::'a ell2)" for a b apply (transfer fixing: a b) by simp show "x / y = x * inverse y" for x y :: "'a ell2" by (simp add: divide_inverse) show "inverse (c *\<^sub>C 1) = inverse c *\<^sub>C (1::'a ell2)" for c :: complex apply transfer by auto qed end subsection \Classical operators\ text \We call an operator mapping \<^term>\ket x\ to \<^term>\ket (\ x)\ or \<^term>\0\ "classical". (The meaning is inspired by the fact that in quantum mechanics, such operators usually correspond to operations with classical interpretation (such as Pauli-X, CNOT, measurement in the computational basis, etc.))\ definition classical_operator :: "('a\'b option) \ 'a ell2 \\<^sub>C\<^sub>L'b ell2" where "classical_operator \ = (let f = (\t. (case \ (inv (ket::'a\_) t) of None \ (0::'b ell2) | Some i \ ket i)) in cblinfun_extension (range (ket::'a\_)) f)" definition "classical_operator_exists \ \ cblinfun_extension_exists (range ket) (\t. case \ (inv ket t) of None \ 0 | Some i \ ket i)" lemma classical_operator_existsI: assumes "\x. B *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" shows "classical_operator_exists \" unfolding classical_operator_exists_def apply (rule cblinfun_extension_existsI[of _ B]) using assms by (auto simp: inv_f_f[OF inj_ket]) -lemma classical_operator_exists_inj: +lemma assumes "inj_map \" - shows "classical_operator_exists \" -proof (unfold classical_operator_exists_def, rule cblinfun_extension_exists_ortho) - show \is_ortho_set (range ket)\ - by simp - show \closure (cspan (range ket)) = UNIV\ - by simp + shows classical_operator_exists_inj: "classical_operator_exists \" + and classical_operator_norm_inj: \norm (classical_operator \) \ 1\ +proof - have \is_orthogonal (case \ x of None \ 0 | Some x' \ ket x') (case \ y of None \ 0 | Some y' \ ket y')\ if \x \ y\ for x y apply (cases \\ x\; cases \\ y\) using that assms by (auto simp add: inj_map_def) - then show \is_orthogonal (case \ (inv ket x) of None \ 0 | Some x' \ ket x') + then have 1: \is_orthogonal (case \ (inv ket x) of None \ 0 | Some x' \ ket x') (case \ (inv ket y) of None \ 0 | Some y' \ ket y')\ if \x \ range ket\ and \y \ range ket\ and \x \ y\ for x y using that by auto + have \norm (case \ x of None \ 0 | Some x \ ket x) \ 1 * norm (ket x)\ for x apply (cases \\ x\) by auto - then show \norm (case \ (inv ket x) of None \ 0 | Some x \ ket x) \ 1 * norm x\ + then have 2: \norm (case \ (inv ket x) of None \ 0 | Some x \ ket x) \ 1 * norm x\ if \x \ range ket\ for x using that by auto + + show \classical_operator_exists \\ + unfolding classical_operator_exists_def + using _ _ 1 2 apply (rule cblinfun_extension_exists_ortho) + by simp_all + + show \norm (classical_operator \) \ 1\ + unfolding classical_operator_def Let_def + using _ _ 1 2 apply (rule cblinfun_extension_exists_ortho_norm) + by simp_all qed lemma classical_operator_exists_finite[simp]: "classical_operator_exists (\ :: _::finite \ _)" unfolding classical_operator_exists_def apply (rule cblinfun_extension_exists_finite_dim) using cindependent_ket apply blast using finite_class.finite_UNIV finite_imageI closed_cspan_range_ket closure_finite_cspan by blast lemma classical_operator_ket: assumes "classical_operator_exists \" shows "(classical_operator \) *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" unfolding classical_operator_def using f_inv_into_f ket_injective rangeI by (metis assms cblinfun_extension_apply classical_operator_exists_def) lemma classical_operator_ket_finite: "(classical_operator \) *\<^sub>V (ket (x::'a::finite)) = (case \ x of Some i \ ket i | None \ 0)" by (rule classical_operator_ket, simp) lemma classical_operator_adjoint[simp]: fixes \ :: "'a \ 'b option" assumes a1: "inj_map \" shows "(classical_operator \)* = classical_operator (inv_map \)" proof- define F where "F = classical_operator (inv_map \)" define G where "G = classical_operator \" have "\F *\<^sub>V ket i, ket j\ = \ket i, G *\<^sub>V ket j\" for i j proof- have w1: "(classical_operator (inv_map \)) *\<^sub>V (ket i) = (case inv_map \ i of Some k \ ket k | None \ 0)" by (simp add: classical_operator_ket classical_operator_exists_inj) have w2: "(classical_operator \) *\<^sub>V (ket j) = (case \ j of Some k \ ket k | None \ 0)" by (simp add: assms classical_operator_ket classical_operator_exists_inj) have "\F *\<^sub>V ket i, ket j\ = \classical_operator (inv_map \) *\<^sub>V ket i, ket j\" unfolding F_def by blast also have "\ = \(case inv_map \ i of Some k \ ket k | None \ 0), ket j\" using w1 by simp also have "\ = \ket i, (case \ j of Some k \ ket k | None \ 0)\" proof(induction "inv_map \ i") case None hence pi1: "None = inv_map \ i". show ?case proof (induction "\ j") case None thus ?case using pi1 by auto next case (Some c) have "c \ i" proof(rule classical) assume "\(c \ i)" hence "c = i" by blast hence "inv_map \ c = inv_map \ i" by simp hence "inv_map \ c = None" by (simp add: pi1) moreover have "inv_map \ c = Some j" using Some.hyps unfolding inv_map_def apply auto by (metis a1 f_inv_into_f inj_map_def option.distinct(1) rangeI) ultimately show ?thesis by simp qed thus ?thesis by (metis None.hyps Some.hyps cinner_zero_left orthogonal_ket option.simps(4) option.simps(5)) qed next case (Some d) hence s1: "Some d = inv_map \ i". show "\case inv_map \ i of None \ 0 | Some a \ ket a, ket j\ = \ket i, case \ j of None \ 0 | Some a \ ket a\" proof(induction "\ j") case None have "d \ j" proof(rule classical) assume "\(d \ j)" hence "d = j" by blast hence "\ d = \ j" by simp hence "\ d = None" by (simp add: None.hyps) moreover have "\ d = Some i" using Some.hyps unfolding inv_map_def apply auto by (metis f_inv_into_f option.distinct(1) option.inject) ultimately show ?thesis by simp qed thus ?case by (metis None.hyps Some.hyps cinner_zero_right orthogonal_ket option.case_eq_if option.simps(5)) next case (Some c) hence s2: "\ j = Some c" by simp have "\ket d, ket j\ = \ket i, ket c\" proof(cases "\ j = Some i") case True hence ij: "Some j = inv_map \ i" unfolding inv_map_def apply auto apply (metis a1 f_inv_into_f inj_map_def option.discI range_eqI) by (metis range_eqI) have "i = c" using True s2 by auto moreover have "j = d" by (metis option.inject s1 ij) ultimately show ?thesis by (simp add: cinner_ket_same) next case False moreover have "\ d = Some i" using s1 unfolding inv_map_def by (metis f_inv_into_f option.distinct(1) option.inject) ultimately have "j \ d" by auto moreover have "i \ c" using False s2 by auto ultimately show ?thesis by (metis orthogonal_ket) qed hence "\case Some d of None \ 0 | Some a \ ket a, ket j\ = \ket i, case Some c of None \ 0 | Some a \ ket a\" by simp thus "\case inv_map \ i of None \ 0 | Some a \ ket a, ket j\ = \ket i, case \ j of None \ 0 | Some a \ ket a\" by (simp add: Some.hyps s1) qed qed also have "\ = \ket i, classical_operator \ *\<^sub>V ket j\" by (simp add: w2) also have "\ = \ket i, G *\<^sub>V ket j\" unfolding G_def by blast finally show ?thesis . qed hence "G* = F" using cinner_ket_adjointI by auto thus ?thesis unfolding G_def F_def . qed lemma fixes \::"'b \ 'c option" and \::"'a \ 'b option" assumes "classical_operator_exists \" assumes "classical_operator_exists \" shows classical_operator_exists_comp[simp]: "classical_operator_exists (\ \\<^sub>m \)" and classical_operator_mult[simp]: "classical_operator \ o\<^sub>C\<^sub>L classical_operator \ = classical_operator (\ \\<^sub>m \)" proof - define C\ C\ C\\ where "C\ = classical_operator \" and "C\ = classical_operator \" and "C\\ = classical_operator (\ \\<^sub>m \)" have C\x: "C\ *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" for x unfolding C\_def using \classical_operator_exists \\ by (rule classical_operator_ket) have C\x: "C\ *\<^sub>V (ket x) = (case \ x of Some i \ ket i | None \ 0)" for x unfolding C\_def using \classical_operator_exists \\ by (rule classical_operator_ket) have C\\x': "(C\ o\<^sub>C\<^sub>L C\) *\<^sub>V (ket x) = (case (\ \\<^sub>m \) x of Some i \ ket i | None \ 0)" for x apply (simp add: scaleC_cblinfun.rep_eq C\x) apply (cases "\ x") by (auto simp: C\x) thus \classical_operator_exists (\ \\<^sub>m \)\ by (rule classical_operator_existsI) hence "C\\ *\<^sub>V (ket x) = (case (\ \\<^sub>m \) x of Some i \ ket i | None \ 0)" for x unfolding C\\_def by (rule classical_operator_ket) with C\\x' have "(C\ o\<^sub>C\<^sub>L C\) *\<^sub>V (ket x) = C\\ *\<^sub>V (ket x)" for x by simp thus "C\ o\<^sub>C\<^sub>L C\ = C\\" by (simp add: equal_ket) qed lemma classical_operator_Some[simp]: "classical_operator (Some::'a\_) = id_cblinfun" proof- have "(classical_operator Some) *\<^sub>V (ket i) = id_cblinfun *\<^sub>V (ket i)" for i::'a apply (subst classical_operator_ket) apply (rule classical_operator_exists_inj) by auto thus ?thesis using equal_ket[where A = "classical_operator (Some::'a \ _ option)" and B = "id_cblinfun::'a ell2 \\<^sub>C\<^sub>L _"] by blast qed lemma isometry_classical_operator[simp]: fixes \::"'a \ 'b" assumes a1: "inj \" shows "isometry (classical_operator (Some o \))" proof - have b0: "inj_map (Some \ \)" by (simp add: a1) have b0': "inj_map (inv_map (Some \ \))" by simp have b1: "inv_map (Some \ \) \\<^sub>m (Some \ \) = Some" apply (rule ext) unfolding inv_map_def o_def using assms unfolding inj_def inv_def by auto have b3: "classical_operator (inv_map (Some \ \)) o\<^sub>C\<^sub>L classical_operator (Some \ \) = classical_operator (inv_map (Some \ \) \\<^sub>m (Some \ \))" by (metis b0 b0' b1 classical_operator_Some classical_operator_exists_inj classical_operator_mult) show ?thesis unfolding isometry_def apply (subst classical_operator_adjoint) using b0 by (auto simp add: b1 b3) qed lemma unitary_classical_operator[simp]: fixes \::"'a \ 'b" assumes a1: "bij \" shows "unitary (classical_operator (Some o \))" proof (unfold unitary_def, rule conjI) have "inj \" using a1 bij_betw_imp_inj_on by auto hence "isometry (classical_operator (Some o \))" by simp hence "classical_operator (Some \ \)* o\<^sub>C\<^sub>L classical_operator (Some \ \) = id_cblinfun" unfolding isometry_def by simp thus \classical_operator (Some \ \)* o\<^sub>C\<^sub>L classical_operator (Some \ \) = id_cblinfun\ by simp next have "inj \" by (simp add: assms bij_is_inj) have comp: "Some \ \ \\<^sub>m inv_map (Some \ \) = Some" apply (rule ext) unfolding inv_map_def o_def map_comp_def unfolding inv_def apply auto apply (metis \inj \\ inv_def inv_f_f) using bij_def image_iff range_eqI by (metis a1) have "classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (Some \ \)* = classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (inv_map (Some \ \))" by (simp add: \inj \\) also have "\ = classical_operator ((Some \ \) \\<^sub>m (inv_map (Some \ \)))" by (simp add: \inj \\ classical_operator_exists_inj) also have "\ = classical_operator (Some::'b\_)" using comp by simp also have "\ = (id_cblinfun:: 'b ell2 \\<^sub>C\<^sub>L _)" by simp finally show "classical_operator (Some \ \) o\<^sub>C\<^sub>L classical_operator (Some \ \)* = id_cblinfun". qed unbundle no_lattice_syntax unbundle no_cblinfun_notation end diff --git a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy --- a/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy +++ b/thys/Complex_Bounded_Operators/Complex_Vector_Spaces.thy @@ -1,3125 +1,3158 @@ section \\Complex_Vector_Spaces\ -- Complex Vector Spaces\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Vector_Spaces imports "HOL-Analysis.Elementary_Topology" "HOL-Analysis.Operator_Norm" "HOL-Analysis.Elementary_Normed_Spaces" "HOL-Library.Set_Algebras" "HOL-Analysis.Starlike" "HOL-Types_To_Sets.Types_To_Sets" "HOL-Library.Complemented_Lattices" Extra_Vector_Spaces Extra_Ordered_Fields Extra_Operator_Norm Extra_General Complex_Vector_Spaces0 begin bundle notation_norm begin notation norm ("\_\") end unbundle lattice_syntax subsection \Misc\ (* Should rather be in Extra_Vector_Spaces but then complex_vector.span_image_scale' does not exist for some reason. Ideally this would replace Vector_Spaces.vector_space.span_image_scale. *) lemma (in vector_space) span_image_scale': \ \Strengthening of @{thm [source] vector_space.span_image_scale} without the condition \finite S\\ assumes nz: "\x. x \ S \ c x \ 0" shows "span ((\x. c x *s x) ` S) = span S" proof have \((\x. c x *s x) ` S) \ span S\ by (metis (mono_tags, lifting) image_subsetI in_mono local.span_superset local.subspace_scale local.subspace_span) then show \span ((\x. c x *s x) ` S) \ span S\ by (simp add: local.span_minimal) next have \x \ span ((\x. c x *s x) ` S)\ if \x \ S\ for x proof - have \x = inverse (c x) *s c x *s x\ by (simp add: nz that) moreover have \c x *s x \ (\x. c x *s x) ` S\ using that by blast ultimately show ?thesis by (metis local.span_base local.span_scale) qed then show \span S \ span ((\x. c x *s x) ` S)\ by (simp add: local.span_minimal subsetI) qed lemma (in scaleC) scaleC_real: assumes "r\\" shows "r *\<^sub>C x = Re r *\<^sub>R x" unfolding scaleR_scaleC using assms by simp lemma of_complex_of_real_eq [simp]: "of_complex (of_real n) = of_real n" unfolding of_complex_def of_real_def unfolding scaleR_scaleC by simp lemma Complexs_of_real [simp]: "of_real r \ \" unfolding Complexs_def of_real_def of_complex_def apply (subst scaleR_scaleC) by simp lemma Reals_in_Complexs: "\ \ \" unfolding Reals_def by auto lemma (in bounded_clinear) bounded_linear: "bounded_linear f" by standard lemma clinear_times: "clinear (\x. c * x)" for c :: "'a::complex_algebra" by (auto simp: clinearI distrib_left) lemma (in clinear) linear: \linear f\ by standard lemma bounded_clinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_clinear f" using assms by (auto intro!: exI bounded_clinear.intro clinearI simp: bounded_clinear_axioms_def) lemma bounded_clinear_id[simp]: \bounded_clinear id\ by (simp add: id_def) definition cbilinear :: \('a::complex_vector \ 'b::complex_vector \ 'c::complex_vector) \ bool\ where \cbilinear = (\ f. (\ y. clinear (\ x. f x y)) \ (\ x. clinear (\ y. f x y)) )\ lemma cbilinear_add_left: assumes \cbilinear f\ shows \f (a + b) c = f a c + f b c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_add_right: assumes \cbilinear f\ shows \f a (b + c) = f a b + f a c\ by (smt (verit, del_insts) assms cbilinear_def complex_vector.linear_add) lemma cbilinear_times: fixes g' :: \'a::complex_vector \ complex\ and g :: \'b::complex_vector \ complex\ assumes \\ x y. h x y = (g' x)*(g y)\ and \clinear g\ and \clinear g'\ shows \cbilinear h\ proof - have w1: "h (b1 + b2) y = h b1 y + h b2 y" for b1 :: 'a and b2 :: 'a and y proof- have \h (b1 + b2) y = g' (b1 + b2) * g y\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = (g' b1 + g' b2) * g y\ using \clinear g'\ unfolding clinear_def by (simp add: assms(3) complex_vector.linear_add) also have \\ = g' b1 * g y + g' b2 * g y\ by (simp add: ring_class.ring_distribs(2)) also have \\ = h b1 y + h b2 y\ using assms(1) by auto finally show ?thesis by blast qed have w2: "h (r *\<^sub>C b) y = r *\<^sub>C h b y" for r :: complex and b :: 'a and y proof- have \h (r *\<^sub>C b) y = g' (r *\<^sub>C b) * g y\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' b * g y)\ by (simp add: assms(3) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h b y)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "clinear (\x. h x y)" for y :: 'b unfolding clinear_def by (meson clinearI clinear_def w1 w2) hence t2: "\y. clinear (\x. h x y)" by simp have v1: "h x (b1 + b2) = h x b1 + h x b2" for b1 :: 'b and b2 :: 'b and x proof- have \h x (b1 + b2) = g' x * g (b1 + b2)\ using \\ x y. h x y = (g' x)*(g y)\ by auto also have \\ = g' x * (g b1 + g b2)\ using \clinear g'\ unfolding clinear_def by (simp add: assms(2) complex_vector.linear_add) also have \\ = g' x * g b1 + g' x * g b2\ by (simp add: ring_class.ring_distribs(1)) also have \\ = h x b1 + h x b2\ using assms(1) by auto finally show ?thesis by blast qed have v2: "h x (r *\<^sub>C b) = r *\<^sub>C h x b" for r :: complex and b :: 'b and x proof- have \h x (r *\<^sub>C b) = g' x * g (r *\<^sub>C b)\ by (simp add: assms(1)) also have \\ = r *\<^sub>C (g' x * g b)\ by (simp add: assms(2) complex_vector.linear_scale) also have \\ = r *\<^sub>C (h x b)\ by (simp add: assms(1)) finally show ?thesis by blast qed have "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) (h x)" for x :: 'a using v1 v2 by (meson clinearI clinear_def) hence t1: "\x. clinear (h x)" unfolding clinear_def by simp show ?thesis unfolding cbilinear_def by (simp add: t1 t2) qed lemma csubspace_is_subspace: "csubspace A \ subspace A" apply (rule subspaceI) by (auto simp: complex_vector.subspace_def scaleR_scaleC) lemma span_subset_cspan: "span A \ cspan A" unfolding span_def complex_vector.span_def by (simp add: csubspace_is_subspace hull_antimono) lemma cindependent_implies_independent: assumes "cindependent (S::'a::complex_vector set)" shows "independent S" using assms unfolding dependent_def complex_vector.dependent_def using span_subset_cspan by blast lemma cspan_singleton: "cspan {x} = {\ *\<^sub>C x| \. True}" proof - have \cspan {x} = {y. y\cspan {x}}\ by auto also have \\ = {\ *\<^sub>C x| \. True}\ apply (subst complex_vector.span_breakdown_eq) by auto finally show ?thesis by - qed lemma cspan_as_span: "cspan (B::'a::complex_vector set) = span (B \ scaleC \ ` B)" proof auto let ?cspan = complex_vector.span let ?rspan = real_vector.span fix \ assume cspan: "\ \ ?cspan B" have "\B' r. finite B' \ B' \ B \ \ = (\b\B'. r b *\<^sub>C b)" using complex_vector.span_explicit[of B] cspan by auto then obtain B' r where "finite B'" and "B' \ B" and \_explicit: "\ = (\b\B'. r b *\<^sub>C b)" by atomize_elim define R where "R = B \ scaleC \ ` B" have x2: "(case x of (b, i) \ if i then Im (r b) *\<^sub>R \ *\<^sub>C b else Re (r b) *\<^sub>R b) \ span (B \ (*\<^sub>C) \ ` B)" if "x \ B' \ (UNIV::bool set)" for x :: "'a \ bool" using that \B' \ B\ by (auto simp add: real_vector.span_base real_vector.span_scale subset_iff) have x1: "\ = (\x\B'. \i\UNIV. if i then Im (r x) *\<^sub>R \ *\<^sub>C x else Re (r x) *\<^sub>R x)" if "\b. r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" using that by (simp add: UNIV_bool \_explicit) moreover have "r b *\<^sub>C b = Re (r b) *\<^sub>R b + Im (r b) *\<^sub>R \ *\<^sub>C b" for b using complex_eq scaleC_add_left scaleC_scaleC scaleR_scaleC by (metis (no_types, lifting) complex_of_real_i i_complex_of_real) ultimately have "\ = (\(b,i)\(B'\UNIV). if i then Im (r b) *\<^sub>R (\ *\<^sub>C b) else Re (r b) *\<^sub>R b)" by (simp add: sum.cartesian_product) also have "\ \ ?rspan R" unfolding R_def using x2 by (rule real_vector.span_sum) finally show "\ \ ?rspan R" by - next let ?cspan = complex_vector.span let ?rspan = real_vector.span define R where "R = B \ scaleC \ ` B" fix \ assume rspan: "\ \ ?rspan R" have "subspace {a. a \ cspan B}" by (rule real_vector.subspaceI, auto simp add: complex_vector.span_zero complex_vector.span_add_eq2 complex_vector.span_scale scaleR_scaleC) moreover have "x \ cspan B" if "x \ R" for x :: 'a using that R_def complex_vector.span_base complex_vector.span_scale by fastforce ultimately show "\ \ ?cspan B" using real_vector.span_induct rspan by blast qed lemma isomorphic_equal_cdim: assumes lin_f: \clinear f\ assumes inj_f: \inj_on f (cspan S)\ assumes im_S: \f ` S = T\ shows \cdim S = cdim T\ proof - obtain SB where SB_span: "cspan SB = cspan S" and indep_SB: \cindependent SB\ by (metis complex_vector.basis_exists complex_vector.span_mono complex_vector.span_span subset_antisym) with lin_f inj_f have indep_fSB: \cindependent (f ` SB)\ apply (rule_tac complex_vector.linear_independent_injective_image) by auto from lin_f have \cspan (f ` SB) = f ` cspan SB\ by (meson complex_vector.linear_span_image) also from SB_span lin_f have \\ = cspan T\ by (metis complex_vector.linear_span_image im_S) finally have \cdim T = card (f ` SB)\ using indep_fSB complex_vector.dim_eq_card by blast also have \\ = card SB\ apply (rule card_image) using inj_f by (metis SB_span complex_vector.linear_inj_on_span_iff_independent_image indep_fSB lin_f) also have \\ = cdim S\ using indep_SB SB_span by (metis complex_vector.dim_eq_card) finally show ?thesis by simp qed lemma cindependent_inter_scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 1" shows "B \ (*\<^sub>C) c ` B = {}" proof (rule classical, cases \c = 0\) case True then show ?thesis using a1 by (auto simp add: complex_vector.dependent_zero) next case False assume "\(B \ (*\<^sub>C) c ` B = {})" hence "B \ (*\<^sub>C) c ` B \ {}" by blast then obtain x where u1: "x \ B \ (*\<^sub>C) c ` B" by blast then obtain b where u2: "x = b" and u3: "b\B" by blast then obtain b' where u2': "x = c *\<^sub>C b'" and u3': "b'\B" using u1 by blast have g1: "b = c *\<^sub>C b'" using u2 and u2' by simp hence "b \ complex_vector.span {b'}" using False by (simp add: complex_vector.span_base complex_vector.span_scale) hence "b = b'" by (metis u3' a1 complex_vector.dependent_def complex_vector.span_base complex_vector.span_scale insertE insert_Diff u2 u2' u3) hence "b' = c *\<^sub>C b'" using g1 by blast thus ?thesis by (metis a1 a3 complex_vector.dependent_zero complex_vector.scale_right_imp_eq mult_cancel_right2 scaleC_scaleC u3') qed lemma real_independent_from_complex_independent: assumes "cindependent (B::'a::complex_vector set)" defines "B' == ((*\<^sub>C) \ ` B)" shows "independent (B \ B')" proof (rule notI) assume \dependent (B \ B')\ then obtain T f0 x where [simp]: \finite T\ and \T \ B \ B'\ and f0_sum: \(\v\T. f0 v *\<^sub>R v) = 0\ and x: \x \ T\ and f0_x: \f0 x \ 0\ by (auto simp: real_vector.dependent_explicit) define f T1 T2 T' f' x' where \f v = (if v \ T then f0 v else 0)\ and \T1 = T \ B\ and \T2 = scaleC (-\) ` (T \ B')\ and \T' = T1 \ T2\ and \f' v = f v + \ * f (\ *\<^sub>C v)\ and \x' = (if x \ T1 then x else -\ *\<^sub>C x)\ for v have \B \ B' = {}\ by (simp add: assms cindependent_inter_scaleC_cindependent) have \T' \ B\ by (auto simp: T'_def T1_def T2_def B'_def) have [simp]: \finite T'\ \finite T1\ \finite T2\ by (auto simp add: T'_def T1_def T2_def) have f_sum: \(\v\T. f v *\<^sub>R v) = 0\ unfolding f_def using f0_sum by auto have f_x: \f x \ 0\ using f0_x x by (auto simp: f_def) have f'_sum: \(\v\T'. f' v *\<^sub>C v) = 0\ proof - have \(\v\T'. f' v *\<^sub>C v) = (\v\T'. complex_of_real (f v) *\<^sub>C v) + (\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ by (auto simp: f'_def sum.distrib scaleC_add_left) also have \(\v\T'. complex_of_real (f v) *\<^sub>C v) = (\v\T1. f v *\<^sub>R v)\ (is \_ = ?left\) apply (auto simp: T'_def scaleR_scaleC intro!: sum.mono_neutral_cong_right) using T'_def T1_def \T' \ B\ f_def by auto also have \(\v\T'. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v) = (\v\T2. (\ * complex_of_real (f (\ *\<^sub>C v))) *\<^sub>C v)\ (is \_ = ?right\) apply (auto simp: T'_def intro!: sum.mono_neutral_cong_right) by (smt (z3) B'_def IntE IntI T1_def T2_def \f \ \v. if v \ T then f0 v else 0\ add.inverse_inverse complex_vector.vector_space_axioms i_squared imageI mult_minus_left vector_space.vector_space_assms(3) vector_space.vector_space_assms(4)) also have \?right = (\v\T\B'. f v *\<^sub>R v)\ (is \_ = ?right\) apply (rule sum.reindex_cong[symmetric, where l=\scaleC \\]) apply (auto simp: T2_def image_image scaleR_scaleC) using inj_on_def by fastforce also have \?left + ?right = (\v\T. f v *\<^sub>R v)\ apply (subst sum.union_disjoint[symmetric]) using \B \ B' = {}\ \T \ B \ B'\ apply (auto simp: T1_def) by (metis Int_Un_distrib Un_Int_eq(4) sup.absorb_iff1) also have \\ = 0\ by (rule f_sum) finally show ?thesis by - qed have x': \x' \ T'\ using \T \ B \ B'\ x by (auto simp: x'_def T'_def T1_def T2_def) have f'_x': \f' x' \ 0\ using Complex_eq Complex_eq_0 f'_def f_x x'_def by auto from \finite T'\ \T' \ B\ f'_sum x' f'_x' have \cdependent B\ using complex_vector.independent_explicit_module by blast with assms show False by auto qed lemma crepresentation_from_representation: assumes a1: "cindependent B" and a2: "b \ B" and a3: "finite B" shows "crepresentation B \ b = (representation (B \ (*\<^sub>C) \ ` B) \ b) + \ *\<^sub>C (representation (B \ (*\<^sub>C) \ ` B) \ (\ *\<^sub>C b))" proof (cases "\ \ cspan B") define B' where "B' = B \ (*\<^sub>C) \ ` B" case True define r where "r v = real_vector.representation B' \ v" for v define r' where "r' v = real_vector.representation B' \ (\ *\<^sub>C v)" for v define f where "f v = r v + \ *\<^sub>C r' v" for v define g where "g v = crepresentation B \ v" for v have "(\v | g v \ 0. g v *\<^sub>C v) = \" unfolding g_def using Collect_cong Collect_mono_iff DiffD1 DiffD2 True a1 complex_vector.finite_representation complex_vector.sum_nonzero_representation_eq sum.mono_neutral_cong_left by fastforce moreover have "finite {v. g v \ 0}" unfolding g_def by (simp add: complex_vector.finite_representation) moreover have "v \ B" if "g v \ 0" for v using that unfolding g_def by (simp add: complex_vector.representation_ne_zero) ultimately have rep1: "(\v\B. g v *\<^sub>C v) = \" unfolding g_def using a3 True a1 complex_vector.sum_representation_eq by blast have l0': "inj ((*\<^sub>C) \::'a \'a)" unfolding inj_def by simp have l0: "inj ((*\<^sub>C) (- \)::'a \'a)" unfolding inj_def by simp have l1: "(*\<^sub>C) (- \) ` B \ B = {}" using cindependent_inter_scaleC_cindependent[where B=B and c = "- \"] by (metis Int_commute a1 add.inverse_inverse complex_i_not_one i_squared mult_cancel_left1 neg_equal_0_iff_equal) have l2: "B \ (*\<^sub>C) \ ` B = {}" by (simp add: a1 cindependent_inter_scaleC_cindependent) have rr1: "r (\ *\<^sub>C v) = r' v" for v unfolding r_def r'_def by simp have k1: "independent B'" unfolding B'_def using a1 real_independent_from_complex_independent by simp have "\ \ span B'" using B'_def True cspan_as_span by blast have "v \ B'" if "r v \ 0" for v unfolding r_def using r_def real_vector.representation_ne_zero that by auto have "finite B'" unfolding B'_def using a3 by simp have "(\v\B'. r v *\<^sub>R v) = \" unfolding r_def using True Real_Vector_Spaces.real_vector.sum_representation_eq[where B = B' and basis = B' and v = \] by (smt Real_Vector_Spaces.dependent_raw_def \\ \ Real_Vector_Spaces.span B'\ \finite B'\ equalityD2 k1) have d1: "(\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v)) = (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using l0' by (metis (mono_tags, lifting) inj_eq inj_on_def sum.reindex_cong) have "(\v\B. (r v + \ * (r' v)) *\<^sub>C v) = (\v\B. r v *\<^sub>C v + (\ * r' v) *\<^sub>C v)" by (meson scaleC_left.add) also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. (\ * r' v) *\<^sub>C v)" using sum.distrib by fastforce also have "\ = (\v\B. r v *\<^sub>C v) + (\v\B. \ *\<^sub>C (r' v *\<^sub>C v))" by auto also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. \ *\<^sub>C (r (\ *\<^sub>C v) *\<^sub>R v))" unfolding r'_def r_def by (metis (mono_tags, lifting) scaleR_scaleC sum.cong) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\B. r (\ *\<^sub>C v) *\<^sub>R (\ *\<^sub>C v))" by (metis (no_types, lifting) complex_vector.scale_left_commute scaleR_scaleC) also have "\ = (\v\B. r v *\<^sub>R v) + (\v\(*\<^sub>C) \ ` B. r v *\<^sub>R v)" using d1 by simp also have "\ = \" using l2 \(\v\B'. r v *\<^sub>R v) = \\ unfolding B'_def by (simp add: a3 sum.union_disjoint) finally have "(\v\B. f v *\<^sub>C v) = \" unfolding r'_def r_def f_def by simp hence "0 = (\v\B. f v *\<^sub>C v) - (\v\B. crepresentation B \ v *\<^sub>C v)" using rep1 unfolding g_def by simp also have "\ = (\v\B. f v *\<^sub>C v - crepresentation B \ v *\<^sub>C v)" by (simp add: sum_subtractf) also have "\ = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)" by (metis scaleC_left.diff) finally have "0 = (\v\B. (f v - crepresentation B \ v) *\<^sub>C v)". hence "(\v\B. (f v - crepresentation B \ v) *\<^sub>C v) = 0" by simp hence "f b - crepresentation B \ b = 0" using a1 a2 a3 complex_vector.independentD[where s = B and t = B and u = "\v. f v - crepresentation B \ v" and v = b] order_refl by smt hence "crepresentation B \ b = f b" by simp thus ?thesis unfolding f_def r_def r'_def B'_def by auto next define B' where "B' = B \ (*\<^sub>C) \ ` B" case False have b2: "\ \ real_vector.span B'" unfolding B'_def using False cspan_as_span by auto have "\ \ complex_vector.span B" using False by blast have "crepresentation B \ b = 0" unfolding complex_vector.representation_def by (simp add: False) moreover have "real_vector.representation B' \ b = 0" unfolding real_vector.representation_def by (simp add: b2) moreover have "real_vector.representation B' \ ((*\<^sub>C) \ b) = 0" unfolding real_vector.representation_def by (simp add: b2) ultimately show ?thesis unfolding B'_def by simp qed lemma CARD_1_vec_0[simp]: \(\ :: _ ::{complex_vector,CARD_1}) = 0\ by auto lemma scaleC_cindependent: assumes a1: "cindependent (B::'a::complex_vector set)" and a3: "c \ 0" shows "cindependent ((*\<^sub>C) c ` B)" proof- have "u y = 0" if g1: "y\S" and g2: "(\x\S. u x *\<^sub>C x) = 0" and g3: "finite S" and g4: "S\(*\<^sub>C) c ` B" for u y S proof- define v where "v x = u (c *\<^sub>C x)" for x obtain S' where "S'\B" and S_S': "S = (*\<^sub>C) c ` S'" by (meson g4 subset_imageE) have "inj ((*\<^sub>C) c::'a\_)" unfolding inj_def using a3 by auto hence "finite S'" using S_S' finite_imageD g3 subset_inj_on by blast have "t \ (*\<^sub>C) (inverse c) ` S" if "t \ S'" for t proof- have "c *\<^sub>C t \ S" using \S = (*\<^sub>C) c ` S'\ that by blast hence "(inverse c) *\<^sub>C (c *\<^sub>C t) \ (*\<^sub>C) (inverse c) ` S" by blast moreover have "(inverse c) *\<^sub>C (c *\<^sub>C t) = t" by (simp add: a3) ultimately show ?thesis by simp qed moreover have "t \ S'" if "t \ (*\<^sub>C) (inverse c) ` S" for t proof- obtain t' where "t = (inverse c) *\<^sub>C t'" and "t' \ S" using \t \ (*\<^sub>C) (inverse c) ` S\ by auto have "c *\<^sub>C t = c *\<^sub>C ((inverse c) *\<^sub>C t')" using \t = (inverse c) *\<^sub>C t'\ by simp also have "\ = (c * (inverse c)) *\<^sub>C t'" by simp also have "\ = t'" by (simp add: a3) finally have "c *\<^sub>C t = t'". thus ?thesis using \t' \ S\ using \S = (*\<^sub>C) c ` S'\ a3 complex_vector.scale_left_imp_eq by blast qed ultimately have "S' = (*\<^sub>C) (inverse c) ` S" by blast hence "inverse c *\<^sub>C y \ S'" using that(1) by blast have t: "inj (((*\<^sub>C) c)::'a \ _)" using a3 complex_vector.injective_scale[where c = c] by blast have "0 = (\x\(*\<^sub>C) c ` S'. u x *\<^sub>C x)" using \S = (*\<^sub>C) c ` S'\ that(2) by auto also have "\ = (\x\S'. v x *\<^sub>C (c *\<^sub>C x))" unfolding v_def using t Groups_Big.comm_monoid_add_class.sum.reindex[where h = "((*\<^sub>C) c)" and A = S' and g = "\x. u x *\<^sub>C x"] subset_inj_on by auto also have "\ = c *\<^sub>C (\x\S'. v x *\<^sub>C x)" by (metis (mono_tags, lifting) complex_vector.scale_left_commute scaleC_right.sum sum.cong) finally have "0 = c *\<^sub>C (\x\S'. v x *\<^sub>C x)". hence "(\x\S'. v x *\<^sub>C x) = 0" using a3 by auto hence "v (inverse c *\<^sub>C y) = 0" using \inverse c *\<^sub>C y \ S'\ \finite S'\ \S' \ B\ a1 complex_vector.independentD by blast thus "u y = 0" unfolding v_def by (simp add: a3) qed thus ?thesis using complex_vector.dependent_explicit by (simp add: complex_vector.dependent_explicit ) qed lemma cspan_eqI: assumes \\a. a\A \ a\cspan B\ assumes \\b. b\B \ b\cspan A\ shows \cspan A = cspan B\ apply (rule complex_vector.span_subspace[rotated]) apply (rule complex_vector.span_minimal) using assms by auto lemma (in bounded_cbilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by standard subsection \Antilinear maps and friends\ locale antilinear = additive f for f :: "'a::complex_vector \ 'b::complex_vector" + assumes scaleC: "f (scaleC r x) = cnj r *\<^sub>C f x" sublocale antilinear \ linear proof (rule linearI) show "f (b1 + b2) = f b1 + f b2" for b1 :: 'a and b2 :: 'a by (simp add: add) show "f (r *\<^sub>R b) = r *\<^sub>R f b" for r :: real and b :: 'a unfolding scaleR_scaleC by (subst scaleC, simp) qed lemma antilinear_imp_scaleC: fixes D :: "complex \ 'a::complex_vector" assumes "antilinear D" obtains d where "D = (\x. cnj x *\<^sub>C d)" proof - interpret clinear "D o cnj" apply standard apply auto apply (simp add: additive.add assms antilinear.axioms(1)) using assms antilinear.scaleC by fastforce obtain d where "D o cnj = (\x. x *\<^sub>C d)" using clinear_axioms complex_vector.linear_imp_scale by blast then have \D = (\x. cnj x *\<^sub>C d)\ by (metis comp_apply complex_cnj_cnj) then show ?thesis by (rule that) qed corollary complex_antilinearD: fixes f :: "complex \ complex" assumes "antilinear f" obtains c where "f = (\x. c * cnj x)" by (rule antilinear_imp_scaleC [OF assms]) (force simp: scaleC_conv_of_complex) lemma antilinearI: assumes "\x y. f (x + y) = f x + f y" and "\c x. f (c *\<^sub>C x) = cnj c *\<^sub>C f x" shows "antilinear f" by standard (rule assms)+ lemma antilinear_o_antilinear: "antilinear f \ antilinear g \ clinear (g o f)" apply (rule clinearI) apply (simp add: additive.add antilinear_def) by (simp add: antilinear.scaleC) lemma clinear_o_antilinear: "antilinear f \ clinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) lemma antilinear_o_clinear: "clinear f \ antilinear g \ antilinear (g o f)" apply (rule antilinearI) apply (simp add: additive.add complex_vector.linear_add antilinear_def) by (simp add: complex_vector.linear_scale antilinear.scaleC) locale bounded_antilinear = antilinear f for f :: "'a::complex_normed_vector \ 'b::complex_normed_vector" + assumes bounded: "\K. \x. norm (f x) \ norm x * K" lemma bounded_antilinearI: assumes \\b1 b2. f (b1 + b2) = f b1 + f b2\ assumes \\r b. f (r *\<^sub>C b) = cnj r *\<^sub>C f b\ assumes \\x. norm (f x) \ norm x * K\ shows "bounded_antilinear f" using assms by (auto intro!: exI bounded_antilinear.intro antilinearI simp: bounded_antilinear_axioms_def) sublocale bounded_antilinear \ real: bounded_linear \ \Gives access to all lemmas from \<^locale>\linear\ using prefix \real.\\ apply standard by (fact bounded) lemma (in bounded_antilinear) bounded_linear: "bounded_linear f" by (fact real.bounded_linear) lemma (in bounded_antilinear) antilinear: "antilinear f" by (fact antilinear_axioms) lemma bounded_antilinear_intro: assumes "\x y. f (x + y) = f x + f y" and "\r x. f (scaleC r x) = scaleC (cnj r) (f x)" and "\x. norm (f x) \ norm x * K" shows "bounded_antilinear f" by standard (blast intro: assms)+ lemma bounded_antilinear_0[simp]: \bounded_antilinear (\_. 0)\ by (rule bounded_antilinear_intro[where K=0], auto) lemma cnj_bounded_antilinear[simp]: "bounded_antilinear cnj" apply (rule bounded_antilinear_intro [where K = 1]) by auto lemma bounded_antilinear_o_bounded_antilinear: assumes "bounded_antilinear f" and "bounded_antilinear g" shows "bounded_clinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_antilinear g by fact fix b1 b2 b r show "f (g (b1 + b2)) = f (g b1) + f (g b2)" by (simp add: f.add g.add) show "f (g (r *\<^sub>C b)) = r *\<^sub>C f (g b)" by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_antilinear_o_bounded_clinear: assumes "bounded_antilinear f" and "bounded_clinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_antilinear f by fact interpret g: bounded_clinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x by (simp add: f.scaleC g.scaleC) have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bounded_clinear_o_bounded_antilinear: assumes "bounded_clinear f" and "bounded_antilinear g" shows "bounded_antilinear (\x. f (g x))" proof interpret f: bounded_clinear f by fact interpret g: bounded_antilinear g by fact show "f (g (x + y)) = f (g x) + f (g y)" for x y by (simp only: f.add g.add) show "f (g (scaleC r x)) = scaleC (cnj r) (f (g x))" for r x using f.scaleC g.scaleC by fastforce have "bounded_linear (\x. f (g x))" using f.bounded_linear g.bounded_linear by (rule bounded_linear_compose) then show "\K. \x. norm (f (g x)) \ norm x * K" by (rule bounded_linear.bounded) qed lemma bij_clinear_imp_inv_clinear: "clinear (inv f)" if a1: "clinear f" and a2: "bij f" proof fix b1 b2 r b show "inv f (b1 + b2) = inv f b1 + inv f b2" by (simp add: a1 a2 bij_is_inj bij_is_surj complex_vector.linear_add inv_f_eq surj_f_inv_f) show "inv f (r *\<^sub>C b) = r *\<^sub>C inv f b" using that by (smt bij_inv_eq_iff clinear_def complex_vector.linear_scale) qed locale bounded_sesquilinear = fixes prod :: "'a::complex_normed_vector \ 'b::complex_normed_vector \ 'c::complex_normed_vector" (infixl "**" 70) assumes add_left: "prod (a + a') b = prod a b + prod a' b" and add_right: "prod a (b + b') = prod a b + prod a b'" and scaleC_left: "prod (r *\<^sub>C a) b = (cnj r) *\<^sub>C (prod a b)" and scaleC_right: "prod a (r *\<^sub>C b) = r *\<^sub>C (prod a b)" and bounded: "\K. \a b. norm (prod a b) \ norm a * norm b * K" sublocale bounded_sesquilinear \ real: bounded_bilinear \ \Gives access to all lemmas from \<^locale>\linear\ using prefix \real.\\ apply standard by (auto simp: add_left add_right scaleC_left scaleC_right bounded scaleR_scaleC) lemma (in bounded_sesquilinear) bounded_bilinear[simp]: "bounded_bilinear prod" by intro_locales lemma (in bounded_sesquilinear) bounded_antilinear_left: "bounded_antilinear (\a. prod a b)" apply standard apply (auto simp add: scaleC_left add_left) by (metis ab_semigroup_mult_class.mult_ac(1) bounded) lemma (in bounded_sesquilinear) bounded_clinear_right: "bounded_clinear (\b. prod a b)" apply standard apply (auto simp add: scaleC_right add_right) by (metis ab_semigroup_mult_class.mult_ac(1) ordered_field_class.sign_simps(34) real.pos_bounded) lemma (in bounded_sesquilinear) comp1: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x. prod (g x))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (g (a + a')) b = prod (g a) b + prod (g a') b" by (simp add: add add_left) show "prod (g a) (b + b') = prod (g a) b + prod (g a) b'" by (simp add: add add_right) show "prod (g (r *\<^sub>C a)) b = cnj r *\<^sub>C prod (g a) b" by (simp add: scaleC scaleC_left) show "prod (g a) (r *\<^sub>C b) = r *\<^sub>C prod (g a) b" by (simp add: scaleC_right) interpret bi: bounded_bilinear \(\x. prod (g x))\ by (simp add: bounded_linear real.comp1) show "\K. \a b. norm (prod (g a) b) \ norm a * norm b * K" using bi.bounded by blast qed lemma (in bounded_sesquilinear) comp2: assumes \bounded_clinear g\ shows \bounded_sesquilinear (\x y. prod x (g y))\ proof interpret bounded_clinear g by fact fix a a' b b' r show "prod (a + a') (g b) = prod a (g b) + prod a' (g b)" by (simp add: add add_left) show "prod a (g (b + b')) = prod a (g b) + prod a (g b')" by (simp add: add add_right) show "prod (r *\<^sub>C a) (g b) = cnj r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_left) show "prod a (g (r *\<^sub>C b)) = r *\<^sub>C prod a (g b)" by (simp add: scaleC scaleC_right) interpret bi: bounded_bilinear \(\x y. prod x (g y))\ apply (rule bounded_bilinear.flip) using _ bounded_linear apply (rule bounded_bilinear.comp1) using bounded_bilinear by (rule bounded_bilinear.flip) show "\K. \a b. norm (prod a (g b)) \ norm a * norm b * K" using bi.bounded by blast qed lemma (in bounded_sesquilinear) comp: "bounded_clinear f \ bounded_clinear g \ bounded_sesquilinear (\x y. prod (f x) (g y))" using comp1 bounded_sesquilinear.comp2 by auto lemma bounded_clinear_const_scaleR: fixes c :: real assumes \bounded_clinear f\ shows \bounded_clinear (\ x. c *\<^sub>R f x )\ proof- have \bounded_clinear (\ x. (complex_of_real c) *\<^sub>C f x )\ by (simp add: assms bounded_clinear_const_scaleC) thus ?thesis by (simp add: scaleR_scaleC) qed lemma bounded_linear_bounded_clinear: \bounded_linear A \ \c x. A (c *\<^sub>C x) = c *\<^sub>C A x \ bounded_clinear A\ apply standard by (simp_all add: linear_simps bounded_linear.bounded) lemma comp_bounded_clinear: fixes A :: \'b::complex_normed_vector \ 'c::complex_normed_vector\ and B :: \'a::complex_normed_vector \ 'b\ assumes \bounded_clinear A\ and \bounded_clinear B\ shows \bounded_clinear (A \ B)\ by (metis clinear_compose assms(1) assms(2) bounded_clinear_axioms_def bounded_clinear_compose bounded_clinear_def o_def) lemmas isCont_scaleC [simp] = bounded_bilinear.isCont [OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] subsection \Misc 2\ lemma summable_on_scaleC_left [intro]: fixes c :: \'a :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. f x *\<^sub>C c) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>C c) = (\y. y *\<^sub>C c) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleC_left.additive_axioms) lemma summable_on_scaleC_right [intro]: fixes f :: \'a \ 'b :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "(\x. c *\<^sub>C f x) summable_on A" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. c *\<^sub>C f x) = (\y. c *\<^sub>C y) o f\], simp add: o_def) apply (rule summable_on_comm_additive) using assms by (auto simp add: scaleC_right.additive_axioms) lemma infsum_scaleC_left: fixes c :: \'a :: complex_normed_vector\ assumes "c \ 0 \ f summable_on A" shows "infsum (\x. f x *\<^sub>C c) A = infsum f A *\<^sub>C c" apply (cases \c \ 0\) apply (subst asm_rl[of \(\x. f x *\<^sub>C c) = (\y. y *\<^sub>C c) o f\], simp add: o_def) apply (rule infsum_comm_additive) using assms by (auto simp add: scaleC_left.additive_axioms) lemma infsum_scaleC_right: fixes f :: \'a \ 'b :: complex_normed_vector\ shows "infsum (\x. c *\<^sub>C f x) A = c *\<^sub>C infsum f A" proof - consider (summable) \f summable_on A\ | (c0) \c = 0\ | (not_summable) \\ f summable_on A\ \c \ 0\ by auto then show ?thesis proof cases case summable then show ?thesis apply (subst asm_rl[of \(\x. c *\<^sub>C f x) = (\y. c *\<^sub>C y) o f\], simp add: o_def) apply (rule infsum_comm_additive) using summable by (auto simp add: scaleC_right.additive_axioms) next case c0 then show ?thesis by auto next case not_summable have \\ (\x. c *\<^sub>C f x) summable_on A\ proof (rule notI) assume \(\x. c *\<^sub>C f x) summable_on A\ then have \(\x. inverse c *\<^sub>C c *\<^sub>C f x) summable_on A\ using summable_on_scaleC_right by blast then have \f summable_on A\ using not_summable by auto with not_summable show False by simp qed then show ?thesis by (simp add: infsum_not_exists not_summable(1)) qed qed lemmas sums_of_complex = bounded_linear.sums [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas summable_of_complex = bounded_linear.summable [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas suminf_of_complex = bounded_linear.suminf [OF bounded_clinear_of_complex[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_left = bounded_linear.sums[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_left = bounded_linear.summable[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_left = bounded_linear.suminf[OF bounded_clinear_scaleC_left[THEN bounded_clinear.bounded_linear]] lemmas sums_scaleC_right = bounded_linear.sums[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas summable_scaleC_right = bounded_linear.summable[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemmas suminf_scaleC_right = bounded_linear.suminf[OF bounded_clinear_scaleC_right[THEN bounded_clinear.bounded_linear]] lemma closed_scaleC: fixes S::\'a::complex_normed_vector set\ and a :: complex assumes \closed S\ shows \closed ((*\<^sub>C) a ` S)\ proof (cases \a = 0\) case True then show ?thesis apply (cases \S = {}\) by (auto simp: image_constant) next case False then have \(*\<^sub>C) a ` S = (*\<^sub>C) (inverse a) -` S\ by (auto simp add: rev_image_eqI) moreover have \closed ((*\<^sub>C) (inverse a) -` S)\ by (simp add: assms continuous_closed_vimage) ultimately show ?thesis by simp qed lemma closure_scaleC: fixes S::\'a::complex_normed_vector set\ shows \closure ((*\<^sub>C) a ` S) = (*\<^sub>C) a ` closure S\ proof have \closed (closure S)\ by simp show "closure ((*\<^sub>C) a ` S) \ (*\<^sub>C) a ` closure S" by (simp add: closed_scaleC closure_minimal closure_subset image_mono) have "x \ closure ((*\<^sub>C) a ` S)" if "x \ (*\<^sub>C) a ` closure S" for x :: 'a proof- obtain t where \x = ((*\<^sub>C) a) t\ and \t \ closure S\ using \x \ (*\<^sub>C) a ` closure S\ by auto have \\s. (\n. s n \ S) \ s \ t\ using \t \ closure S\ Elementary_Topology.closure_sequential by blast then obtain s where \\n. s n \ S\ and \s \ t\ by blast have \(\ n. scaleC a (s n) \ ((*\<^sub>C) a ` S))\ using \\n. s n \ S\ by blast moreover have \(\ n. scaleC a (s n)) \ x\ proof- have \isCont (scaleC a) t\ by simp thus ?thesis using \s \ t\ \x = ((*\<^sub>C) a) t\ by (simp add: isCont_tendsto_compose) qed ultimately show ?thesis using Elementary_Topology.closure_sequential by metis qed thus "(*\<^sub>C) a ` closure S \ closure ((*\<^sub>C) a ` S)" by blast qed lemma onorm_scalarC: fixes f :: \'a::complex_normed_vector \ 'b::complex_normed_vector\ assumes a1: \bounded_clinear f\ shows \onorm (\ x. r *\<^sub>C (f x)) = (cmod r) * onorm f\ proof- have \(norm (f x)) / norm x \ onorm f\ for x using a1 by (simp add: bounded_clinear.bounded_linear le_onorm) hence t2: \bdd_above {(norm (f x)) / norm x | x. True}\ by fastforce have \continuous_on UNIV ( (*) w ) \ for w::real by simp hence \isCont ( ((*) (cmod r)) ) x\ for x by simp hence t3: \continuous (at_left (Sup {(norm (f x)) / norm x | x. True})) ((*) (cmod r))\ using Elementary_Topology.continuous_at_imp_continuous_within by blast have \{(norm (f x)) / norm x | x. True} \ {}\ by blast moreover have \mono ((*) (cmod r))\ by (simp add: monoI ordered_comm_semiring_class.comm_mult_left_mono) ultimately have \Sup {((*) (cmod r)) ((norm (f x)) / norm x) | x. True} = ((*) (cmod r)) (Sup {(norm (f x)) / norm x | x. True})\ using t2 t3 by (simp add: continuous_at_Sup_mono full_SetCompr_eq image_image) hence \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (cmod r) * (Sup {(norm (f x)) / norm x | x. True})\ by blast moreover have \Sup {(cmod r) * ((norm (f x)) / norm x) | x. True} = (SUP x. cmod r * norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) moreover have \(Sup {(norm (f x)) / norm x | x. True}) = (SUP x. norm (f x) / norm x)\ by (simp add: full_SetCompr_eq) ultimately have t1: "(SUP x. cmod r * norm (f x) / norm x) = cmod r * (SUP x. norm (f x) / norm x)" by simp have \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. norm ( (\ t. r *\<^sub>C (f t)) x) / norm x)\ by (simp add: onorm_def) hence \onorm (\ x. r *\<^sub>C (f x)) = (SUP x. (cmod r) * (norm (f x)) / norm x)\ by simp also have \... = (cmod r) * (SUP x. (norm (f x)) / norm x)\ using t1. finally show ?thesis by (simp add: onorm_def) qed lemma onorm_scaleC_left_lemma: fixes f :: "'a::complex_normed_vector" assumes r: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" proof (rule onorm_bound) fix x have "norm (r x *\<^sub>C f) = norm (r x) * norm f" by simp also have "\ \ onorm r * norm x * norm f" by (simp add: bounded_clinear.bounded_linear mult.commute mult_left_mono onorm r) finally show "norm (r x *\<^sub>C f) \ onorm r * norm f * norm x" by (simp add: ac_simps) show "0 \ onorm r * norm f" by (simp add: bounded_clinear.bounded_linear onorm_pos_le r) qed lemma onorm_scaleC_left: fixes f :: "'a::complex_normed_vector" assumes f: "bounded_clinear r" shows "onorm (\x. r x *\<^sub>C f) = onorm r * norm f" proof (cases "f = 0") assume "f \ 0" show ?thesis proof (rule order_antisym) show "onorm (\x. r x *\<^sub>C f) \ onorm r * norm f" using f by (rule onorm_scaleC_left_lemma) next have bl1: "bounded_clinear (\x. r x *\<^sub>C f)" by (metis bounded_clinear_scaleC_const f) have x1:"bounded_clinear (\x. r x * norm f)" by (metis bounded_clinear_mult_const f) have "onorm r \ onorm (\x. r x * complex_of_real (norm f)) / norm f" if "onorm r \ onorm (\x. r x * complex_of_real (norm f)) * cmod (1 / complex_of_real (norm f))" and "f \ 0" using that by (metis complex_of_real_cmod complex_of_real_nn_iff field_class.field_divide_inverse inverse_eq_divide nice_ordered_field_class.zero_le_divide_1_iff norm_ge_zero of_real_1 of_real_divide of_real_eq_iff) hence "onorm r \ onorm (\x. r x * norm f) * inverse (norm f)" using \f \ 0\ onorm_scaleC_left_lemma[OF x1, of "inverse (norm f)"] by (simp add: inverse_eq_divide) also have "onorm (\x. r x * norm f) \ onorm (\x. r x *\<^sub>C f)" proof (rule onorm_bound) have "bounded_linear (\x. r x *\<^sub>C f)" using bl1 bounded_clinear.bounded_linear by auto thus "0 \ onorm (\x. r x *\<^sub>C f)" by (rule Operator_Norm.onorm_pos_le) show "cmod (r x * complex_of_real (norm f)) \ onorm (\x. r x *\<^sub>C f) * norm x" for x :: 'b by (smt \bounded_linear (\x. r x *\<^sub>C f)\ complex_of_real_cmod complex_of_real_nn_iff complex_scaleC_def norm_ge_zero norm_scaleC of_real_eq_iff onorm) qed finally show "onorm r * norm f \ onorm (\x. r x *\<^sub>C f)" using \f \ 0\ by (simp add: inverse_eq_divide pos_le_divide_eq mult.commute) qed qed (simp add: onorm_zero) subsection \Finite dimension and canonical basis\ lemma vector_finitely_spanned: assumes \z \ cspan T\ shows \\ S. finite S \ S \ T \ z \ cspan S\ proof- have \\ S r. finite S \ S \ T \ z = (\a\S. r a *\<^sub>C a)\ using complex_vector.span_explicit[where b = "T"] assms by auto then obtain S r where \finite S\ and \S \ T\ and \z = (\a\S. r a *\<^sub>C a)\ by blast thus ?thesis by (meson complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset subset_iff) qed setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a set \ 'a set\)\ class cfinite_dim = complex_vector + assumes cfinitely_spanned: "\S::'a set. finite S \ cspan S = UNIV" class basis_enum = complex_vector + fixes canonical_basis :: "'a list" assumes distinct_canonical_basis[simp]: "distinct canonical_basis" and is_cindependent_set[simp]: "cindependent (set canonical_basis)" and is_generator_set[simp]: "cspan (set canonical_basis) = UNIV" setup \Sign.add_const_constraint ("Complex_Vector_Spaces0.cindependent", SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cdependent\, SOME \<^typ>\'a::complex_vector set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\cspan\, SOME \<^typ>\'a::complex_vector set \ 'a set\)\ lemma cdim_UNIV_basis_enum[simp]: \cdim (UNIV::'a::basis_enum set) = length (canonical_basis::'a list)\ apply (subst is_generator_set[symmetric]) apply (subst complex_vector.dim_span_eq_card_independent) apply (rule is_cindependent_set) using distinct_canonical_basis distinct_card by blast lemma finite_basis: "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = UNIV" proof - from cfinitely_spanned obtain S :: \'a set\ where \finite S\ and \cspan S = UNIV\ by auto from complex_vector.maximal_independent_subset obtain B :: \'a set\ where \B \ S\ and \cindependent B\ and \S \ cspan B\ by metis moreover have \finite B\ using \B \ S\ \finite S\ by (meson finite_subset) moreover have \cspan B = UNIV\ using \cspan S = UNIV\ \S \ cspan B\ by (metis complex_vector.span_eq top_greatest) ultimately show ?thesis by auto qed instance basis_enum \ cfinite_dim apply intro_classes apply (rule exI[of _ \set canonical_basis\]) using is_cindependent_set is_generator_set by auto lemma cindependent_cfinite_dim_finite: assumes \cindependent (S::'a::cfinite_dim set)\ shows \finite S\ by (metis assms cfinitely_spanned complex_vector.independent_span_bound top_greatest) lemma cfinite_dim_finite_subspace_basis: assumes \csubspace X\ shows "\basis::'a::cfinite_dim set. finite basis \ cindependent basis \ cspan basis = X" by (meson assms cindependent_cfinite_dim_finite complex_vector.basis_exists complex_vector.span_subspace) text \The following auxiliary lemma (\finite_span_complete_aux\) shows more or less the same as \finite_span_representation_bounded\, \finite_span_complete\ below (see there for an intuition about the mathematical content of the lemmas). However, there is one difference: Here we additionally assume here that there is a bijection rep/abs between a finite type \<^typ>\'basis\ and the set $B$. This is needed to be able to use results about euclidean spaces that are formulated w.r.t. the type class \<^class>\finite\ Since we anyway assume that $B$ is finite, this added assumption does not make the lemma weaker. However, we cannot derive the existence of \<^typ>\'basis\ inside the proof (HOL does not support such reasoning). Therefore we have the type \<^typ>\'basis\ as an explicit assumption and remove it using @{attribute internalize_sort} after the proof.\ lemma finite_span_complete_aux: fixes b :: "'b::real_normed_vector" and B :: "'b set" and rep :: "'basis::finite \ 'b" and abs :: "'b \ 'basis" assumes t: "type_definition rep abs B" and t1: "finite B" and t2: "b\B" and t3: "independent B" shows "\D>0. \\. norm (representation B \ b) \ norm \ * D" and "complete (span B)" proof - define repr where "repr = real_vector.representation B" define repr' where "repr' \ = Abs_euclidean_space (repr \ o rep)" for \ define comb where "comb l = (\b\B. l b *\<^sub>R b)" for l define comb' where "comb' l = comb (Rep_euclidean_space l o abs)" for l have comb_cong: "comb x = comb y" if "\z. z\B \ x z = y z" for x y unfolding comb_def using that by auto have comb_repr[simp]: "comb (repr \) = \" if "\ \ real_vector.span B" for \ using \comb \ \l. \b\B. l b *\<^sub>R b\ local.repr_def real_vector.sum_representation_eq t1 t3 that by fastforce have w5:"(\b | (b \ B \ x b \ 0) \ b \ B. x b *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" for x using \finite B\ by (smt DiffD1 DiffD2 mem_Collect_eq real_vector.scale_eq_0_iff subset_eq sum.mono_neutral_left) have "representation B (\b\B. x b *\<^sub>R b) = (\b. if b \ B then x b else 0)" for x proof (rule real_vector.representation_eqI) show "independent B" by (simp add: t3) show "(\b\B. x b *\<^sub>R b) \ span B" by (meson real_vector.span_scale real_vector.span_sum real_vector.span_superset subset_iff) show "b \ B" if "(if b \ B then x b else 0) \ 0" for b :: 'b using that by meson show "finite {b. (if b \ B then x b else 0) \ 0}" using t1 by auto show "(\b | (if b \ B then x b else 0) \ 0. (if b \ B then x b else 0) *\<^sub>R b) = (\b\B. x b *\<^sub>R b)" using w5 by simp qed hence repr_comb[simp]: "repr (comb x) = (\b. if b\B then x b else 0)" for x unfolding repr_def comb_def. have repr_bad[simp]: "repr \ = (\_. 0)" if "\ \ real_vector.span B" for \ unfolding repr_def using that by (simp add: real_vector.representation_def) have [simp]: "repr' \ = 0" if "\ \ real_vector.span B" for \ unfolding repr'_def repr_bad[OF that] apply transfer by auto have comb'_repr'[simp]: "comb' (repr' \) = \" if "\ \ real_vector.span B" for \ proof - have x1: "(repr \ \ rep \ abs) z = repr \ z" if "z \ B" for z unfolding o_def using t that type_definition.Abs_inverse by fastforce have "comb' (repr' \) = comb ((repr \ \ rep) \ abs)" unfolding comb'_def repr'_def by (subst Abs_euclidean_space_inverse; simp) also have "\ = comb (repr \)" using x1 comb_cong by blast also have "\ = \" using that by simp finally show ?thesis by - qed have t1: "Abs_euclidean_space (Rep_euclidean_space t) = t" if "\x. rep x \ B" for t::"'a euclidean_space" apply (subst Rep_euclidean_space_inverse) by simp have "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x y else 0) = x" for x using type_definition.Rep[OF t] apply simp using t1 by blast hence "Abs_euclidean_space (\y. if rep y \ B then Rep_euclidean_space x (abs (rep y)) else 0) = x" for x apply (subst type_definition.Rep_inverse[OF t]) by simp hence repr'_comb'[simp]: "repr' (comb' x) = x" for x unfolding comb'_def repr'_def o_def by simp have sphere: "compact (sphere 0 d :: 'basis euclidean_space set)" for d using compact_sphere by blast have "complete (UNIV :: 'basis euclidean_space set)" by (simp add: complete_UNIV) have "(\b\B. (Rep_euclidean_space (x + y) \ abs) b *\<^sub>R b) = (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b) + (\b\B. (Rep_euclidean_space y \ abs) b *\<^sub>R b)" for x :: "'basis euclidean_space" and y :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: scaleR_add_left sum.distrib) moreover have "(\b\B. (Rep_euclidean_space (c *\<^sub>R x) \ abs) b *\<^sub>R b) = c *\<^sub>R (\b\B. (Rep_euclidean_space x \ abs) b *\<^sub>R b)" for c :: real and x :: "'basis euclidean_space" apply (transfer fixing: abs) by (simp add: real_vector.scale_sum_right) ultimately have blin_comb': "bounded_linear comb'" unfolding comb_def comb'_def by (rule bounded_linearI') hence "continuous_on X comb'" for X by (simp add: linear_continuous_on) hence "compact (comb' ` sphere 0 d)" for d using sphere by (rule compact_continuous_image) hence compact_norm_comb': "compact (norm ` comb' ` sphere 0 1)" using compact_continuous_image continuous_on_norm_id by blast have not0: "0 \ norm ` comb' ` sphere 0 1" proof (rule ccontr, simp) assume "0 \ norm ` comb' ` sphere 0 1" then obtain x where nc0: "norm (comb' x) = 0" and x: "x \ sphere 0 1" by auto hence "comb' x = 0" by simp hence "repr' (comb' x) = 0" unfolding repr'_def o_def repr_def apply simp by (smt repr'_comb' blin_comb' dist_0_norm linear_simps(3) mem_sphere norm_zero x) hence "x = 0" by auto with x show False by simp qed have "closed (norm ` comb' ` sphere 0 1)" using compact_imp_closed compact_norm_comb' by blast moreover have "0 \ norm ` comb' ` sphere 0 1" by (simp add: not0) ultimately have "\d>0. \x\norm ` comb' ` sphere 0 1. d \ dist 0 x" by (meson separate_point_closed) then obtain d where d: "x\norm ` comb' ` sphere 0 1 \ d \ dist 0 x" and "d > 0" for x by metis define D where "D = 1/d" hence "D > 0" using \d>0\ unfolding D_def by auto have "x \ d" if "x\norm ` comb' ` sphere 0 1" for x using d that apply auto by fastforce hence *: "norm (comb' x) \ d" if "norm x = 1" for x using that by auto have norm_comb': "norm (comb' x) \ d * norm x" for x proof (cases "x=0") show "d * norm x \ norm (comb' x)" if "x = 0" using that by simp show "d * norm x \ norm (comb' x)" if "x \ 0" using that using *[of "(1/norm x) *\<^sub>R x"] unfolding linear_simps(5)[OF blin_comb'] apply auto by (simp add: le_divide_eq) qed have *: "norm (repr' \) \ norm \ * D" for \ proof (cases "\ \ real_vector.span B") show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that unfolding D_def using norm_comb'[of "repr' \"] \d>0\ by (simp_all add: linordered_field_class.mult_imp_le_div_pos mult.commute) show "norm (repr' \) \ norm \ * D" if "\ \ span B" using that \0 < D\ by auto qed hence "norm (Rep_euclidean_space (repr' \) (abs b)) \ norm \ * D" for \ proof - have "(Rep_euclidean_space (repr' \) (abs b)) = repr' \ \ euclidean_space_basis_vector (abs b)" apply (transfer fixing: abs b) by auto also have "\\\ \ norm (repr' \)" apply (rule Basis_le_norm) unfolding Basis_euclidean_space_def by simp also have "\ \ norm \ * D" using * by auto finally show ?thesis by simp qed hence "norm (repr \ b) \ norm \ * D" for \ unfolding repr'_def by (smt \comb' \ \l. comb (Rep_euclidean_space l \ abs)\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' comp_apply norm_le_zero_iff repr_bad repr_comb) thus "\D>0. \\. norm (repr \ b) \ norm \ * D" using \D>0\ by auto from \d>0\ have complete_comb': "complete (comb' ` UNIV)" proof (rule complete_isometric_image) show "subspace (UNIV::'basis euclidean_space set)" by simp show "bounded_linear comb'" by (simp add: blin_comb') show "\x\UNIV. d * norm x \ norm (comb' x)" by (simp add: norm_comb') show "complete (UNIV::'basis euclidean_space set)" by (simp add: \complete UNIV\) qed have range_comb': "comb' ` UNIV = real_vector.span B" proof (auto simp: image_def) show "comb' x \ real_vector.span B" for x by (metis comb'_def comb_cong comb_repr local.repr_def repr_bad repr_comb real_vector.representation_zero real_vector.span_zero) next fix \ assume "\ \ real_vector.span B" then obtain f where f: "comb f = \" apply atomize_elim unfolding span_finite[OF \finite B\] comb_def by auto define f' where "f' b = (if b\B then f b else 0)" for b :: 'b have f': "comb f' = \" unfolding f[symmetric] apply (rule comb_cong) unfolding f'_def by simp define x :: "'basis euclidean_space" where "x = Abs_euclidean_space (f' o rep)" have "\ = comb' x" by (metis (no_types, lifting) \\ \ span B\ \repr' \ \\. Abs_euclidean_space (repr \ \ rep)\ comb'_repr' f' fun.map_cong repr_comb t type_definition.Rep_range x_def) thus "\x. \ = comb' x" by auto qed from range_comb' complete_comb' show "complete (real_vector.span B)" by simp qed lemma finite_span_complete[simp]: fixes A :: "'a::real_normed_vector set" assumes "finite A" shows "complete (span A)" text \The span of a finite set is complete.\ proof (cases "A \ {} \ A \ {0}") case True obtain B where BT: "real_vector.span B = real_vector.span A" and "independent B" and "finite B" by (meson True assms finite_subset real_vector.maximal_independent_subset real_vector.span_eq real_vector.span_superset subset_trans) have "B\{}" apply (rule ccontr, simp) using BT True by (metis real_vector.span_superset real_vector.span_empty subset_singletonD) (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) { (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux, otherwise "internalize_sort" below fails *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) note finite_span_complete_aux(2)[internalize_sort "'basis::finite"] note this[OF basisT_finite t] } note this[cancel_type_definition, OF \B\{}\ \finite B\ _ \independent B\] hence "complete (real_vector.span B)" using \B\{}\ by auto thus "complete (real_vector.span A)" unfolding BT by simp next case False thus ?thesis using complete_singleton by auto qed lemma finite_span_representation_bounded: fixes B :: "'a::real_normed_vector set" assumes "finite B" and "independent B" shows "\D>0. \\ b. abs (representation B \ b) \ norm \ * D" text \ Assume $B$ is a finite linear independent set of vectors (in a real normed vector space). Let $\alpha^\psi_b$ be the coefficients of $\psi$ expressed as a linear combination over $B$. Then $\alpha$ is is uniformly cblinfun (i.e., $\lvert\alpha^\psi_b \leq D \lVert\psi\rVert\psi$ for some $D$ independent of $\psi,b$). (This also holds when $b$ is not in the span of $B$ because of the way \real_vector.representation\ is defined in this corner case.)\ proof (cases "B\{}") case True (* The following generalizes finite_span_complete_aux to hold without the assumption that 'basis has type class finite *) define repr where "repr = real_vector.representation B" { (* Step 1: Create a fake type definition by introducing a new type variable 'basis and then assuming the existence of the morphisms Rep/Abs to B This is then roughly equivalent to "typedef 'basis = B" *) (* The type variable 'basisT must not be the same as the one used in finite_span_complete_aux (I.e., we cannot call it 'basis) *) assume "\(Rep :: 'basisT\'a) Abs. type_definition Rep Abs B" then obtain rep :: "'basisT \ 'a" and abs :: "'a \ 'basisT" where t: "type_definition rep abs B" by auto (* Step 2: We show that our fake typedef 'basisT could be instantiated as type class finite *) have basisT_finite: "class.finite TYPE('basisT)" apply intro_classes using \finite B\ t by (metis (mono_tags, opaque_lifting) ex_new_if_finite finite_imageI image_eqI type_definition_def) (* Step 3: We take the finite_span_complete_aux and remove the requirement that 'basis::finite (instead, a precondition "class.finite TYPE('basisT)" is introduced) *) note finite_span_complete_aux(1)[internalize_sort "'basis::finite"] (* Step 4: We instantiate the premises *) note this[OF basisT_finite t] } (* Now we have the desired fact, except that it still assumes that B is isomorphic to some type 'basis together with the assumption that there are morphisms between 'basis and B. 'basis and that premise are removed using cancel_type_definition *) note this[cancel_type_definition, OF True \finite B\ _ \independent B\] hence d2:"\D. \\. D>0 \ norm (repr \ b) \ norm \ * D" if \b\B\ for b by (simp add: repr_def that True) have d1: " (\b. b \ B \ \D. \\. 0 < D \ norm (repr \ b) \ norm \ * D) \ \D. \b \. b \ B \ 0 < D b \ norm (repr \ b) \ norm \ * D b" apply (rule choice) by auto then obtain D where D: "D b > 0 \ norm (repr \ b) \ norm \ * D b" if "b\B" for b \ apply atomize_elim using d2 by blast hence Dpos: "D b > 0" and Dbound: "norm (repr \ b) \ norm \ * D b" if "b\B" for b \ using that by auto define Dall where "Dall = Max (D`B)" have "Dall > 0" unfolding Dall_def using \finite B\ \B\{}\ Dpos by (metis (mono_tags, lifting) Max_in finite_imageI image_iff image_is_empty) have "Dall \ D b" if "b\B" for b unfolding Dall_def using \finite B\ that by auto with Dbound have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ using that by (smt mult_left_mono norm_not_less_zero) moreover have "norm (repr \ b) \ norm \ * Dall" if "b\B" for b \ unfolding repr_def using real_vector.representation_ne_zero True by (metis calculation empty_subsetI less_le_trans local.repr_def norm_ge_zero norm_zero not_less subsetI subset_antisym) ultimately show "\D>0. \\ b. abs (repr \ b) \ norm \ * D" using \Dall > 0\ real_norm_def by metis next case False thus ?thesis unfolding repr_def using real_vector.representation_ne_zero[of B] using nice_ordered_field_class.linordered_field_no_ub by fastforce qed hide_fact finite_span_complete_aux lemma finite_cspan_complete[simp]: fixes B :: "'a::complex_normed_vector set" assumes "finite B" shows "complete (cspan B)" by (simp add: assms cspan_as_span) lemma finite_span_closed[simp]: fixes B :: "'a::real_normed_vector set" assumes "finite B" shows "closed (real_vector.span B)" by (simp add: assms complete_imp_closed) lemma finite_cspan_closed[simp]: fixes S::\'a::complex_normed_vector set\ assumes a1: \finite S\ shows \closed (cspan S)\ by (simp add: assms complete_imp_closed) lemma closure_finite_cspan: fixes T::\'a::complex_normed_vector set\ assumes \finite T\ shows \closure (cspan T) = cspan T\ by (simp add: assms) lemma finite_cspan_crepresentation_bounded: fixes B :: "'a::complex_normed_vector set" assumes a1: "finite B" and a2: "cindependent B" shows "\D>0. \\ b. norm (crepresentation B \ b) \ norm \ * D" proof - define B' where "B' = (B \ scaleC \ ` B)" have independent_B': "independent B'" using B'_def \cindependent B\ by (simp add: real_independent_from_complex_independent a1) have "finite B'" unfolding B'_def using \finite B\ by simp obtain D' where "D' > 0" and D': "norm (real_vector.representation B' \ b) \ norm \ * D'" for \ b apply atomize_elim using independent_B' \finite B'\ by (simp add: finite_span_representation_bounded) define D where "D = 2*D'" from \D' > 0\ have \D > 0\ unfolding D_def by simp have "norm (crepresentation B \ b) \ norm \ * D" for \ b proof (cases "b\B") case True have d3: "norm \ = 1" by simp have "norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm \ * norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_scaleC by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d3 by simp finally have d2:"norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b))) = norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))". have "norm (crepresentation B \ b) = norm (complex_of_real (real_vector.representation B' \ b) + \ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" by (simp add: B'_def True a1 a2 crepresentation_from_representation) also have "\ \ norm (complex_of_real (real_vector.representation B' \ b)) + norm (\ *\<^sub>C complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using norm_triangle_ineq by blast also have "\ = norm (complex_of_real (real_vector.representation B' \ b)) + norm (complex_of_real (real_vector.representation B' \ (\ *\<^sub>C b)))" using d2 by simp also have "\ = norm (real_vector.representation B' \ b) + norm (real_vector.representation B' \ (\ *\<^sub>C b))" by simp also have "\ \ norm \ * D' + norm \ * D'" by (rule add_mono; rule D') also have "\ \ norm \ * D" unfolding D_def by linarith finally show ?thesis by auto next case False hence "crepresentation B \ b = 0" using complex_vector.representation_ne_zero by blast thus ?thesis by (smt \0 < D\ norm_ge_zero norm_zero split_mult_pos_le) qed with \D > 0\ show ?thesis by auto qed lemma bounded_clinear_finite_dim[simp]: fixes f :: \'a::{cfinite_dim,complex_normed_vector} \ 'b::complex_normed_vector\ assumes \clinear f\ shows \bounded_clinear f\ proof - include notation_norm obtain basis :: \'a set\ where b1: "complex_vector.span basis = UNIV" and b2: "cindependent basis" and b3:"finite basis" using finite_basis by auto have "\C>0. \\ b. cmod (crepresentation basis \ b) \ \\\ * C" using finite_cspan_crepresentation_bounded[where B = basis] b2 b3 by blast then obtain C where s1: "cmod (crepresentation basis \ b) \ \\\ * C" and s2: "C > 0" for \ b by blast define M where "M = C * (\a\basis. \f a\)" have "\f x\ \ \x\ * M" for x proof- define r where "r b = crepresentation basis x b" for b have x_span: "x \ complex_vector.span basis" by (simp add: b1) have f0: "v \ basis" if "r v \ 0" for v using complex_vector.representation_ne_zero r_def that by auto have w:"{a|a. r a \ 0} \ basis" using f0 by blast hence f1: "finite {a|a. r a \ 0}" using b3 rev_finite_subset by auto have f2: "(\a| r a \ 0. r a *\<^sub>C a) = x" unfolding r_def using b2 complex_vector.sum_nonzero_representation_eq x_span Collect_cong by fastforce have g1: "(\a\basis. crepresentation basis x a *\<^sub>C a) = x" by (simp add: b2 b3 complex_vector.sum_representation_eq x_span) have f3: "(\a\basis. r a *\<^sub>C a) = x" unfolding r_def by (simp add: g1) hence "f x = f (\a\basis. r a *\<^sub>C a)" by simp also have "\ = (\a\basis. r a *\<^sub>C f a)" by (smt (verit, ccfv_SIG) assms complex_vector.linear_scale complex_vector.linear_sum sum.cong) finally have "f x = (\a\basis. r a *\<^sub>C f a)". hence "\f x\ = \(\a\basis. r a *\<^sub>C f a)\" by simp also have "\ \ (\a\basis. \r a *\<^sub>C f a\)" by (simp add: sum_norm_le) also have "\ \ (\a\basis. \r a\ * \f a\)" by simp also have "\ \ (\a\basis. \x\ * C * \f a\)" using sum_mono s1 unfolding r_def by (simp add: sum_mono mult_right_mono) also have "\ \ \x\ * C * (\a\basis. \f a\)" using sum_distrib_left by (smt sum.cong) also have "\ = \x\ * M" unfolding M_def by linarith finally show ?thesis . qed thus ?thesis using assms bounded_clinear_def bounded_clinear_axioms_def by blast qed lemma summable_on_scaleR_left_converse: \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here.\ fixes f :: \'b \ real\ and c :: \'a :: real_normed_vector\ assumes \c \ 0\ assumes \(\x. f x *\<^sub>R c) summable_on A\ shows \f summable_on A\ proof - define fromR toR T where \fromR x = x *\<^sub>R c\ and \toR = inv fromR\ and \T = range fromR\ for x :: real have \additive fromR\ by (simp add: fromR_def additive.intro scaleR_left_distrib) have \inj fromR\ by (simp add: fromR_def \c \ 0\ inj_def) have toR_fromR: \toR (fromR x) = x\ for x by (simp add: \inj fromR\ toR_def) have fromR_toR: \fromR (toR x) = x\ if \x \ T\ for x by (metis T_def f_inv_into_f that toR_def) have 1: \sum (toR \ (fromR \ f)) F = toR (sum (fromR \ f) F)\ if \finite F\ for F by (simp add: o_def additive.sum[OF \additive fromR\, symmetric] toR_fromR) have 2: \sum (fromR \ f) F \ T\ if \finite F\ for F by (simp add: o_def additive.sum[OF \additive fromR\, symmetric] T_def) have 3: \(toR \ toR x) (at x within T)\ for x proof (cases \x \ T\) case True have \dist (toR y) (toR x) < e\ if \y\T\ \e>0\ \dist y x < e * norm c\ for e y proof - obtain x' y' where x: \x = fromR x'\ and y: \y = fromR y'\ using T_def True \y \ T\ by blast have \dist (toR y) (toR x) = dist (fromR (toR y)) (fromR (toR x)) / norm c\ by (auto simp: dist_real_def fromR_def \c \ 0\) also have \\ = dist y x / norm c\ using \x\T\ \y\T\ by (simp add: fromR_toR) also have \\ < e\ using \dist y x < e * norm c\ by (simp add: divide_less_eq that(2)) finally show ?thesis by (simp add: x y toR_fromR) qed then show ?thesis apply (auto simp: tendsto_iff at_within_def eventually_inf_principal eventually_nhds_metric) by (metis assms(1) div_0 divide_less_eq zero_less_norm_iff) next case False have \T = span {c}\ by (simp add: T_def fromR_def span_singleton) then have \closed T\ by simp with False have \x \ closure T\ by simp then have \(at x within T) = bot\ by (rule not_in_closure_trivial_limitI) then show ?thesis by simp qed have 4: \(fromR \ f) summable_on A\ by (simp add: assms(2) fromR_def summable_on_cong) have \(toR o (fromR o f)) summable_on A\ using 1 2 3 4 by (rule summable_on_comm_additive_general[where T=T]) with toR_fromR show ?thesis by (auto simp: o_def) qed lemma infsum_scaleR_left: \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here. It is a strengthening of @{thm [source] infsum_scaleR_left}.\ fixes c :: \'a :: real_normed_vector\ shows "infsum (\x. f x *\<^sub>R c) A = infsum f A *\<^sub>R c" proof (cases \f summable_on A\) case True then show ?thesis apply (subst asm_rl[of \(\x. f x *\<^sub>R c) = (\y. y *\<^sub>R c) o f\], simp add: o_def) apply (rule infsum_comm_additive) using True by (auto simp add: scaleR_left.additive_axioms) next case False then have \\ (\x. f x *\<^sub>R c) summable_on A\ if \c \ 0\ using summable_on_scaleR_left_converse[where A=A and f=f and c=c] using that by auto with False show ?thesis apply (cases \c = 0\) by (auto simp add: infsum_not_exists) qed lemma infsum_of_real: shows \(\\<^sub>\x\A. of_real (f x) :: 'b::{real_normed_vector, real_algebra_1}) = of_real (\\<^sub>\x\A. f x)\ \ \This result has nothing to do with the bounded operator library but it uses @{thm [source] finite_span_closed} so it is proven here.\ unfolding of_real_def by (rule infsum_scaleR_left) subsection \Closed subspaces\ lemma csubspace_INF[simp]: "(\x. x \ A \ csubspace x) \ csubspace (\A)" by (simp add: complex_vector.subspace_Inter) locale closed_csubspace = fixes A::"('a::{complex_vector,topological_space}) set" assumes subspace: "csubspace A" assumes closed: "closed A" declare closed_csubspace.subspace[simp] lemma closure_is_csubspace[simp]: fixes A::"('a::complex_normed_vector) set" assumes \csubspace A\ shows \csubspace (closure A)\ proof- have "x \ closure A \ y \ closure A \ x+y \ closure A" for x y proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast assume \y\(closure A)\ then obtain yy where \\ n::nat. yy n \ A\ and \yy \ y\ using closure_sequential by blast have \\ n::nat. (xx n) + (yy n) \ A\ using \\n. xx n \ A\ \\n. yy n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) hence \(\ n. (xx n) + (yy n)) \ x + y\ using \xx \ x\ \yy \ y\ by (simp add: tendsto_add) thus ?thesis using \\ n::nat. (xx n) + (yy n) \ A\ by (meson closure_sequential) qed moreover have "x\(closure A) \ c *\<^sub>C x \ (closure A)" for x c proof- assume \x\(closure A)\ then obtain xx where \\ n::nat. xx n \ A\ and \xx \ x\ using closure_sequential by blast have \\ n::nat. c *\<^sub>C (xx n) \ A\ using \\n. xx n \ A\ assms complex_vector.subspace_def by (simp add: complex_vector.subspace_def) have \isCont (\ t. c *\<^sub>C t) x\ using bounded_clinear.bounded_linear bounded_clinear_scaleC_right linear_continuous_at by auto hence \(\ n. c *\<^sub>C (xx n)) \ c *\<^sub>C x\ using \xx \ x\ by (simp add: isCont_tendsto_compose) thus ?thesis using \\ n::nat. c *\<^sub>C (xx n) \ A\ by (meson closure_sequential) qed moreover have "0 \ (closure A)" using assms closure_subset complex_vector.subspace_def by (metis in_mono) ultimately show ?thesis by (simp add: complex_vector.subspaceI) qed lemma csubspace_set_plus: assumes \csubspace A\ and \csubspace B\ shows \csubspace (A + B)\ proof - define C where \C = {\+\| \ \. \\A \ \\B}\ have "x\C \ y\C \ x+y\C" for x y using C_def assms(1) assms(2) complex_vector.subspace_add complex_vector.subspace_sums by blast moreover have "c *\<^sub>C x \ C" if \x\C\ for x c proof - have "csubspace C" by (simp add: C_def assms(1) assms(2) complex_vector.subspace_sums) then show ?thesis using that by (simp add: complex_vector.subspace_def) qed moreover have "0 \ C" using \C = {\ + \ |\ \. \ \ A \ \ \ B}\ add.inverse_neutral add_uminus_conv_diff assms(1) assms(2) diff_0 mem_Collect_eq add.right_inverse by (metis (mono_tags, lifting) complex_vector.subspace_0) ultimately show ?thesis unfolding C_def complex_vector.subspace_def by (smt mem_Collect_eq set_plus_elim set_plus_intro) qed lemma closed_csubspace_0[simp]: "closed_csubspace ({0} :: ('a::{complex_vector,t1_space}) set)" proof- have \csubspace {0}\ using add.right_neutral complex_vector.subspace_def scaleC_right.zero by blast moreover have "closed ({0} :: 'a set)" by simp ultimately show ?thesis by (simp add: closed_csubspace_def) qed lemma closed_csubspace_UNIV[simp]: "closed_csubspace (UNIV::('a::{complex_vector,topological_space}) set)" proof- have \csubspace UNIV\ by simp moreover have \closed UNIV\ by simp ultimately show ?thesis unfolding closed_csubspace_def by auto qed lemma closed_csubspace_inter[simp]: assumes "closed_csubspace A" and "closed_csubspace B" shows "closed_csubspace (A\B)" proof- obtain C where \C = A \ B\ by blast have \csubspace C\ proof- have "x\C \ y\C \ x+y\C" for x y by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "x\C \ c *\<^sub>C x \ C" for x c by (metis IntD1 IntD2 IntI \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def) moreover have "0 \ C" using \C = A \ B\ assms(1) assms(2) complex_vector.subspace_def closed_csubspace_def by fastforce ultimately show ?thesis by (simp add: complex_vector.subspace_def) qed moreover have \closed C\ using \C = A \ B\ by (simp add: assms(1) assms(2) closed_Int closed_csubspace.closed) ultimately show ?thesis using \C = A \ B\ by (simp add: closed_csubspace_def) qed lemma closed_csubspace_INF[simp]: assumes a1: "\A\\. closed_csubspace A" shows "closed_csubspace (\\)" proof- have \csubspace (\\)\ by (simp add: assms closed_csubspace.subspace complex_vector.subspace_Inter) moreover have \closed (\\)\ by (simp add: assms closed_Inter closed_csubspace.closed) ultimately show ?thesis by (simp add: closed_csubspace.intro) qed typedef (overloaded) ('a::"{complex_vector,topological_space}") ccsubspace = \{S::'a set. closed_csubspace S}\ morphisms space_as_set Abs_clinear_space using Complex_Vector_Spaces.closed_csubspace_UNIV by blast setup_lifting type_definition_ccsubspace lemma csubspace_space_as_set[simp]: \csubspace (space_as_set S)\ by (metis closed_csubspace_def mem_Collect_eq space_as_set) lemma closed_space_as_set[simp]: \closed (space_as_set S)\ apply transfer by (simp add: closed_csubspace.closed) +lemma zero_space_as_set[simp]: \0 \ space_as_set A\ + by (simp add: complex_vector.subspace_0) + instantiation ccsubspace :: (complex_normed_vector) scaleC begin lift_definition scaleC_ccsubspace :: "complex \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>C) c ` S" proof show "csubspace ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>C) c ` S)" if "closed_csubspace S" for c :: complex and S :: "'a set" using that by (simp add: closed_scaleC closed_csubspace.closed) qed lift_definition scaleR_ccsubspace :: "real \ 'a ccsubspace \ 'a ccsubspace" is "\c S. (*\<^sub>R) c ` S" proof show "csubspace ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that using bounded_clinear_def bounded_clinear_scaleC_right scaleR_scaleC by (simp add: scaleR_scaleC closed_csubspace.subspace complex_vector.linear_subspace_image) show "closed ((*\<^sub>R) r ` S)" if "closed_csubspace S" for r :: real and S :: "'a set" using that by (simp add: closed_scaling closed_csubspace.closed) qed instance proof show "((*\<^sub>R) r::'a ccsubspace \ _) = (*\<^sub>C) (complex_of_real r)" for r :: real by (simp add: scaleR_scaleC scaleC_ccsubspace_def scaleR_ccsubspace_def) qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") bot begin lift_definition bot_ccsubspace :: \'a ccsubspace\ is \{0}\ by simp instance.. end lemma zero_cblinfun_image[simp]: "0 *\<^sub>C S = bot" for S :: "_ ccsubspace" proof transfer have "(0::'b) \ (\x. 0) ` S" if "closed_csubspace S" for S::"'b set" using that unfolding closed_csubspace_def by (simp add: complex_vector.linear_subspace_image complex_vector.module_hom_zero complex_vector.subspace_0) thus "(*\<^sub>C) 0 ` S = {0::'b}" if "closed_csubspace (S::'b set)" for S :: "'b set" using that by (auto intro !: exI [of _ 0]) qed lemma csubspace_scaleC_invariant: fixes a S assumes \a \ 0\ and \csubspace S\ shows \(*\<^sub>C) a ` S = S\ proof- have \x \ (*\<^sub>C) a ` S \ x \ S\ for x using assms(2) complex_vector.subspace_scale by blast moreover have \x \ S \ x \ (*\<^sub>C) a ` S\ for x proof - assume "x \ S" hence "\c aa. (c / a) *\<^sub>C aa \ S \ c *\<^sub>C aa = x" using assms(2) complex_vector.subspace_def scaleC_one by metis hence "\aa. aa \ S \ a *\<^sub>C aa = x" using assms(1) by auto thus ?thesis by (meson image_iff) qed ultimately show ?thesis by blast qed lemma ccsubspace_scaleC_invariant[simp]: "a \ 0 \ a *\<^sub>C S = S" for S :: "_ ccsubspace" apply transfer by (simp add: closed_csubspace.subspace csubspace_scaleC_invariant) instantiation ccsubspace :: ("{complex_vector,topological_space}") "top" begin lift_definition top_ccsubspace :: \'a ccsubspace\ is \UNIV\ by simp instance .. end lemma space_as_set_bot[simp]: \space_as_set bot = {0}\ by (rule bot_ccsubspace.rep_eq) lemma ccsubspace_top_not_bot[simp]: "(top::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ bot" (* The type class t1_space is needed because the definition of bot in ccsubspace needs it *) by (metis UNIV_not_singleton bot_ccsubspace.rep_eq top_ccsubspace.rep_eq) lemma ccsubspace_bot_not_top[simp]: "(bot::'a::{complex_vector,t1_space,not_singleton} ccsubspace) \ top" using ccsubspace_top_not_bot by metis instantiation ccsubspace :: ("{complex_vector,topological_space}") "Inf" begin lift_definition Inf_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\ S. \ S\ proof fix S :: "'a set set" assume closed: "closed_csubspace x" if \x \ S\ for x show "csubspace (\ S::'a set)" by (simp add: closed closed_csubspace.subspace) show "closed (\ S::'a set)" by (simp add: closed closed_csubspace.closed) qed instance .. end lift_definition ccspan :: "'a::complex_normed_vector set \ 'a ccsubspace" is "\G. closure (cspan G)" proof (rule closed_csubspace.intro) fix S :: "'a set" show "csubspace (closure (cspan S))" by (simp add: closure_is_csubspace) show "closed (closure (cspan S))" by simp qed lemma ccspan_canonical_basis[simp]: "ccspan (set canonical_basis) = top" using ccspan.rep_eq space_as_set_inject top_ccsubspace.rep_eq closure_UNIV is_generator_set by metis lemma ccspan_Inf_def: \ccspan A = Inf {S. A \ space_as_set S}\ for A::\('a::cbanach) set\ proof- have \x \ space_as_set (ccspan A) \ x \ space_as_set (Inf {S. A \ space_as_set S})\ for x::'a proof- assume \x \ space_as_set (ccspan A)\ hence "x \ closure (cspan A)" by (simp add: ccspan.rep_eq) hence \x \ closure (complex_vector.span A)\ unfolding ccspan_def by simp hence \\ y::nat \ 'a. (\ n. y n \ (complex_vector.span A)) \ y \ x\ by (simp add: closure_sequential) then obtain y where \\ n. y n \ (complex_vector.span A)\ and \y \ x\ by blast have \y n \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ for n using \\ n. y n \ (complex_vector.span A)\ by auto have \closed_csubspace S \ closed S\ for S::\'a set\ by (simp add: closed_csubspace.closed) hence \closed ( \ {S. (complex_vector.span A) \ S \ closed_csubspace S})\ by simp hence \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using \y \ x\ using \\n. y n \ \ {S. complex_vector.span A \ S \ closed_csubspace S}\ closed_sequentially by blast moreover have \{S. A \ S \ closed_csubspace S} \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ using Collect_mono_iff by (simp add: Collect_mono_iff closed_csubspace.subspace complex_vector.span_minimal) ultimately have \x \ \ {S. A \ S \ closed_csubspace S}\ by blast moreover have "(x::'a) \ \ {x. A \ x \ closed_csubspace x}" if "(x::'a) \ \ {S. A \ S \ closed_csubspace S}" for x :: 'a and A :: "'a set" using that by simp ultimately show \x \ space_as_set (Inf {S. A \ space_as_set S})\ apply transfer. qed moreover have \x \ space_as_set (Inf {S. A \ space_as_set S}) \ x \ space_as_set (ccspan A)\ for x::'a proof- assume \x \ space_as_set (Inf {S. A \ space_as_set S})\ hence \x \ \ {S. A \ S \ closed_csubspace S}\ apply transfer by blast moreover have \{S. (complex_vector.span A) \ S \ closed_csubspace S} \ {S. A \ S \ closed_csubspace S}\ using Collect_mono_iff complex_vector.span_superset by fastforce ultimately have \x \ \ {S. (complex_vector.span A) \ S \ closed_csubspace S}\ by blast thus \x \ space_as_set (ccspan A)\ by (metis (no_types, lifting) Inter_iff space_as_set closure_subset mem_Collect_eq ccspan.rep_eq) qed ultimately have \space_as_set (ccspan A) = space_as_set (Inf {S. A \ space_as_set S})\ by blast thus ?thesis using space_as_set_inject by auto qed lemma cspan_singleton_scaleC[simp]: "(a::complex)\0 \ cspan { a *\<^sub>C \ } = cspan {\}" for \::"'a::complex_vector" by (smt complex_vector.dependent_single complex_vector.independent_insert complex_vector.scale_eq_0_iff complex_vector.span_base complex_vector.span_redundant complex_vector.span_scale doubleton_eq_iff insert_absorb insert_absorb2 insert_commute singletonI) lemma closure_is_closed_csubspace[simp]: fixes S::\'a::complex_normed_vector set\ assumes \csubspace S\ shows \closed_csubspace (closure S)\ proof- fix x y :: 'a and c :: complex have "x + y \ closure S" if "x \ closure S" and "y \ closure S" proof- have \\ r. (\ n::nat. r n \ S) \ r \ x\ using closure_sequential that(1) by auto then obtain r where \\ n::nat. r n \ S\ and \r \ x\ by blast have \\ s. (\ n::nat. s n \ S) \ s \ y\ using closure_sequential that(2) by auto then obtain s where \\ n::nat. s n \ S\ and \s \ y\ by blast have \\ n::nat. r n + s n \ S\ using \\n. r n \ S\ \\n. s n \ S\ assms complex_vector.subspace_add by blast moreover have \(\ n. r n + s n) \ x + y\ by (simp add: \r \ x\ \s \ y\ tendsto_add) ultimately show ?thesis using assms that(1) that(2) by (simp add: complex_vector.subspace_add) qed moreover have "c *\<^sub>C x \ closure S" if "x \ closure S" proof- have \\ y. (\ n::nat. y n \ S) \ y \ x\ using Elementary_Topology.closure_sequential that by auto then obtain y where \\ n::nat. y n \ S\ and \y \ x\ by blast have \isCont (scaleC c) x\ by simp hence \(\ n. scaleC c (y n)) \ scaleC c x\ using \y \ x\ by (simp add: isCont_tendsto_compose) from \\ n::nat. y n \ S\ have \\ n::nat. scaleC c (y n) \ S\ using assms complex_vector.subspace_scale by auto thus ?thesis using assms that by (simp add: complex_vector.subspace_scale) qed moreover have "0 \ closure S" by (simp add: assms complex_vector.subspace_0) moreover have "closed (closure S)" by auto ultimately show ?thesis by (simp add: assms closed_csubspace_def) qed lemma ccspan_singleton_scaleC[simp]: "(a::complex)\0 \ ccspan {a *\<^sub>C \} = ccspan {\}" apply transfer by simp lemma clinear_continuous_at: assumes \bounded_clinear f\ shows \isCont f x\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_at) lemma clinear_continuous_within: assumes \bounded_clinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_clinear.bounded_linear linear_continuous_within) lemma antilinear_continuous_at: assumes \bounded_antilinear f\ shows \isCont f x\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_at) lemma antilinear_continuous_within: assumes \bounded_antilinear f\ shows \continuous (at x within s) f\ by (simp add: assms bounded_antilinear.bounded_linear linear_continuous_within) lemma bounded_clinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_clinear A\ and \bounded_clinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - have eq': \A t = B t\ if \t \ cspan G\ for t using _ _ that eq apply (rule complex_vector.linear_eq_on) by (auto simp: assms bounded_clinear.clinear) have \A t - B t = 0\ using _ _ t apply (rule continuous_constant_on_closure) by (auto simp add: eq' assms(1) assms(2) clinear_continuous_at continuous_at_imp_continuous_on) then show ?thesis by auto qed instantiation ccsubspace :: ("{complex_vector,topological_space}") "order" begin lift_definition less_eq_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_eq_ccsubspace_def[code del] lift_definition less_ccsubspace :: \'a ccsubspace \ 'a ccsubspace \ bool\ is \(\)\. declare less_ccsubspace_def[code del] instance proof fix x y z :: "'a ccsubspace" show "(x < y) = (x \ y \ \ y \ x)" by (simp add: less_eq_ccsubspace.rep_eq less_le_not_le less_ccsubspace.rep_eq) show "x \ x" by (simp add: less_eq_ccsubspace.rep_eq) show "x \ z" if "x \ y" and "y \ z" using that less_eq_ccsubspace.rep_eq by auto show "x = y" if "x \ y" and "y \ x" using that by (simp add: space_as_set_inject less_eq_ccsubspace.rep_eq) qed end lemma ccspan_leqI: assumes \M \ space_as_set S\ shows \ccspan M \ S\ using assms apply transfer by (simp add: closed_csubspace.closed closure_minimal complex_vector.span_minimal) lemma ccspan_mono: assumes \A \ B\ shows \ccspan A \ ccspan B\ apply (transfer fixing: A B) by (simp add: assms closure_mono complex_vector.span_mono) lemma bounded_sesquilinear_add: \bounded_sesquilinear (\ x y. A x y + B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "A (a + a') b + B (a + a') b = (A a b + B a b) + (A a' b + B a' b)" by (simp add: bounded_sesquilinear.add_left that(1) that(2)) show \A a (b + b') + B a (b + b') = (A a b + B a b) + (A a b' + B a b')\ by (simp add: bounded_sesquilinear.add_right that(1) that(2)) show \A (r *\<^sub>C a) b + B (r *\<^sub>C a) b = cnj r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_left scaleC_add_right that(1) that(2)) show \A a (r *\<^sub>C b) + B a (r *\<^sub>C b) = r *\<^sub>C (A a b + B a b)\ by (simp add: bounded_sesquilinear.scaleC_right scaleC_add_right that(1) that(2)) show \\K. \a b. norm (A a b + B a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \\ KB. \ a b. norm (B a b) \ norm a * norm b * KB\ by (simp add: bounded_sesquilinear.bounded that(2)) then obtain KB where \\ a b. norm (B a b) \ norm a * norm b * KB\ by blast have \norm (A a b + B a b) \ norm a * norm b * (KA + KB)\ for a b proof- have \norm (A a b + B a b) \ norm (A a b) + norm (B a b)\ using norm_triangle_ineq by blast also have \\ \ norm a * norm b * KA + norm a * norm b * KB\ using \\ a b. norm (A a b) \ norm a * norm b * KA\ \\ a b. norm (B a b) \ norm a * norm b * KB\ using add_mono by blast also have \\= norm a * norm b * (KA + KB)\ by (simp add: mult.commute ring_class.ring_distribs(2)) finally show ?thesis by blast qed thus ?thesis by blast qed qed lemma bounded_sesquilinear_uminus: \bounded_sesquilinear (\ x y. - A x y)\ if \bounded_sesquilinear A\ proof fix a a' :: 'a and b b' :: 'b and r :: complex show "- A (a + a') b = (- A a b) + (- A a' b)" by (simp add: bounded_sesquilinear.add_left that) show \- A a (b + b') = (- A a b) + (- A a b')\ by (simp add: bounded_sesquilinear.add_right that) show \- A (r *\<^sub>C a) b = cnj r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_left that) show \- A a (r *\<^sub>C b) = r *\<^sub>C (- A a b)\ by (simp add: bounded_sesquilinear.scaleC_right that) show \\K. \a b. norm (- A a b) \ norm a * norm b * K\ proof- have \\ KA. \ a b. norm (A a b) \ norm a * norm b * KA\ by (simp add: bounded_sesquilinear.bounded that(1)) then obtain KA where \\ a b. norm (A a b) \ norm a * norm b * KA\ by blast have \norm (- A a b) \ norm a * norm b * KA\ for a b by (simp add: \\a b. norm (A a b) \ norm a * norm b * KA\) thus ?thesis by blast qed qed lemma bounded_sesquilinear_diff: \bounded_sesquilinear (\ x y. A x y - B x y)\ if \bounded_sesquilinear A\ and \bounded_sesquilinear B\ proof - have \bounded_sesquilinear (\ x y. - B x y)\ using that(2) by (rule bounded_sesquilinear_uminus) then have \bounded_sesquilinear (\ x y. A x y + (- B x y))\ using that(1) by (rule bounded_sesquilinear_add[rotated]) then show ?thesis by auto qed lemma ccsubspace_leI: assumes t1: "space_as_set A \ space_as_set B" shows "A \ B" using t1 apply transfer by - lemma ccspan_of_empty[simp]: "ccspan {} = bot" proof transfer show "closure (cspan {}) = {0::'a}" by simp qed instantiation ccsubspace :: ("{complex_vector,topological_space}") inf begin lift_definition inf_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" is "(\)" by simp instance .. end lemma space_as_set_inf[simp]: "space_as_set (A \ B) = space_as_set A \ space_as_set B" by (rule inf_ccsubspace.rep_eq) instantiation ccsubspace :: ("{complex_vector,topological_space}") order_top begin instance proof show "a \ \" for a :: "'a ccsubspace" apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") order_bot begin instance proof show "(\::'a ccsubspace) \ a" for a :: "'a ccsubspace" apply transfer apply auto using closed_csubspace.subspace complex_vector.subspace_0 by blast qed end instantiation ccsubspace :: ("{complex_vector,topological_space}") semilattice_inf begin instance proof fix x y z :: \'a ccsubspace\ show "x \ y \ x" apply transfer by simp show "x \ y \ y" apply transfer by simp show "x \ y \ z" if "x \ y" and "x \ z" using that apply transfer by simp qed end instantiation ccsubspace :: ("{complex_vector,t1_space}") zero begin definition zero_ccsubspace :: "'a ccsubspace" where [simp]: "zero_ccsubspace = bot" lemma zero_ccsubspace_transfer[transfer_rule]: \pcr_ccsubspace (=) {0} 0\ unfolding zero_ccsubspace_def by transfer_prover instance .. end lemma ccspan_0[simp]: \ccspan {0} = 0\ apply transfer by simp definition \rel_ccsubspace R x y = rel_set R (space_as_set x) (space_as_set y)\ lemma left_unique_rel_ccsubspace[transfer_rule]: \left_unique (rel_ccsubspace R)\ if \left_unique R\ proof (rule left_uniqueI) fix S T :: \'a ccsubspace\ and U assume assms: \rel_ccsubspace R S U\ \rel_ccsubspace R T U\ have \space_as_set S = space_as_set T\ apply (rule left_uniqueD) using that apply (rule left_unique_rel_set) using assms unfolding rel_ccsubspace_def by auto then show \S = T\ by (simp add: space_as_set_inject) qed lemma right_unique_rel_ccsubspace[transfer_rule]: \right_unique (rel_ccsubspace R)\ if \right_unique R\ by (metis rel_ccsubspace_def right_unique_def right_unique_rel_set space_as_set_inject that) lemma bi_unique_rel_ccsubspace[transfer_rule]: \bi_unique (rel_ccsubspace R)\ if \bi_unique R\ by (metis (no_types, lifting) bi_unique_def bi_unique_rel_set rel_ccsubspace_def space_as_set_inject that) lemma converse_rel_ccsubspace: \conversep (rel_ccsubspace R) = rel_ccsubspace (conversep R)\ by (auto simp: rel_ccsubspace_def[abs_def]) lemma space_as_set_top[simp]: \space_as_set top = UNIV\ by (rule top_ccsubspace.rep_eq) lemma ccsubspace_eqI: assumes \\x. x \ space_as_set S \ x \ space_as_set T\ shows \S = T\ by (metis Abs_clinear_space_cases Abs_clinear_space_inverse antisym assms subsetI) +lemma ccspan_remove_0: \ccspan (A - {0}) = ccspan A\ + apply transfer + by auto + +lemma sgn_in_spaceD: \\ \ space_as_set A\ if \sgn \ \ space_as_set A\ and \\ \ 0\ + for \ :: \_ :: complex_normed_vector\ + using that + apply (transfer fixing: \) + by (metis closed_csubspace.subspace complex_vector.subspace_scale divideC_field_simps(1) scaleR_eq_0_iff scaleR_scaleC sgn_div_norm sgn_zero_iff) + +lemma sgn_in_spaceI: \sgn \ \ space_as_set A\ if \\ \ space_as_set A\ + for \ :: \_ :: complex_normed_vector\ + using that by (auto simp: sgn_div_norm scaleR_scaleC complex_vector.subspace_scale) + +lemma ccsubspace_leI_unit: + fixes A B :: \_ :: complex_normed_vector ccsubspace\ + assumes "\\. norm \ = 1 \ \ \ space_as_set A \ \ \ space_as_set B" + shows "A \ B" +proof (rule ccsubspace_leI, rule subsetI) + fix \ assume \A: \\ \ space_as_set A\ + show \\ \ space_as_set B\ + apply (cases \\ = 0\) + apply simp + using assms[of \sgn \\] \A sgn_in_spaceD sgn_in_spaceI + by (auto simp: norm_sgn) +qed + subsection \Closed sums\ definition closed_sum:: \'a::{semigroup_add,topological_space} set \ 'a set \ 'a set\ where \closed_sum A B = closure (A + B)\ notation closed_sum (infixl "+\<^sub>M" 65) lemma closed_sum_comm: \A +\<^sub>M B = B +\<^sub>M A\ for A B :: "_::ab_semigroup_add" by (simp add: add.commute closed_sum_def) lemma closed_sum_left_subset: \0 \ B \ A \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.right_neutral closed_sum_def closure_subset in_mono set_plus_intro subsetI) lemma closed_sum_right_subset: \0 \ A \ B \ A +\<^sub>M B\ for A B :: "_::monoid_add" by (metis add.left_neutral closed_sum_def closure_subset set_plus_intro subset_iff) lemma finite_cspan_closed_csubspace: assumes "finite (S::'a::complex_normed_vector set)" shows "closed_csubspace (cspan S)" by (simp add: assms closed_csubspace.intro) lemma closed_sum_is_sup: fixes A B C:: \('a::{complex_vector,topological_space}) set\ assumes \closed_csubspace C\ assumes \A \ C\ and \B \ C\ shows \(A +\<^sub>M B) \ C\ proof - have \A + B \ C\ using assms unfolding set_plus_def using closed_csubspace.subspace complex_vector.subspace_add by blast then show \(A +\<^sub>M B) \ C\ unfolding closed_sum_def using \closed_csubspace C\ by (simp add: closed_csubspace.closed closure_minimal) qed lemma closed_subspace_closed_sum: fixes A B::"('a::complex_normed_vector) set" assumes a1: \csubspace A\ and a2: \csubspace B\ shows \closed_csubspace (A +\<^sub>M B)\ using a1 a2 closed_sum_def by (metis closure_is_closed_csubspace csubspace_set_plus) lemma closed_sum_assoc: fixes A B C::"'a::real_normed_vector set" shows \A +\<^sub>M (B +\<^sub>M C) = (A +\<^sub>M B) +\<^sub>M C\ proof - have \A + closure B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \A +\<^sub>M (B +\<^sub>M C) = closure (A + (B + C))\ unfolding closed_sum_def by (meson antisym_conv closed_closure closure_minimal closure_mono closure_subset equalityD1 set_plus_mono2) moreover have \closure A + B \ closure (A + B)\ for A B :: "'a set" by (meson closure_subset closure_sum dual_order.trans order_refl set_plus_mono2) then have \(A +\<^sub>M B) +\<^sub>M C = closure ((A + B) + C)\ unfolding closed_sum_def by (meson closed_closure closure_minimal closure_mono closure_subset eq_iff set_plus_mono2) ultimately show ?thesis by (simp add: ab_semigroup_add_class.add_ac(1)) qed lemma closed_sum_zero_left[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \{0} +\<^sub>M A = closure A\ unfolding closed_sum_def by (metis add.left_neutral set_zero) lemma closed_sum_zero_right[simp]: fixes A :: \('a::{monoid_add, topological_space}) set\ shows \A +\<^sub>M {0} = closure A\ unfolding closed_sum_def by (metis add.right_neutral set_zero) lemma closed_sum_closure_right[simp]: fixes A B :: \'a::real_normed_vector set\ shows \A +\<^sub>M closure B = A +\<^sub>M B\ by (metis closed_sum_assoc closed_sum_def closed_sum_zero_right closure_closure) lemma closed_sum_closure_left[simp]: fixes A B :: \'a::real_normed_vector set\ shows \closure A +\<^sub>M B = A +\<^sub>M B\ by (simp add: closed_sum_comm) lemma closed_sum_mono_left: assumes \A \ B\ shows \A +\<^sub>M C \ B +\<^sub>M C\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) lemma closed_sum_mono_right: assumes \A \ B\ shows \C +\<^sub>M A \ C +\<^sub>M B\ by (simp add: assms closed_sum_def closure_mono set_plus_mono2) instantiation ccsubspace :: (complex_normed_vector) sup begin lift_definition sup_ccsubspace :: "'a ccsubspace \ 'a ccsubspace \ 'a ccsubspace" \ \Note that \<^term>\A+B\ would not be a closed subspace, we need the closure. See, e.g., \<^url>\https://math.stackexchange.com/a/1786792/403528\.\ is "\A B::'a set. A +\<^sub>M B" by (simp add: closed_subspace_closed_sum) instance .. end lemma closed_sum_cspan[simp]: shows \cspan X +\<^sub>M cspan Y = closure (cspan (X \ Y))\ by (smt (verit, best) Collect_cong closed_sum_def complex_vector.span_Un set_plus_def) lemma closure_image_closed_sum: assumes \bounded_linear U\ shows \closure (U ` (A +\<^sub>M B)) = closure (U ` A) +\<^sub>M closure (U ` B)\ proof - have \closure (U ` (A +\<^sub>M B)) = closure (U ` closure (closure A + closure B))\ unfolding closed_sum_def by (smt (verit, best) closed_closure closure_minimal closure_mono closure_subset closure_sum set_plus_mono2 subset_antisym) also have \\ = closure (U ` (closure A + closure B))\ using assms closure_bounded_linear_image_subset_eq by blast also have \\ = closure (U ` closure A + U ` closure B)\ apply (subst image_set_plus) by (simp_all add: assms bounded_linear.linear) also have \\ = closure (closure (U ` A) + closure (U ` B))\ by (smt (verit, ccfv_SIG) assms closed_closure closure_bounded_linear_image_subset closure_bounded_linear_image_subset_eq closure_minimal closure_mono closure_sum dual_order.eq_iff set_plus_mono2) also have \\ = closure (U ` A) +\<^sub>M closure (U ` B)\ using closed_sum_def by blast finally show ?thesis by - qed lemma ccspan_union: "ccspan A \ ccspan B = ccspan (A \ B)" apply transfer by simp instantiation ccsubspace :: (complex_normed_vector) "Sup" begin lift_definition Sup_ccsubspace::\'a ccsubspace set \ 'a ccsubspace\ is \\S. closure (complex_vector.span (Union S))\ proof show "csubspace (closure (complex_vector.span (\ S::'a set)))" if "\x::'a set. x \ S \ closed_csubspace x" for S :: "'a set set" using that by (simp add: closure_is_closed_csubspace) show "closed (closure (complex_vector.span (\ S::'a set)))" if "\x. (x::'a set) \ S \ closed_csubspace x" for S :: "'a set set" using that by simp qed instance.. end instance ccsubspace :: ("{complex_normed_vector}") semilattice_sup proof fix x y z :: \'a ccsubspace\ show \x \ sup x y\ apply transfer by (simp add: closed_csubspace_def closed_sum_left_subset complex_vector.subspace_0) show "y \ sup x y" apply transfer by (simp add: closed_csubspace_def closed_sum_right_subset complex_vector.subspace_0) show "sup x y \ z" if "x \ z" and "y \ z" using that apply transfer apply (rule closed_sum_is_sup) by auto qed instance ccsubspace :: ("{complex_normed_vector}") complete_lattice proof show "Inf A \ x" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by auto have b1: "z \ \ A" if "Ball A closed_csubspace" and "closed_csubspace z" and "(\x. closed_csubspace x \ x \ A \ z \ x)" for z::"'a set" and A using that by auto show "z \ Inf A" if "\x::'a ccsubspace. x \ A \ z \ x" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer using b1 by blast show "x \ Sup A" if "x \ A" for x :: "'a ccsubspace" and A :: "'a ccsubspace set" using that apply transfer by (meson Union_upper closure_subset complex_vector.span_superset dual_order.trans) show "Sup A \ z" if "\x::'a ccsubspace. x \ A \ x \ z" for A :: "'a ccsubspace set" and z :: "'a ccsubspace" using that apply transfer proof - fix A :: "'a set set" and z :: "'a set" assume A_closed: "Ball A closed_csubspace" assume "closed_csubspace z" assume in_z: "\x. closed_csubspace x \ x \ A \ x \ z" from A_closed in_z have \V \ z\ if \V \ A\ for V by (simp add: that) then have \\ A \ z\ by (simp add: Sup_le_iff) with \closed_csubspace z\ show "closure (cspan (\ A)) \ z" by (simp add: closed_csubspace_def closure_minimal complex_vector.span_def subset_hull) qed show "Inf {} = (top::'a ccsubspace)" using \\z A. (\x. x \ A \ z \ x) \ z \ Inf A\ top.extremum_uniqueI by auto show "Sup {} = (bot::'a ccsubspace)" using \\z A. (\x. x \ A \ x \ z) \ Sup A \ z\ bot.extremum_uniqueI by auto qed instantiation ccsubspace :: (complex_normed_vector) comm_monoid_add begin definition plus_ccsubspace :: "'a ccsubspace \ _ \ _" where [simp]: "plus_ccsubspace = sup" instance proof fix a b c :: \'a ccsubspace\ show "a + b + c = a + (b + c)" using sup.assoc by auto show "a + b = b + a" by (simp add: sup.commute) show "0 + a = a" by (simp add: zero_ccsubspace_def) qed end lemma ccsubspace_plus_sup: "y \ x \ z \ x \ y + z \ x" for x y z :: "'a::complex_normed_vector ccsubspace" unfolding plus_ccsubspace_def by auto lemma ccsubspace_Sup_empty: "Sup {} = (0::_ ccsubspace)" unfolding zero_ccsubspace_def by auto lemma ccsubspace_add_right_incr[simp]: "a \ a + c" for a::"_ ccsubspace" by (simp add: add_increasing2) lemma ccsubspace_add_left_incr[simp]: "a \ c + a" for a::"_ ccsubspace" by (simp add: add_increasing) +lemma sum_bot_ccsubspace[simp]: \(\x\X. \) = (\ :: _ ccsubspace)\ + by (simp flip: zero_ccsubspace_def) + subsection \Conjugate space\ typedef 'a conjugate_space = "UNIV :: 'a set" morphisms from_conjugate_space to_conjugate_space .. setup_lifting type_definition_conjugate_space instantiation conjugate_space :: (complex_vector) complex_vector begin lift_definition scaleC_conjugate_space :: \complex \ 'a conjugate_space \ 'a conjugate_space\ is \\c x. cnj c *\<^sub>C x\. lift_definition scaleR_conjugate_space :: \real \ 'a conjugate_space \ 'a conjugate_space\ is \\r x. r *\<^sub>R x\. lift_definition plus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(+)". lift_definition uminus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is \\x. -x\. lift_definition zero_conjugate_space :: "'a conjugate_space" is 0. lift_definition minus_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ 'a conjugate_space" is "(-)". instance apply (intro_classes; transfer) by (simp_all add: scaleR_scaleC scaleC_add_right scaleC_left.add) end instantiation conjugate_space :: (complex_normed_vector) complex_normed_vector begin lift_definition sgn_conjugate_space :: "'a conjugate_space \ 'a conjugate_space" is "sgn". lift_definition norm_conjugate_space :: "'a conjugate_space \ real" is norm. lift_definition dist_conjugate_space :: "'a conjugate_space \ 'a conjugate_space \ real" is dist. lift_definition uniformity_conjugate_space :: "('a conjugate_space \ 'a conjugate_space) filter" is uniformity. lift_definition open_conjugate_space :: "'a conjugate_space set \ bool" is "open". instance apply (intro_classes; transfer) by (simp_all add: dist_norm sgn_div_norm open_uniformity uniformity_dist norm_triangle_ineq) end instantiation conjugate_space :: (cbanach) cbanach begin instance apply intro_classes unfolding Cauchy_def convergent_def LIMSEQ_def apply transfer using Cauchy_convergent unfolding Cauchy_def convergent_def LIMSEQ_def by metis end lemma bounded_antilinear_to_conjugate_space[simp]: \bounded_antilinear to_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma bounded_antilinear_from_conjugate_space[simp]: \bounded_antilinear from_conjugate_space\ by (rule bounded_antilinear_intro[where K=1]; transfer; auto) lemma antilinear_to_conjugate_space[simp]: \antilinear to_conjugate_space\ by (rule antilinearI; transfer, auto) lemma antilinear_from_conjugate_space[simp]: \antilinear from_conjugate_space\ by (rule antilinearI; transfer, auto) lemma cspan_to_conjugate_space[simp]: "cspan (to_conjugate_space ` X) = to_conjugate_space ` cspan X" unfolding complex_vector.span_def complex_vector.subspace_def hull_def apply transfer apply simp by (metis (no_types, opaque_lifting) complex_cnj_cnj) lemma surj_to_conjugate_space[simp]: "surj to_conjugate_space" by (meson surj_def to_conjugate_space_cases) lemmas has_derivative_scaleC[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_cbilinear_scaleC[THEN bounded_cbilinear.bounded_bilinear]] lemma norm_to_conjugate_space[simp]: \norm (to_conjugate_space x) = norm x\ by (fact norm_conjugate_space.abs_eq) lemma norm_from_conjugate_space[simp]: \norm (from_conjugate_space x) = norm x\ by (simp add: norm_conjugate_space.rep_eq) lemma closure_to_conjugate_space: \closure (to_conjugate_space ` X) = to_conjugate_space ` closure X\ proof - have 1: \to_conjugate_space ` closure X \ closure (to_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = to_conjugate_space ` from_conjugate_space ` closure (to_conjugate_space ` X)\ by (simp add: from_conjugate_space_inverse image_image) also have \\ \ to_conjugate_space ` closure (from_conjugate_space ` to_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = to_conjugate_space ` closure X\ by (simp add: to_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma closure_from_conjugate_space: \closure (from_conjugate_space ` X) = from_conjugate_space ` closure X\ proof - have 1: \from_conjugate_space ` closure X \ closure (from_conjugate_space ` X)\ apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) have \\ = from_conjugate_space ` to_conjugate_space ` closure (from_conjugate_space ` X)\ by (simp add: to_conjugate_space_inverse image_image) also have \\ \ from_conjugate_space ` closure (to_conjugate_space ` from_conjugate_space ` X)\ apply (rule image_mono) apply (rule closure_bounded_linear_image_subset) by (simp add: bounded_antilinear.bounded_linear) also have \\ = from_conjugate_space ` closure X\ by (simp add: from_conjugate_space_inverse image_image) finally show ?thesis using 1 by simp qed lemma bounded_antilinear_eq_on: fixes A B :: "'a::complex_normed_vector \ 'b::complex_normed_vector" assumes \bounded_antilinear A\ and \bounded_antilinear B\ and eq: \\x. x \ G \ A x = B x\ and t: \t \ closure (cspan G)\ shows \A t = B t\ proof - let ?A = \\x. A (from_conjugate_space x)\ and ?B = \\x. B (from_conjugate_space x)\ and ?G = \to_conjugate_space ` G\ and ?t = \to_conjugate_space t\ have \bounded_clinear ?A\ and \bounded_clinear ?B\ by (auto intro!: bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear A\] bounded_antilinear_o_bounded_antilinear[OF \bounded_antilinear B\]) moreover from eq have \\x. x \ ?G \ ?A x = ?B x\ by (metis image_iff iso_tuple_UNIV_I to_conjugate_space_inverse) moreover from t have \?t \ closure (cspan ?G)\ by (metis bounded_antilinear.bounded_linear bounded_antilinear_to_conjugate_space closure_bounded_linear_image_subset cspan_to_conjugate_space imageI subsetD) ultimately have \?A ?t = ?B ?t\ by (rule bounded_clinear_eq_on) then show \A t = B t\ by (simp add: to_conjugate_space_inverse) qed instantiation complex :: basis_enum begin definition "canonical_basis = [1::complex]" instance proof show "distinct (canonical_basis::complex list)" by (simp add: canonical_basis_complex_def) show "cindependent (set (canonical_basis::complex list))" unfolding canonical_basis_complex_def by auto show "cspan (set (canonical_basis::complex list)) = UNIV" unfolding canonical_basis_complex_def apply (auto simp add: cspan_raw_def vector_space_over_itself.span_Basis) by (metis complex_scaleC_def complex_vector.span_base complex_vector.span_scale cspan_raw_def insertI1 mult.right_neutral) qed end lemma csubspace_is_convex[simp]: assumes a1: "csubspace M" shows "convex M" proof- have \\x\M. \y\ M. \u. \v. u *\<^sub>C x + v *\<^sub>C y \ M\ using a1 by (simp add: complex_vector.subspace_def) hence \\x\M. \y\M. \u::real. \v::real. u *\<^sub>R x + v *\<^sub>R y \ M\ by (simp add: scaleR_scaleC) hence \\x\M. \y\M. \u\0. \v\0. u + v = 1 \ u *\<^sub>R x + v *\<^sub>R y \M\ by blast thus ?thesis using convex_def by blast qed lemma kernel_is_csubspace[simp]: assumes a1: "clinear f" shows "csubspace (f -` {0})" proof- have w3: \t *\<^sub>C x \ {x. f x = 0}\ if b1: "x \ {x. f x = 0}" for x t by (metis assms complex_vector.linear_subspace_kernel complex_vector.subspace_def that) have \f 0 = 0\ by (simp add: assms complex_vector.linear_0) hence s2: \0 \ {x. f x = 0}\ by blast have w4: "x + y \ {x. f x = 0}" if c1: "x \ {x. f x = 0}" and c2: "y \ {x. f x = 0}" for x y using assms c1 c2 complex_vector.linear_add by fastforce have s4: \c *\<^sub>C t \ {x. f x = 0}\ if "t \ {x. f x = 0}" for t c using that w3 by auto have s5: "u + v \ {x. f x = 0}" if "u \ {x. f x = 0}" and "v \ {x. f x = 0}" for u v using w4 that(1) that(2) by auto have f3: "f -` {b. b = 0 \ b \ {}} = {a. f a = 0}" by blast have "csubspace {a. f a = 0}" by (metis complex_vector.subspace_def s2 s4 s5) thus ?thesis using f3 by auto qed lemma kernel_is_closed_csubspace[simp]: assumes a1: "bounded_clinear f" shows "closed_csubspace (f -` {0})" proof- have \csubspace (f -` {0})\ using assms bounded_clinear.clinear complex_vector.linear_subspace_vimage complex_vector.subspace_single_0 by blast have "L \ {x. f x = 0}" if "r \ L" and "\ n. r n \ {x. f x = 0}" for r and L proof- have d1: \\ n. f (r n) = 0\ using that(2) by auto have \(\ n. f (r n)) \ f L\ using assms clinear_continuous_at continuous_within_tendsto_compose' that(1) by fastforce hence \(\ n. 0) \ f L\ using d1 by simp hence \f L = 0\ using limI by fastforce thus ?thesis by blast qed then have s3: \closed (f -` {0})\ using closed_sequential_limits by force with \csubspace (f -` {0})\ show ?thesis using closed_csubspace.intro by blast qed lemma range_is_clinear[simp]: assumes a1: "clinear f" shows "csubspace (range f)" using assms complex_vector.linear_subspace_image complex_vector.subspace_UNIV by blast lemma ccspan_superset: \A \ space_as_set (ccspan A)\ for A :: \'a::complex_normed_vector set\ apply transfer by (meson closure_subset complex_vector.span_superset subset_trans) subsection \Product is a Complex Vector Space\ (* Follows closely Product_Vector.thy *) instantiation prod :: (complex_vector, complex_vector) complex_vector begin definition scaleC_prod_def: "scaleC r A = (scaleC r (fst A), scaleC r (snd A))" lemma fst_scaleC [simp]: "fst (scaleC r A) = scaleC r (fst A)" unfolding scaleC_prod_def by simp lemma snd_scaleC [simp]: "snd (scaleC r A) = scaleC r (snd A)" unfolding scaleC_prod_def by simp proposition scaleC_Pair [simp]: "scaleC r (a, b) = (scaleC r a, scaleC r b)" unfolding scaleC_prod_def by simp instance proof fix a b :: complex and x y :: "'a \ 'b" show "scaleC a (x + y) = scaleC a x + scaleC a y" by (simp add: scaleC_add_right scaleC_prod_def) show "scaleC (a + b) x = scaleC a x + scaleC b x" by (simp add: Complex_Vector_Spaces.scaleC_prod_def scaleC_left.add) show "scaleC a (scaleC b x) = scaleC (a * b) x" by (simp add: prod_eq_iff) show "scaleC 1 x = x" by (simp add: prod_eq_iff) show \(scaleR :: _ \ _ \ 'a*'b) r = (*\<^sub>C) (complex_of_real r)\ for r by (auto intro!: ext simp: scaleR_scaleC scaleC_prod_def scaleR_prod_def) qed end lemma module_prod_scale_eq_scaleC: "module_prod.scale (*\<^sub>C) (*\<^sub>C) = scaleC" apply (rule ext) apply (rule ext) apply (subst module_prod.scale_def) subgoal by unfold_locales by (simp add: scaleC_prod_def) interpretation complex_vector?: vector_space_prod "scaleC::_\_\'a::complex_vector" "scaleC::_\_\'b::complex_vector" rewrites "scale = ((*\<^sub>C)::_\_\('a \ 'b))" and "module.dependent (*\<^sub>C) = cdependent" and "module.representation (*\<^sub>C) = crepresentation" and "module.subspace (*\<^sub>C) = csubspace" and "module.span (*\<^sub>C) = cspan" and "vector_space.extend_basis (*\<^sub>C) = cextend_basis" and "vector_space.dim (*\<^sub>C) = cdim" and "Vector_Spaces.linear (*\<^sub>C) (*\<^sub>C) = clinear" subgoal by unfold_locales subgoal by (fact module_prod_scale_eq_scaleC) unfolding cdependent_raw_def crepresentation_raw_def csubspace_raw_def cspan_raw_def cextend_basis_raw_def cdim_raw_def clinear_def by (rule refl)+ instance prod :: (complex_normed_vector, complex_normed_vector) complex_normed_vector proof fix c :: complex and x y :: "'a \ 'b" show "norm (c *\<^sub>C x) = cmod c * norm x" unfolding norm_prod_def apply (simp add: power_mult_distrib) apply (simp add: distrib_left [symmetric]) by (simp add: real_sqrt_mult) qed lemma cspan_Times: \cspan (S \ T) = cspan S \ cspan T\ if \0 \ S\ and \0 \ T\ proof have \fst ` cspan (S \ T) \ cspan S\ apply (subst complex_vector.linear_span_image[symmetric]) using that complex_vector.module_hom_fst by auto moreover have \snd ` cspan (S \ T) \ cspan T\ apply (subst complex_vector.linear_span_image[symmetric]) using that complex_vector.module_hom_snd by auto ultimately show \cspan (S \ T) \ cspan S \ cspan T\ by auto show \cspan S \ cspan T \ cspan (S \ T)\ proof fix x assume assm: \x \ cspan S \ cspan T\ then have \fst x \ cspan S\ by auto then obtain t1 r1 where fst_x: \fst x = (\a\t1. r1 a *\<^sub>C a)\ and [simp]: \finite t1\ and \t1 \ S\ by (auto simp add: complex_vector.span_explicit) from assm have \snd x \ cspan T\ by auto then obtain t2 r2 where snd_x: \snd x = (\a\t2. r2 a *\<^sub>C a)\ and [simp]: \finite t2\ and \t2 \ T\ by (auto simp add: complex_vector.span_explicit) define t :: \('a+'b) set\ and r :: \('a+'b) \ complex\ and f :: \('a+'b) \ ('a\'b)\ where \t = t1 <+> t2\ and \r a = (case a of Inl a1 \ r1 a1 | Inr a2 \ r2 a2)\ and \f a = (case a of Inl a1 \ (a1,0) | Inr a2 \ (0,a2))\ for a have \finite t\ by (simp add: t_def) moreover have \f ` t \ S \ T\ using \t1 \ S\ \t2 \ T\ that by (auto simp: f_def t_def) moreover have \(fst x, snd x) = (\a\t. r a *\<^sub>C f a)\ apply (simp only: fst_x snd_x) by (auto simp: t_def sum.Plus r_def f_def sum_prod) ultimately show \x \ cspan (S \ T)\ apply auto by (smt (verit, best) complex_vector.span_scale complex_vector.span_sum complex_vector.span_superset image_subset_iff subset_iff) qed qed lemma onorm_case_prod_plus: \onorm (case_prod plus :: _ \ 'a::{real_normed_vector, not_singleton}) = sqrt 2\ proof - obtain x :: 'a where \x \ 0\ apply atomize_elim by auto show ?thesis apply (rule onormI[where x=\(x,x)\]) using norm_plus_leq_norm_prod apply force using \x \ 0\ by (auto simp add: zero_prod_def norm_prod_def real_sqrt_mult simp flip: scaleR_2) qed subsection \Copying existing theorems into sublocales\ context bounded_clinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = real.continuous lemmas uniform_limit = real.uniform_limit lemmas Cauchy = real.Cauchy end context bounded_antilinear begin interpretation bounded_linear f by (rule bounded_linear) lemmas continuous = real.continuous lemmas uniform_limit = real.uniform_limit end context bounded_cbilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = real.tendsto lemmas isCont = real.isCont lemmas scaleR_right = real.scaleR_right lemmas scaleR_left = real.scaleR_left end context bounded_sesquilinear begin interpretation bounded_bilinear prod by simp lemmas tendsto = real.tendsto lemmas isCont = real.isCont lemmas scaleR_right = real.scaleR_right lemmas scaleR_left = real.scaleR_left end lemmas tendsto_scaleC [tendsto_intros] = bounded_cbilinear.tendsto [OF bounded_cbilinear_scaleC] unbundle no_lattice_syntax end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_General.thy b/thys/Complex_Bounded_Operators/extra/Extra_General.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_General.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_General.thy @@ -1,687 +1,720 @@ section \\Extra_General\ -- General missing things\ theory Extra_General imports "HOL-Library.Cardinality" "HOL-Analysis.Elementary_Topology" "HOL-Analysis.Uniform_Limit" "HOL-Library.Set_Algebras" "HOL-Types_To_Sets.Types_To_Sets" "HOL-Library.Complex_Order" "HOL-Analysis.Infinite_Sum" "HOL-Cardinals.Cardinals" + "HOL-Library.Complemented_Lattices" begin subsection \Misc\ lemma reals_zero_comparable: fixes x::complex assumes "x\\" shows "x \ 0 \ x \ 0" using assms unfolding complex_is_real_iff_compare0 by assumption lemma unique_choice: "\x. \!y. Q x y \ \!f. \x. Q x (f x)" apply (auto intro!: choice ext) by metis lemma image_set_plus: assumes \linear U\ shows \U ` (A + B) = U ` A + U ` B\ unfolding image_def set_plus_def using assms by (force simp: linear_add) consts heterogenous_identity :: \'a \ 'b\ overloading heterogenous_identity_id \ "heterogenous_identity :: 'a \ 'a" begin definition heterogenous_identity_def[simp]: \heterogenous_identity_id = id\ end lemma L2_set_mono2: assumes a1: "finite L" and a2: "K \ L" shows "L2_set f K \ L2_set f L" proof- have "(\i\K. (f i)\<^sup>2) \ (\i\L. (f i)\<^sup>2)" proof (rule sum_mono2) show "finite L" using a1. show "K \ L" using a2. show "0 \ (f b)\<^sup>2" if "b \ L - K" for b :: 'a using that by simp qed hence "sqrt (\i\K. (f i)\<^sup>2) \ sqrt (\i\L. (f i)\<^sup>2)" by (rule real_sqrt_le_mono) thus ?thesis unfolding L2_set_def. qed lemma Sup_real_close: fixes e :: real assumes "0 < e" and S: "bdd_above S" "S \ {}" shows "\x\S. Sup S - e < x" proof - have \Sup (ereal ` S) \ \\ by (metis assms(2) bdd_above_def ereal_less_eq(3) less_SUP_iff less_ereal.simps(4) not_le) moreover have \Sup (ereal ` S) \ -\\ by (simp add: SUP_eq_iff assms(3)) ultimately have Sup_bdd: \\Sup (ereal ` S)\ \ \\ by auto then have \\x'\ereal ` S. Sup (ereal ` S) - ereal e < x'\ apply (rule_tac Sup_ereal_close) using assms by auto then obtain x where \x \ S\ and Sup_x: \Sup (ereal ` S) - ereal e < ereal x\ by auto have \Sup (ereal ` S) = ereal (Sup S)\ using Sup_bdd by (rule ereal_Sup[symmetric]) with Sup_x have \ereal (Sup S - e) < ereal x\ by auto then have \Sup S - e < x\ by auto with \x \ S\ show ?thesis by auto qed text \Improved version of @{attribute internalize_sort}: It is not necessary to specify the sort of the type variable.\ attribute_setup internalize_sort' = \let fun find_tvar thm v = let val tvars = Term.add_tvars (Thm.prop_of thm) [] val tv = case find_first (fn (n,sort) => n=v) tvars of SOME tv => tv | NONE => raise THM ("Type variable " ^ string_of_indexname v ^ " not found", 0, [thm]) in TVar tv end fun internalize_sort_attr (tvar:indexname) = Thm.rule_attribute [] (fn context => fn thm => (snd (Internalize_Sort.internalize_sort (Thm.ctyp_of (Context.proof_of context) (find_tvar thm tvar)) thm))); in Scan.lift Args.var >> internalize_sort_attr end\ "internalize a sort" lemma card_prod_omega: \X *c natLeq =o X\ if \Cinfinite X\ by (simp add: Cinfinite_Cnotzero cprod_infinite1' natLeq_Card_order natLeq_cinfinite natLeq_ordLeq_cinfinite that) lemma countable_leq_natLeq: \|X| \o natLeq\ if \countable X\ using subset_range_from_nat_into[OF that] by (meson card_of_nat ordIso_iff_ordLeq ordLeq_transitive surj_imp_ordLeq) lemma set_Times_plus_distrib: \(A \ B) + (C \ D) = (A + C) \ (B + D)\ by (auto simp: Sigma_def set_plus_def) subsection \Not singleton\ class not_singleton = assumes not_singleton_card: "\x y. x \ y" lemma not_singleton_existence[simp]: \\ x::('a::not_singleton). x \ t\ using not_singleton_card[where ?'a = 'a] by (metis (full_types)) lemma UNIV_not_singleton[simp]: "(UNIV::_::not_singleton set) \ {x}" using not_singleton_existence[of x] by blast lemma UNIV_not_singleton_converse: assumes"\x::'a. UNIV \ {x}" shows "\x::'a. \y. x \ y" using assms by fastforce subclass (in card2) not_singleton apply standard using two_le_card by (meson card_2_iff' obtain_subset_with_card_n) subclass (in perfect_space) not_singleton apply intro_classes by (metis (mono_tags) Collect_cong Collect_mem_eq UNIV_I local.UNIV_not_singleton local.not_open_singleton local.open_subopen) lemma class_not_singletonI_monoid_add: assumes "(UNIV::'a set) \ {0}" shows "class.not_singleton TYPE('a::monoid_add)" proof intro_classes let ?univ = "UNIV :: 'a set" from assms obtain x::'a where "x \ 0" by auto thus "\x y :: 'a. x \ y" by auto qed lemma not_singleton_vs_CARD_1: assumes \\ class.not_singleton TYPE('a)\ shows \class.CARD_1 TYPE('a)\ using assms unfolding class.not_singleton_def class.CARD_1_def by (metis (full_types) One_nat_def UNIV_I card.empty card.insert empty_iff equalityI finite.intros(1) insert_iff subsetI) subsection \\<^class>\CARD_1\\ context CARD_1 begin lemma everything_the_same[simp]: "(x::'a)=y" by (metis (full_types) UNIV_I card_1_singletonE empty_iff insert_iff local.CARD_1) lemma CARD_1_UNIV: "UNIV = {x::'a}" by (metis (full_types) UNIV_I card_1_singletonE local.CARD_1 singletonD) lemma CARD_1_ext: "x (a::'a) = y b \ x = y" proof (rule ext) show "x t = y t" if "x a = y b" for t :: 'a using that apply (subst (asm) everything_the_same[where x=a]) apply (subst (asm) everything_the_same[where x=b]) by simp qed end instance unit :: CARD_1 apply standard by auto instance prod :: (CARD_1, CARD_1) CARD_1 apply intro_classes by (simp add: CARD_1) instance "fun" :: (CARD_1, CARD_1) CARD_1 apply intro_classes by (auto simp add: card_fun CARD_1) lemma enum_CARD_1: "(Enum.enum :: 'a::{CARD_1,enum} list) = [a]" proof - let ?enum = "Enum.enum :: 'a::{CARD_1,enum} list" have "length ?enum = 1" apply (subst card_UNIV_length_enum[symmetric]) by (rule CARD_1) then obtain b where "?enum = [b]" apply atomize_elim apply (cases ?enum, auto) by (metis length_0_conv length_Cons nat.inject) thus "?enum = [a]" by (subst everything_the_same[of _ b], simp) qed +lemma card_not_singleton: \CARD('a::not_singleton) \ 1\ + by (simp add: card_1_singleton_iff) subsection \Topology\ lemma cauchy_filter_metricI: fixes F :: "'a::metric_space filter" assumes "\e. e>0 \ \P. eventually P F \ (\x y. P x \ P y \ dist x y < e)" shows "cauchy_filter F" proof (unfold cauchy_filter_def le_filter_def, auto) fix P :: "'a \ 'a \ bool" assume "eventually P uniformity" then obtain e where e: "e > 0" and P: "dist x y < e \ P (x, y)" for x y unfolding eventually_uniformity_metric by auto obtain P' where evP': "eventually P' F" and P'_dist: "P' x \ P' y \ dist x y < e" for x y apply atomize_elim using assms e by auto from evP' P'_dist P show "eventually P (F \\<^sub>F F)" unfolding eventually_uniformity_metric eventually_prod_filter eventually_filtermap by metis qed lemma cauchy_filter_metric_filtermapI: fixes F :: "'a filter" and f :: "'a\'b::metric_space" assumes "\e. e>0 \ \P. eventually P F \ (\x y. P x \ P y \ dist (f x) (f y) < e)" shows "cauchy_filter (filtermap f F)" proof (rule cauchy_filter_metricI) fix e :: real assume e: "e > 0" with assms obtain P where evP: "eventually P F" and dist: "P x \ P y \ dist (f x) (f y) < e" for x y by atomize_elim auto define P' where "P' y = (\x. P x \ y = f x)" for y have "eventually P' (filtermap f F)" unfolding eventually_filtermap P'_def using evP by (smt eventually_mono) moreover have "P' x \ P' y \ dist x y < e" for x y unfolding P'_def using dist by metis ultimately show "\P. eventually P (filtermap f F) \ (\x y. P x \ P y \ dist x y < e)" by auto qed lemma tendsto_add_const_iff: \ \This is a generalization of \Limits.tendsto_add_const_iff\, the only difference is that the sort here is more general.\ "((\x. c + f x :: 'a::topological_group_add) \ c + d) F \ (f \ d) F" using tendsto_add[OF tendsto_const[of c], of f d] and tendsto_add[OF tendsto_const[of "-c"], of "\x. c + f x" "c + d"] by auto lemma finite_subsets_at_top_minus: assumes "A\B" shows "finite_subsets_at_top (B - A) \ filtermap (\F. F - A) (finite_subsets_at_top B)" proof (rule filter_leI) fix P assume "eventually P (filtermap (\F. F - A) (finite_subsets_at_top B))" then obtain X where "finite X" and "X \ B" and P: "finite Y \ X \ Y \ Y \ B \ P (Y - A)" for Y unfolding eventually_filtermap eventually_finite_subsets_at_top by auto hence "finite (X-A)" and "X-A \ B - A" by auto moreover have "finite Y \ X-A \ Y \ Y \ B - A \ P Y" for Y using P[where Y="Y\X"] \finite X\ \X \ B\ by (metis Diff_subset Int_Diff Un_Diff finite_Un inf.orderE le_sup_iff sup.orderE sup_ge2) ultimately show "eventually P (finite_subsets_at_top (B - A))" unfolding eventually_finite_subsets_at_top by meson qed lemma finite_subsets_at_top_inter: assumes "A\B" shows "filtermap (\F. F \ A) (finite_subsets_at_top B) = finite_subsets_at_top A" proof (subst filter_eq_iff, intro allI iffI) fix P :: "'a set \ bool" assume "eventually P (finite_subsets_at_top A)" then show "eventually P (filtermap (\F. F \ A) (finite_subsets_at_top B))" unfolding eventually_filtermap unfolding eventually_finite_subsets_at_top by (metis Int_subset_iff assms finite_Int inf_le2 subset_trans) next fix P :: "'a set \ bool" assume "eventually P (filtermap (\F. F \ A) (finite_subsets_at_top B))" then obtain X where \finite X\ \X \ B\ and P: \finite Y \ X \ Y \ Y \ B \ P (Y \ A)\ for Y unfolding eventually_filtermap eventually_finite_subsets_at_top by metis have *: \finite Y \ X \ A \ Y \ Y \ A \ P Y\ for Y using P[where Y=\Y \ (B-A)\] apply (subgoal_tac \(Y \ (B - A)) \ A = Y\) apply (smt (verit, best) Int_Un_distrib2 Int_Un_eq(4) P Un_subset_iff \X \ B\ \finite X\ assms finite_UnI inf.orderE sup_ge2) by auto show "eventually P (finite_subsets_at_top A)" unfolding eventually_finite_subsets_at_top apply (rule exI[of _ \X\A\]) by (auto simp: \finite X\ intro!: *) qed lemma tendsto_principal_singleton: shows "(f \ f x) (principal {x})" unfolding tendsto_def eventually_principal by simp lemma complete_singleton: "complete {s::'a::uniform_space}" proof- have "F \ principal {s} \ F \ bot \ cauchy_filter F \ F \ nhds s" for F by (metis eventually_nhds eventually_principal le_filter_def singletonD) thus ?thesis unfolding complete_uniform by simp qed lemma on_closure_eqI: fixes f g :: \'a::topological_space \ 'b::t2_space\ assumes eq: \\x. x \ S \ f x = g x\ assumes xS: \x \ closure S\ assumes cont: \continuous_on UNIV f\ \continuous_on UNIV g\ shows \f x = g x\ proof - define X where \X = {x. f x = g x}\ have \closed X\ using cont by (simp add: X_def closed_Collect_eq) moreover have \S \ X\ by (simp add: X_def eq subsetI) ultimately have \closure S \ X\ using closure_minimal by blast with xS have \x \ X\ by auto then show ?thesis using X_def by blast qed lemma on_closure_leI: fixes f g :: \'a::topological_space \ 'b::linorder_topology\ assumes eq: \\x. x \ S \ f x \ g x\ assumes xS: \x \ closure S\ assumes cont: \continuous_on UNIV f\ \continuous_on UNIV g\ (* Is "isCont f x" "isCont g x" sufficient? *) shows \f x \ g x\ proof - define X where \X = {x. f x \ g x}\ have \closed X\ using cont by (simp add: X_def closed_Collect_le) moreover have \S \ X\ by (simp add: X_def eq subsetI) ultimately have \closure S \ X\ using closure_minimal by blast with xS have \x \ X\ by auto then show ?thesis using X_def by blast qed lemma tendsto_compose_at_within: assumes f: "(f \ y) F" and g: "(g \ z) (at y within S)" and fg: "eventually (\w. f w = y \ g y = z) F" and fS: \\\<^sub>F w in F. f w \ S\ shows "((g \ f) \ z) F" proof (cases \g y = z\) case False then have 1: "(\\<^sub>F a in F. f a \ y)" using fg by force have 2: "(g \ z) (filtermap f F) \ \ (\\<^sub>F a in F. f a \ y)" by (smt (verit, best) eventually_elim2 f fS filterlim_at filterlim_def g tendsto_mono) show ?thesis using "1" "2" tendsto_compose_filtermap by blast next case True have *: ?thesis if \(g \ z) (filtermap f F)\ using that by (simp add: tendsto_compose_filtermap) from g have \(g \ g y) (inf (nhds y) (principal (S-{y})))\ by (simp add: True at_within_def) then have g': \(g \ g y) (inf (nhds y) (principal S))\ using True g tendsto_at_iff_tendsto_nhds_within by blast from f have \filterlim f (nhds y) F\ by - then have f': \filterlim f (inf (nhds y) (principal S)) F\ using fS by (simp add: filterlim_inf filterlim_principal) from f' g' show ?thesis by (simp add: * True filterlim_compose filterlim_filtermap) qed subsection \Sums\ lemma sum_single: assumes "finite A" assumes "\j. j \ i \ j\A \ f j = 0" shows "sum f A = (if i\A then f i else 0)" apply (subst sum.mono_neutral_cong_right[where S=\A \ {i}\ and h=f]) using assms by auto lemma has_sum_comm_additive_general: \ \This is a strengthening of @{thm [source] has_sum_comm_additive_general}.\ fixes f :: \'b :: {comm_monoid_add,topological_space} \ 'c :: {comm_monoid_add,topological_space}\ assumes f_sum: \\F. finite F \ F \ S \ sum (f o g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes inS: \\F. finite F \ sum g F \ T\ assumes cont: \(f \ f x) (at x within T)\ \ \For \<^class>\t2_space\ and \<^term>\T=UNIV\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes infsum: \has_sum g S x\ shows \has_sum (f o g) S (f x)\ proof - have \(sum g \ x) (finite_subsets_at_top S)\ using infsum has_sum_def by blast then have \((f o sum g) \ f x) (finite_subsets_at_top S)\ apply (rule tendsto_compose_at_within[where S=T]) using assms by auto then have \(sum (f o g) \ f x) (finite_subsets_at_top S)\ apply (rule tendsto_cong[THEN iffD1, rotated]) using f_sum by fastforce then show \has_sum (f o g) S (f x)\ using has_sum_def by blast qed lemma summable_on_comm_additive_general: \ \This is a strengthening of @{thm [source] summable_on_comm_additive_general}.\ fixes g :: \'a \ 'b :: {comm_monoid_add,topological_space}\ and f :: \'b \ 'c :: {comm_monoid_add,topological_space}\ assumes \\F. finite F \ F \ S \ sum (f o g) F = f (sum g F)\ \ \Not using \<^const>\additive\ because it would add sort constraint \<^class>\ab_group_add\\ assumes inS: \\F. finite F \ sum g F \ T\ assumes cont: \\x. has_sum g S x \ (f \ f x) (at x within T)\ \ \For \<^class>\t2_space\ and \<^term>\T=UNIV\, this is equivalent to \isCont f x\ by @{thm [source] isCont_def}.\ assumes \g summable_on S\ shows \(f o g) summable_on S\ by (meson assms summable_on_def has_sum_comm_additive_general has_sum_def infsum_tendsto) lemma has_sum_metric: fixes l :: \'a :: {metric_space, comm_monoid_add}\ shows \has_sum f A l \ (\e. e > 0 \ (\X. finite X \ X \ A \ (\Y. finite Y \ X \ Y \ Y \ A \ dist (sum f Y) l < e)))\ unfolding has_sum_def apply (subst tendsto_iff) unfolding eventually_finite_subsets_at_top by simp lemma summable_on_product_finite_left: fixes f :: \'a\'b \ 'c::{topological_comm_monoid_add}\ assumes sum: \\x. x\X \ (\y. f(x,y)) summable_on Y\ assumes \finite X\ shows \f summable_on (X\Y)\ using \finite X\ subset_refl[of X] proof (induction rule: finite_subset_induct') case empty then show ?case by simp next case (insert x F) have *: \bij_betw (Pair x) Y ({x} \ Y)\ apply (rule bij_betwI') by auto from sum[of x] have \f summable_on {x} \ Y\ apply (rule summable_on_reindex_bij_betw[THEN iffD1, rotated]) by (simp_all add: * insert.hyps(2)) then have \f summable_on {x} \ Y \ F \ Y\ apply (rule summable_on_Un_disjoint) using insert by auto then show ?case by (metis Sigma_Un_distrib1 insert_is_Un) qed lemma summable_on_product_finite_right: fixes f :: \'a\'b \ 'c::{topological_comm_monoid_add}\ assumes sum: \\y. y\Y \ (\x. f(x,y)) summable_on X\ assumes \finite Y\ shows \f summable_on (X\Y)\ proof - have \(\(y,x). f(x,y)) summable_on (Y\X)\ apply (rule summable_on_product_finite_left) using assms by auto then show ?thesis apply (subst summable_on_reindex_bij_betw[where g=prod.swap and A=\Y\X\, symmetric]) apply (simp add: bij_betw_def product_swap) by (metis (mono_tags, lifting) case_prod_unfold prod.swap_def summable_on_cong) qed subsection \Complex numbers\ lemma cmod_Re: assumes "x \ 0" shows "cmod x = Re x" using assms unfolding less_eq_complex_def cmod_def by auto lemma abs_complex_real[simp]: "abs x \ \" for x :: complex by (simp add: abs_complex_def) lemma Im_abs[simp]: "Im (abs x) = 0" using abs_complex_real complex_is_Real_iff by blast lemma cnj_x_x: "cnj x * x = (abs x)\<^sup>2" proof (cases x) show "cnj x * x = \x\\<^sup>2" if "x = Complex x1 x2" for x1 :: real and x2 :: real using that by (auto simp: complex_cnj complex_mult abs_complex_def complex_norm power2_eq_square complex_of_real_def) qed lemma cnj_x_x_geq0[simp]: \cnj x * x \ 0\ by (simp add: less_eq_complex_def) lemma complex_of_real_leq_1_iff[iff]: \complex_of_real x \ 1 \ x \ 1\ by (simp add: less_eq_complex_def) lemma x_cnj_x: \x * cnj x = (abs x)\<^sup>2\ by (metis cnj_x_x mult.commute) subsection \List indices and enum\ fun index_of where "index_of x [] = (0::nat)" | "index_of x (y#ys) = (if x=y then 0 else (index_of x ys + 1))" definition "enum_idx (x::'a::enum) = index_of x (enum_class.enum :: 'a list)" lemma index_of_length: "index_of x y \ length y" apply (induction y) by auto lemma index_of_correct: assumes "x \ set y" shows "y ! index_of x y = x" using assms apply (induction y arbitrary: x) by auto lemma enum_idx_correct: "Enum.enum ! enum_idx i = i" proof- have "i \ set enum_class.enum" using UNIV_enum by blast thus ?thesis unfolding enum_idx_def using index_of_correct by metis qed lemma index_of_bound: assumes "y \ []" and "x \ set y" shows "index_of x y < length y" using assms proof(induction y arbitrary: x) case Nil thus ?case by auto next case (Cons a y) show ?case proof(cases "a = x") case True thus ?thesis by auto next case False moreover have "a \ x \ index_of x y < length y" using Cons.IH Cons.prems(2) by fastforce ultimately show ?thesis by auto qed qed lemma enum_idx_bound: "enum_idx x < length (Enum.enum :: 'a list)" for x :: "'a::enum" proof- have p1: "False" if "(Enum.enum :: 'a list) = []" proof- have "(UNIV::'a set) = set ([]::'a list)" using that UNIV_enum by metis also have "\ = {}" by blast finally have "(UNIV::'a set) = {}". thus ?thesis by simp qed have p2: "x \ set (Enum.enum :: 'a list)" using UNIV_enum by auto moreover have "(enum_class.enum::'a list) \ []" using p2 by auto ultimately show ?thesis unfolding enum_idx_def using index_of_bound [where x = x and y = "(Enum.enum :: 'a list)"] by auto qed lemma index_of_nth: assumes "distinct xs" assumes "i < length xs" shows "index_of (xs ! i) xs = i" using assms by (metis gr_implies_not_zero index_of_bound index_of_correct length_0_conv nth_eq_iff_index_eq nth_mem) lemma enum_idx_enum: assumes \i < CARD('a::enum)\ shows \enum_idx (enum_class.enum ! i :: 'a) = i\ unfolding enum_idx_def apply (rule index_of_nth) using assms by (simp_all add: card_UNIV_length_enum enum_distinct) subsection \Filtering lists/sets\ lemma map_filter_map: "List.map_filter f (map g l) = List.map_filter (f o g) l" proof (induction l) show "List.map_filter f (map g []) = List.map_filter (f \ g) []" by (simp add: map_filter_simps) show "List.map_filter f (map g (a # l)) = List.map_filter (f \ g) (a # l)" if "List.map_filter f (map g l) = List.map_filter (f \ g) l" for a :: 'c and l :: "'c list" using that map_filter_simps(1) by (metis comp_eq_dest_lhs list.simps(9)) qed lemma map_filter_Some[simp]: "List.map_filter (\x. Some (f x)) l = map f l" proof (induction l) show "List.map_filter (\x. Some (f x)) [] = map f []" by (simp add: map_filter_simps) show "List.map_filter (\x. Some (f x)) (a # l) = map f (a # l)" if "List.map_filter (\x. Some (f x)) l = map f l" for a :: 'b and l :: "'b list" using that by (simp add: map_filter_simps(1)) qed lemma filter_Un: "Set.filter f (x \ y) = Set.filter f x \ Set.filter f y" unfolding Set.filter_def by auto lemma Set_filter_unchanged: "Set.filter P X = X" if "\x. x\X \ P x" for P and X :: "'z set" using that unfolding Set.filter_def by auto subsection \Maps\ definition "inj_map \ = (\x y. \ x = \ y \ \ x \ None \ x = y)" definition "inv_map \ = (\y. if Some y \ range \ then Some (inv \ (Some y)) else None)" lemma inj_map_total[simp]: "inj_map (Some o \) = inj \" unfolding inj_map_def inj_def by simp lemma inj_map_Some[simp]: "inj_map Some" by (simp add: inj_map_def) lemma inv_map_total: assumes "surj \" shows "inv_map (Some o \) = Some o inv \" proof- have "(if Some y \ range (\x. Some (\ x)) then Some (SOME x. Some (\ x) = Some y) else None) = Some (SOME b. \ b = y)" if "surj \" for y using that by auto hence "surj \ \ (\y. if Some y \ range (\x. Some (\ x)) then Some (SOME x. Some (\ x) = Some y) else None) = (\x. Some (SOME xa. \ xa = x))" by (rule ext) thus ?thesis unfolding inv_map_def o_def inv_def using assms by linarith qed lemma inj_map_map_comp[simp]: assumes a1: "inj_map f" and a2: "inj_map g" shows "inj_map (f \\<^sub>m g)" using a1 a2 unfolding inj_map_def by (metis (mono_tags, lifting) map_comp_def option.case_eq_if option.expand) lemma inj_map_inv_map[simp]: "inj_map (inv_map \)" proof (unfold inj_map_def, rule allI, rule allI, rule impI, erule conjE) fix x y assume same: "inv_map \ x = inv_map \ y" and pix_not_None: "inv_map \ x \ None" have x_pi: "Some x \ range \" using pix_not_None unfolding inv_map_def apply auto by (meson option.distinct(1)) have y_pi: "Some y \ range \" using pix_not_None unfolding same unfolding inv_map_def apply auto by (meson option.distinct(1)) have "inv_map \ x = Some (Hilbert_Choice.inv \ (Some x))" unfolding inv_map_def using x_pi by simp moreover have "inv_map \ y = Some (Hilbert_Choice.inv \ (Some y))" unfolding inv_map_def using y_pi by simp ultimately have "Hilbert_Choice.inv \ (Some x) = Hilbert_Choice.inv \ (Some y)" using same by simp thus "x = y" by (meson inv_into_injective option.inject x_pi y_pi) qed +subsection \Lattices\ + +unbundle lattice_syntax + +text \The following lemma is identical to @{thm [source] Complete_Lattices.uminus_Inf} + except for the more general sort.\ +lemma uminus_Inf: "- (\A) = \(uminus ` A)" for A :: \'a::complete_orthocomplemented_lattice set\ +proof (rule order.antisym) + show "- \A \ \(uminus ` A)" + by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp + show "\(uminus ` A) \ - \A" + by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto +qed + +text \The following lemma is identical to @{thm [source] Complete_Lattices.uminus_INF} + except for the more general sort.\ +lemma uminus_INF: "- (INF x\A. B x) = (SUP x\A. - B x)" for B :: \'a \ 'b::complete_orthocomplemented_lattice\ + by (simp add: uminus_Inf image_image) + +text \The following lemma is identical to @{thm [source] Complete_Lattices.uminus_Sup} + except for the more general sort.\ +lemma uminus_Sup: "- (\A) = \(uminus ` A)" for A :: \'a::complete_orthocomplemented_lattice set\ + by (metis (no_types, lifting) uminus_INF image_cong image_ident ortho_involution) + +text \The following lemma is identical to @{thm [source] Complete_Lattices.uminus_SUP} + except for the more general sort.\ +lemma uminus_SUP: "- (SUP x\A. B x) = (INF x\A. - B x)" for B :: \'a \ 'b::complete_orthocomplemented_lattice\ + by (simp add: uminus_Sup image_image) + + end diff --git a/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy b/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy --- a/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy +++ b/thys/Linear_Recurrences/Linear_Recurrences_Misc.thy @@ -1,227 +1,225 @@ (* File: Linear_Recurrences_Misc.thy Author: Manuel Eberl, TU München *) section \Miscellaneous material required for linear recurrences\ theory Linear_Recurrences_Misc imports Complex_Main "HOL-Computational_Algebra.Computational_Algebra" "HOL-Computational_Algebra.Polynomial_Factorial" begin fun zip_with where "zip_with f (x#xs) (y#ys) = f x y # zip_with f xs ys" | "zip_with f _ _ = []" lemma length_zip_with [simp]: "length (zip_with f xs ys) = min (length xs) (length ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_altdef: "zip_with f xs ys = map (\(x,y). f x y) (zip xs ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_nth [simp]: "n < length xs \ n < length ys \ zip_with f xs ys ! n = f (xs!n) (ys!n)" by (simp add: zip_with_altdef) lemma take_zip_with: "take n (zip_with f xs ys) = zip_with f (take n xs) (take n ys)" proof (induction f xs ys arbitrary: n rule: zip_with.induct) case (1 f x xs y ys n) thus ?case by (cases n) simp_all qed simp_all lemma drop_zip_with: "drop n (zip_with f xs ys) = zip_with f (drop n xs) (drop n ys)" proof (induction f xs ys arbitrary: n rule: zip_with.induct) case (1 f x xs y ys n) thus ?case by (cases n) simp_all qed simp_all lemma map_zip_with: "map f (zip_with g xs ys) = zip_with (\x y. f (g x y)) xs ys" by (induction g xs ys rule: zip_with.induct) simp_all lemma zip_with_map: "zip_with f (map g xs) (map h ys) = zip_with (\x y. f (g x) (h y)) xs ys" by (induction "\x y. f (g x) (h y)" xs ys rule: zip_with.induct) simp_all lemma zip_with_map_left: "zip_with f (map g xs) ys = zip_with (\x y. f (g x) y) xs ys" using zip_with_map[of f g xs "\x. x" ys] by simp lemma zip_with_map_right: "zip_with f xs (map g ys) = zip_with (\x y. f x (g y)) xs ys" using zip_with_map[of f "\x. x" xs g ys] by simp lemma zip_with_swap: "zip_with (\x y. f y x) xs ys = zip_with f ys xs" by (induction f ys xs rule: zip_with.induct) simp_all lemma set_zip_with: "set (zip_with f xs ys) = (\(x,y). f x y) ` set (zip xs ys)" by (induction f xs ys rule: zip_with.induct) simp_all lemma zip_with_Pair: "zip_with Pair (xs :: 'a list) (ys :: 'b list) = zip xs ys" by (induction "Pair :: 'a \ 'b \ _" xs ys rule: zip_with.induct) simp_all lemma zip_with_altdef': "zip_with f xs ys = [f (xs!i) (ys!i). i \ [0.. [0.. 0" shows "card {x. poly p x = 0} \ degree p" using assms proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) show ?case proof (cases "\x. poly p x = 0") case False hence "{x. poly p x = 0} = {}" by blast thus ?thesis by simp next case True then obtain x where x: "poly p x = 0" by blast hence "[:-x, 1:] dvd p" by (subst (asm) poly_eq_0_iff_dvd) then obtain q where q: "p = [:-x, 1:] * q" by (auto simp: dvd_def) with \p \ 0\ have [simp]: "q \ 0" by auto have deg: "degree p = Suc (degree q)" by (subst q, subst degree_mult_eq) auto have "card {x. poly p x = 0} \ card (insert x {x. poly q x = 0})" by (intro card_mono) (auto intro: poly_roots_finite simp: q) also have "\ \ Suc (card {x. poly q x = 0})" by (rule card_insert_le_m1) auto also from deg have "card {x. poly q x = 0} \ degree q" using \p \ 0\ and q by (intro less) auto also have "Suc \ = degree p" by (simp add: deg) finally show ?thesis by - simp_all qed qed lemma poly_eqI_degree: fixes p q :: "'a :: {comm_ring_1, ring_no_zero_divisors} poly" assumes "\x. x \ A \ poly p x = poly q x" assumes "card A > degree p" "card A > degree q" shows "p = q" proof (rule ccontr) assume neq: "p \ q" have "degree (p - q) \ max (degree p) (degree q)" by (rule degree_diff_le_max) also from assms have "\ < card A" by linarith also have "\ \ card {x. poly (p - q) x = 0}" using neq and assms by (intro card_mono poly_roots_finite) auto finally have "degree (p - q) < card {x. poly (p - q) x = 0}" . moreover have "degree (p - q) \ card {x. poly (p - q) x = 0}" using neq by (intro card_poly_roots_bound) auto ultimately show False by linarith qed lemma poly_root_order_induct [case_names 0 no_roots root]: fixes p :: "'a :: idom poly" assumes "P 0" "\p. (\x. poly p x \ 0) \ P p" "\p x n. n > 0 \ poly p x \ 0 \ P p \ P ([:-x, 1:] ^ n * p)" shows "P p" proof (induction "degree p" arbitrary: p rule: less_induct) case (less p) consider "p = 0" | "p \ 0" "\x. poly p x = 0" | "\x. poly p x \ 0" by blast thus ?case proof cases case 3 with assms(2)[of p] show ?thesis by simp next case 2 then obtain x where x: "poly p x = 0" by auto have "[:-x, 1:] ^ order x p dvd p" by (intro order_1) then obtain q where q: "p = [:-x, 1:] ^ order x p * q" by (auto simp: dvd_def) with 2 have [simp]: "q \ 0" by auto have order_pos: "order x p > 0" using \p \ 0\ and x by (auto simp: order_root) have "order x p = order x p + order x q" by (subst q, subst order_mult) (auto simp: order_power_n_n) hence [simp]: "order x q = 0" by simp have deg: "degree p = order x p + degree q" by (subst q, subst degree_mult_eq) (auto simp: degree_power_eq) with order_pos have "degree q < degree p" by simp hence "P q" by (rule less) with order_pos have "P ([:-x, 1:] ^ order x p * q)" by (intro assms(3)) (auto simp: order_root) with q show ?thesis by simp qed (simp_all add: assms(1)) qed lemma complex_poly_decompose: "smult (lead_coeff p) (\z|poly p z = 0. [:-z, 1:] ^ order z p) = (p :: complex poly)" proof (induction p rule: poly_root_order_induct) case (no_roots p) show ?case proof (cases "degree p = 0") case False hence "\constant (poly p)" by (subst constant_degree) with fundamental_theorem_of_algebra and no_roots show ?thesis by blast qed (auto elim!: degree_eq_zeroE) next case (root p x n) from root have *: "{z. poly ([:- x, 1:] ^ n * p) z = 0} = insert x {z. poly p z = 0}" by auto have "smult (lead_coeff ([:-x, 1:] ^ n * p)) (\z|poly ([:-x,1:] ^ n * p) z = 0. [:-z, 1:] ^ order z ([:- x, 1:] ^ n * p)) = [:- x, 1:] ^ order x ([:- x, 1:] ^ n * p) * smult (lead_coeff p) (\z\{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p))" by (subst *, subst prod.insert) (insert root, auto intro: poly_roots_finite simp: mult_ac lead_coeff_mult lead_coeff_power) also have "order x ([:- x, 1:] ^ n * p) = n" using root by (subst order_mult) (auto simp: order_power_n_n order_0I) also have "(\z\{z. poly p z = 0}. [:- z, 1:] ^ order z ([:- x, 1:] ^ n * p)) = (\z\{z. poly p z = 0}. [:- z, 1:] ^ order z p)" proof (intro prod.cong refl, goal_cases) case (1 y) with root have "order y ([:-x,1:] ^ n) = 0" by (intro order_0I) auto thus ?case using root by (subst order_mult) auto qed also note root.IH finally show ?case . qed simp_all lemma normalize_field: "normalize (x :: 'a :: {normalization_semidom,field}) = (if x = 0 then 0 else 1)" by (auto simp: normalize_1_iff dvd_field_iff) lemma unit_factor_field [simp]: "unit_factor (x :: 'a :: {normalization_semidom,field}) = x" using unit_factor_mult_normalize[of x] normalize_field[of x] by (simp split: if_splits) lemma coprime_linear_poly: fixes c :: "'a :: field_gcd" assumes "c \ c'" shows "coprime [:c,1:] [:c',1:]" proof - have "gcd [:c,1:] [:c',1:] = gcd ([:c,1:] - [:c',1:]) [:c',1:]" by (rule gcd_diff1 [symmetric]) also have "[:c,1:] - [:c',1:] = [:c-c':]" by simp also from assms have "gcd \ [:c',1:] = normalize [:c-c':]" by (intro gcd_proj1_if_dvd) (auto simp: const_poly_dvd_iff dvd_field_iff) also from assms have "\ = 1" by (simp add: normalize_poly_def) finally show "coprime [:c,1:] [:c',1:]" by (simp add: gcd_eq_1_imp_coprime) qed lemma coprime_linear_poly': fixes c :: "'a :: field_gcd" assumes "c \ c'" "c \ 0" "c' \ 0" shows "coprime [:1,c:] [:1,c':]" proof - have "gcd [:1,c:] [:1,c':] = gcd ([:1,c:] mod [:1,c':]) [:1,c':]" by simp - also have "eucl_rel_poly [:1, c:] [:1, c':] ([:c/c':], [:1-c/c':])" - using assms by (auto simp add: eucl_rel_poly_iff one_pCons) - hence "[:1,c:] mod [:1,c':] = [:1 - c / c':]" - by (rule mod_poly_eq) + also have \[:1,c:] mod [:1,c':] = [:1 - c / c':]\ + using \c' \ 0\ by (simp add: mod_pCons) also from assms have "gcd \ [:1,c':] = normalize ([:1 - c / c':])" by (intro gcd_proj1_if_dvd) (auto simp: const_poly_dvd_iff dvd_field_iff) also from assms have "\ = 1" by (auto simp: normalize_poly_def) finally show ?thesis by (rule gcd_eq_1_imp_coprime) qed end diff --git a/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy b/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy --- a/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy +++ b/thys/Polynomial_Interpolation/Ring_Hom_Poly.thy @@ -1,450 +1,455 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) section \Connecting Polynomials with Homomorphism Locales\ theory Ring_Hom_Poly imports "HOL-Computational_Algebra.Euclidean_Algorithm" Ring_Hom Missing_Polynomial begin text\@{term poly} as a homomorphism. Note that types differ.\ interpretation poly_hom: comm_semiring_hom "\p. poly p a" by (unfold_locales, auto) interpretation poly_hom: comm_ring_hom "\p. poly p a".. interpretation poly_hom: idom_hom "\p. poly p a".. text \@{term "(\\<^sub>p)"} as a homomorphism.\ interpretation pcompose_hom: comm_semiring_hom "\q. q \\<^sub>p p" using pcompose_add pcompose_mult pcompose_1 by (unfold_locales, auto) interpretation pcompose_hom: comm_ring_hom "\q. q \\<^sub>p p" .. interpretation pcompose_hom: idom_hom "\q. q \\<^sub>p p" .. definition eval_poly :: "('a \ 'b :: comm_semiring_1) \ 'a :: zero poly \ 'b \ 'b" where [code del]: "eval_poly h p = poly (map_poly h p)" lemma eval_poly_code[code]: "eval_poly h p x = fold_coeffs (\ a b. h a + x * b) p 0" by (induct p, auto simp: eval_poly_def) lemma eval_poly_as_sum: fixes h :: "'a :: zero \ 'b :: comm_semiring_1" assumes "h 0 = 0" shows "eval_poly h p x = (\i\degree p. x^i * h (coeff p i))" unfolding eval_poly_def proof (induct p) case 0 show ?case using assms by simp next case (pCons a p) thus ?case proof (cases "p = 0") case True show ?thesis by (simp add: True map_poly_simps assms) next case False show ?thesis unfolding degree_pCons_eq[OF False] unfolding sum.atMost_Suc_shift unfolding map_poly_pCons[OF pCons(1)] by (simp add: pCons(2) sum_distrib_left mult.assoc) qed qed lemma coeff_const: "coeff [: a :] i = (if i = 0 then a else 0)" by (metis coeff_monom monom_0) lemma x_as_monom: "[:0,1:] = monom 1 1" by (simp add: monom_0 monom_Suc) lemma x_pow_n: "monom 1 1 ^ n = monom 1 n" by (induct n) (simp_all add: monom_0 monom_Suc) lemma map_poly_eval_poly: assumes h0: "h 0 = 0" shows "map_poly h p = eval_poly (\ a. [: h a :]) p [:0,1:]" (is "?mp = ?ep") proof (rule poly_eqI) fix i :: nat have 2: "(\x\i. \xa\degree p. (if xa = x then 1 else 0) * coeff [:h (coeff p xa):] (i - x)) = h (coeff p i)" (is "sum ?f ?s = ?r") proof - have "sum ?f ?s = ?f i + sum ?f ({..i} - {i})" by (rule sum.remove[of _ i], auto) also have "sum ?f ({..i} - {i}) = 0" by (rule sum.neutral, intro ballI, rule sum.neutral, auto simp: coeff_const) also have "?f i = (\xa\degree p. (if xa = i then 1 else 0) * h (coeff p xa))" (is "_ = ?m") unfolding coeff_const by simp also have "\ = ?r" proof (cases "i \ degree p") case True show ?thesis by (subst sum.remove[of _ i], insert True, auto) next case False hence [simp]: "coeff p i = 0" using le_degree by blast show ?thesis by (subst sum.neutral, auto simp: h0) qed finally show ?thesis by simp qed have h'0: "[: h 0 :] = 0" using h0 by auto show "coeff ?mp i = coeff ?ep i" unfolding coeff_map_poly[of h, OF h0] unfolding eval_poly_as_sum[of "\a. [: h a :]", OF h'0] unfolding coeff_sum unfolding x_as_monom x_pow_n coeff_mult unfolding sum.swap[of _ _ "{..degree p}"] unfolding coeff_monom using 2 by auto qed lemma smult_as_map_poly: "smult a = map_poly ((*) a)" by (rule ext, rule poly_eqI, subst coeff_map_poly, auto) subsection \@{const map_poly} of Homomorphisms\ context zero_hom begin text \We will consider @{term hom} is always simpler than @{term "map_poly hom"}.\ lemma map_poly_hom_monom[simp]: "map_poly hom (monom a i) = monom (hom a) i" by(rule map_poly_monom, auto) lemma coeff_map_poly_hom[simp]: "coeff (map_poly hom p) i = hom (coeff p i)" by (rule coeff_map_poly, rule hom_zero) end locale map_poly_zero_hom = base: zero_hom begin sublocale zero_hom "map_poly hom" by (unfold_locales, auto) end text \@{const map_poly} preserves homomorphisms over addition.\ context comm_monoid_add_hom begin lemma map_poly_hom_add[hom_distribs]: "map_poly hom (p + q) = map_poly hom p + map_poly hom q" by (rule map_poly_add; simp add: hom_distribs) end locale map_poly_comm_monoid_add_hom = base: comm_monoid_add_hom begin sublocale comm_monoid_add_hom "map_poly hom" by (unfold_locales, auto simp:hom_distribs) end text \To preserve homomorphisms over multiplication, it demands commutative ring homomorphisms.\ context comm_semiring_hom begin lemma map_poly_pCons_hom[hom_distribs]: "map_poly hom (pCons a p) = pCons (hom a) (map_poly hom p)" unfolding map_poly_simps by auto lemma map_poly_hom_smult[hom_distribs]: "map_poly hom (smult c p) = smult (hom c) (map_poly hom p)" by (induct p, auto simp: hom_distribs) lemma poly_map_poly[simp]: "poly (map_poly hom p) (hom x) = hom (poly p x)" by (induct p; simp add: hom_distribs) end locale map_poly_comm_semiring_hom = base: comm_semiring_hom begin sublocale map_poly_comm_monoid_add_hom.. sublocale comm_semiring_hom "map_poly hom" proof show "map_poly hom 1 = 1" by simp fix p q show "map_poly hom (p * q) = map_poly hom p * map_poly hom q" by (induct p, auto simp: hom_distribs) qed end locale map_poly_comm_ring_hom = base: comm_ring_hom begin sublocale map_poly_comm_semiring_hom.. sublocale comm_ring_hom "map_poly hom".. end locale map_poly_idom_hom = base: idom_hom begin sublocale map_poly_comm_ring_hom.. sublocale idom_hom "map_poly hom".. end subsubsection \Injectivity\ locale map_poly_inj_zero_hom = base: inj_zero_hom begin sublocale inj_zero_hom "map_poly hom" proof (unfold_locales) fix p q :: "'a poly" assume "map_poly hom p = map_poly hom q" from cong[of "\p. coeff p _", OF refl this] show "p = q" by (auto intro: poly_eqI) qed simp end locale map_poly_inj_comm_monoid_add_hom = base: inj_comm_monoid_add_hom begin sublocale map_poly_comm_monoid_add_hom.. sublocale map_poly_inj_zero_hom.. sublocale inj_comm_monoid_add_hom "map_poly hom".. end locale map_poly_inj_comm_semiring_hom = base: inj_comm_semiring_hom begin sublocale map_poly_comm_semiring_hom.. sublocale map_poly_inj_zero_hom.. sublocale inj_comm_semiring_hom "map_poly hom".. end locale map_poly_inj_comm_ring_hom = base: inj_comm_ring_hom begin sublocale map_poly_inj_comm_semiring_hom.. sublocale inj_comm_ring_hom "map_poly hom".. end locale map_poly_inj_idom_hom = base: inj_idom_hom begin sublocale map_poly_inj_comm_ring_hom.. sublocale inj_idom_hom "map_poly hom".. end lemma degree_map_poly_le: "degree (map_poly f p) \ degree p" by(induct p;auto) lemma coeffs_map_poly: (* An exact variant *) assumes "f (lead_coeff p) = 0 \ p = 0" shows "coeffs (map_poly f p) = map f (coeffs p)" unfolding coeffs_map_poly using assms by (simp add:coeffs_def) lemma degree_map_poly: (* An exact variant *) assumes "f (lead_coeff p) = 0 \ p = 0" shows "degree (map_poly f p) = degree p" unfolding degree_eq_length_coeffs unfolding coeffs_map_poly[of f, OF assms] by simp context zero_hom_0 begin lemma degree_map_poly_hom[simp]: "degree (map_poly hom p) = degree p" by (rule degree_map_poly, auto) lemma coeffs_map_poly_hom[simp]: "coeffs (map_poly hom p) = map hom (coeffs p)" by (rule coeffs_map_poly, auto) lemma hom_lead_coeff[simp]: "lead_coeff (map_poly hom p) = hom (lead_coeff p)" by simp end context comm_semiring_hom begin interpretation map_poly_hom: map_poly_comm_semiring_hom.. lemma poly_map_poly_0[simp]: "poly (map_poly hom p) 0 = hom (poly p 0)" (is "?l = ?r") proof- have "?l = poly (map_poly hom p) (hom 0)" by auto then show ?thesis unfolding poly_map_poly. qed lemma poly_map_poly_1[simp]: "poly (map_poly hom p) 1 = hom (poly p 1)" (is "?l = ?r") proof- have "?l = poly (map_poly hom p) (hom 1)" by auto then show ?thesis unfolding poly_map_poly. qed lemma map_poly_hom_as_monom_sum: "(\j\degree p. monom (hom (coeff p j)) j) = map_poly hom p" proof - show ?thesis by (subst(6) poly_as_sum_of_monoms'[OF le_refl, symmetric], simp add: hom_distribs) qed lemma map_poly_pcompose[hom_distribs]: "map_poly hom (f \\<^sub>p g) = map_poly hom f \\<^sub>p map_poly hom g" by (induct f arbitrary: g; auto simp: hom_distribs) end context comm_semiring_hom begin lemma eval_poly_0[simp]: "eval_poly hom 0 x = 0" unfolding eval_poly_def by simp lemma eval_poly_monom: "eval_poly hom (monom a n) x = hom a * x ^ n" unfolding eval_poly_def unfolding map_poly_monom[of hom, OF hom_zero] using poly_monom. lemma poly_map_poly_eval_poly: "poly (map_poly hom p) = eval_poly hom p" unfolding eval_poly_def.. lemma map_poly_eval_poly: "map_poly hom p = eval_poly (\ a. [: hom a :]) p [:0,1:]" by (rule map_poly_eval_poly, simp) lemma degree_extension: assumes "degree p \ n" shows "(\i\degree p. x ^ i * hom (coeff p i)) = (\i\n. x ^ i * hom (coeff p i))" (is "?l = ?r") proof - let ?f = "\ i. x ^ i * hom (coeff p i)" define m where "m = n - degree p" have n: "n = degree p + m" unfolding m_def using assms by auto have "?r = (\ i \ degree p + m. ?f i)" unfolding n .. also have "\ = ?l + sum ?f {Suc (degree p) .. degree p + m}" by (subst sum.union_disjoint[symmetric], auto intro: sum.cong) also have "sum ?f {Suc (degree p) .. degree p + m} = 0" by (rule sum.neutral, auto simp: coeff_eq_0) finally show ?thesis by simp qed lemma eval_poly_add[simp]: "eval_poly hom (p + q) x = eval_poly hom p x + eval_poly hom q x" unfolding eval_poly_def hom_distribs.. lemma eval_poly_sum: "eval_poly hom (\k\A. p k) x = (\k\A. eval_poly hom (p k) x)" proof (induct A rule: infinite_finite_induct) case (insert a A) show ?case unfolding sum.insert[OF insert(1-2)] insert(3)[symmetric] by simp qed (auto simp: eval_poly_def) lemma eval_poly_poly: "eval_poly hom p (hom x) = hom (poly p x)" unfolding eval_poly_def by auto end context comm_ring_hom begin interpretation map_poly_hom: map_poly_comm_ring_hom.. lemma pseudo_divmod_main_hom: "pseudo_divmod_main (hom lc) (map_poly hom q) (map_poly hom r) (map_poly hom d) dr i = map_prod (map_poly hom) (map_poly hom) (pseudo_divmod_main lc q r d dr i)" proof- show ?thesis by (induct lc q r d dr i rule:pseudo_divmod_main.induct, auto simp: Let_def hom_distribs) qed end lemma(in inj_comm_ring_hom) pseudo_divmod_hom: "pseudo_divmod (map_poly hom p) (map_poly hom q) = map_prod (map_poly hom) (map_poly hom) (pseudo_divmod p q)" unfolding pseudo_divmod_def using pseudo_divmod_main_hom[of _ 0] by (cases "q = 0",auto) lemma(in inj_idom_hom) pseudo_mod_hom: "pseudo_mod (map_poly hom p) (map_poly hom q) = map_poly hom (pseudo_mod p q)" using pseudo_divmod_hom unfolding pseudo_mod_def by auto lemma(in idom_hom) map_poly_pderiv[hom_distribs]: "map_poly hom (pderiv p) = pderiv (map_poly hom p)" proof (induct p rule: pderiv.induct) case (1 a p) then show ?case unfolding pderiv.simps map_poly_pCons_hom by (cases "p = 0", auto simp: hom_distribs) qed context field_hom begin -lemma map_poly_pdivmod[hom_distribs]: - "map_prod (map_poly hom) (map_poly hom) (p div q, p mod q) = - (map_poly hom p div map_poly hom q, map_poly hom p mod map_poly hom q)" - (is "?l = ?r") +lemma dvd_map_poly_hom_imp_dvd: \map_poly hom x dvd map_poly hom y \ x dvd y\ + by (smt (verit, del_insts) degree_map_poly_hom hom_0 hom_div hom_lead_coeff hom_one hom_power map_poly_hom_smult map_poly_zero mod_eq_0_iff_dvd mod_poly_def pseudo_mod_hom) + +lemma map_poly_pdivmod [hom_distribs]: + \map_prod (map_poly hom) (map_poly hom) (p div q, p mod q) = + (map_poly hom p div map_poly hom q, map_poly hom p mod map_poly hom q)\ proof - - let ?mp = "map_poly hom" - interpret map_poly_hom: map_poly_idom_hom.. - obtain r s where dm: "(p div q, p mod q) = (r, s)" - by force - hence r: "r = p div q" and s: "s = p mod q" - by simp_all - from dm [folded pdivmod_pdivmodrel] have "eucl_rel_poly p q (r, s)" - by auto - from this[unfolded eucl_rel_poly_iff] - have eq: "p = r * q + s" and cond: "(if q = 0 then r = 0 else s = 0 \ degree s < degree q)" by auto - from arg_cong[OF eq, of ?mp, unfolded map_poly_add] - have eq: "?mp p = ?mp q * ?mp r + ?mp s" by (auto simp: hom_distribs) - from cond have cond: "(if ?mp q = 0 then ?mp r = 0 else ?mp s = 0 \ degree (?mp s) < degree (?mp q))" + let ?mp = \map_poly hom\ + interpret map_poly_hom: map_poly_idom_hom .. + have \(?mp p div ?mp q, ?mp p mod ?mp q) = (?mp (p div q), ?mp (p mod q))\ + proof (cases \?mp q\ \?mp (p div q)\ \?mp (p mod q)\ \?mp p\ rule: euclidean_relation_polyI) + case by0 + then show ?case + by simp + next + case divides + then have \q \ 0\ \q dvd p\ + by (auto dest: dvd_map_poly_hom_imp_dvd) + from \q dvd p\ obtain r where \p = q * r\ .. + with \q \ 0\ show ?case + by (simp add: map_poly_hom.hom_mult) + next + case euclidean_relation + with degree_mod_less_degree [of q p] show ?case + by (auto simp flip: map_poly_hom.hom_mult map_poly_hom_add) + qed + then show ?thesis by simp - from eq cond have "eucl_rel_poly (?mp p) (?mp q) (?mp r, ?mp s)" - unfolding eucl_rel_poly_iff by auto - from this[unfolded pdivmod_pdivmodrel] - show ?thesis unfolding dm prod.simps by simp qed lemma map_poly_div[hom_distribs]: "map_poly hom (p div q) = map_poly hom p div map_poly hom q" using map_poly_pdivmod[of p q] by simp lemma map_poly_mod[hom_distribs]: "map_poly hom (p mod q) = map_poly hom p mod map_poly hom q" using map_poly_pdivmod[of p q] by simp end locale field_hom' = field_hom hom for hom :: "'a :: {field_gcd} \ 'b :: {field_gcd}" begin lemma map_poly_normalize[hom_distribs]: "map_poly hom (normalize p) = normalize (map_poly hom p)" by (simp add: normalize_poly_def hom_distribs) lemma map_poly_gcd[hom_distribs]: "map_poly hom (gcd p q) = gcd (map_poly hom p) (map_poly hom q)" by (induct p q rule: eucl_induct) (simp_all add: map_poly_normalize ac_simps hom_distribs) end definition div_poly :: "'a :: euclidean_semiring \ 'a poly \ 'a poly" where "div_poly a p = map_poly (\ c. c div a) p" lemma smult_div_poly: assumes "\ c. c \ set (coeffs p) \ a dvd c" shows "smult a (div_poly a p) = p" unfolding smult_as_map_poly div_poly_def by (subst map_poly_map_poly, force, subst map_poly_idI, insert assms, auto) lemma coeff_div_poly: "coeff (div_poly a f) n = coeff f n div a" unfolding div_poly_def by (rule coeff_map_poly, auto) locale map_poly_inj_idom_divide_hom = base: inj_idom_divide_hom begin sublocale map_poly_idom_hom .. sublocale map_poly_inj_zero_hom .. sublocale inj_idom_hom "map_poly hom" .. lemma divide_poly_main_hom: defines "hh \ map_poly hom" shows "hh (divide_poly_main lc f g h i j) = divide_poly_main (hom lc) (hh f) (hh g) (hh h) i j" unfolding hh_def proof (induct j arbitrary: lc f g h i) case (Suc j lc f g h i) let ?h = "map_poly hom" show ?case unfolding divide_poly_main.simps Let_def unfolding base.coeff_map_poly_hom base.hom_div[symmetric] base.hom_mult[symmetric] base.eq_iff if_distrib[of ?h] hom_zero by (rule if_cong[OF refl _ refl], subst Suc, simp add: hom_minus hom_add hom_mult) qed simp sublocale inj_idom_divide_hom "map_poly hom" proof fix f g :: "'a poly" let ?h = "map_poly hom" show "?h (f div g) = (?h f) div (?h g)" unfolding divide_poly_def if_distrib[of ?h] divide_poly_main_hom by simp qed lemma order_hom: "order (hom x) (map_poly hom f) = order x f" unfolding Polynomial.order_def unfolding hom_dvd_iff[symmetric] unfolding hom_power by (simp add: base.hom_uminus) end subsection \Example Interpretations\ abbreviation "of_int_poly \ map_poly of_int" interpretation of_int_poly_hom: map_poly_comm_semiring_hom of_int.. interpretation of_int_poly_hom: map_poly_comm_ring_hom of_int.. interpretation of_int_poly_hom: map_poly_idom_hom of_int.. interpretation of_int_poly_hom: map_poly_inj_comm_ring_hom "of_int :: int \ 'a :: {comm_ring_1,ring_char_0}" .. interpretation of_int_poly_hom: map_poly_inj_idom_hom "of_int :: int \ 'a :: {idom,ring_char_0}" .. text \The following operations are homomorphic w.r.t. only @{class monoid_add}.\ interpretation pCons_0_hom: injective "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: zero_hom_0 "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: inj_comm_monoid_add_hom "pCons 0" by (unfold_locales, auto) interpretation pCons_0_hom: inj_ab_group_add_hom "pCons 0" by (unfold_locales, auto) interpretation monom_hom: injective "\x. monom x d" by (unfold_locales, auto) interpretation monom_hom: inj_monoid_add_hom "\x. monom x d" by (unfold_locales, auto simp: add_monom) interpretation monom_hom: inj_comm_monoid_add_hom "\x. monom x d".. end diff --git a/thys/Subresultants/Subresultant.thy b/thys/Subresultants/Subresultant.thy --- a/thys/Subresultants/Subresultant.thy +++ b/thys/Subresultants/Subresultant.thy @@ -1,2778 +1,2778 @@ section \Subresultants and the subresultant PRS\ text \This theory contains most of the soundness proofs of the subresultant PRS algorithm, where we closely follow the papers of Brown \cite{Brown} and Brown and Traub \cite{BrownTraub}. This is in contrast to a similar Coq formalization of Mahboubi \cite{Mahboubi06} which is based on polynomial determinants. Whereas the current file only contains an algorithm to compute the resultant of two polynomials efficiently, there is another theory ``Subresultant-Gcd'' which also contains the algorithm to compute the GCD of two polynomials via the subresultant algorithm. In both algorithms we integrate Lazard's optimization in the dichotomous version, but not the second optmization described by Ducos \cite{Ducos}.\ theory Subresultant imports Resultant_Prelim Dichotomous_Lazard Binary_Exponentiation More_Homomorphisms Coeff_Int begin subsection \Algorithm\ locale div_exp_param = fixes div_exp :: "'a :: idom_divide \ 'a \ nat \ 'a" begin partial_function(tailrec) subresultant_prs_main where "subresultant_prs_main f g c = (let m = degree f; n = degree g; lf = lead_coeff f; lg = lead_coeff g; \ = m - n; d = div_exp lg c \; h = pseudo_mod f g in if h = 0 then (g,d) else subresultant_prs_main g (sdiv_poly h ((-1) ^ (\ + 1) * lf * (c ^ \))) d)" definition subresultant_prs where "subresultant_prs f g = (let h = pseudo_mod f g; \ = (degree f - degree g); d = lead_coeff g ^ \ in if h = 0 then (g,d) else subresultant_prs_main g ((- 1) ^ (\ + 1) * h) d)" definition resultant_impl_main where "resultant_impl_main G1 G2 = (if G2 = 0 then (if degree G1 = 0 then 1 else 0) else case subresultant_prs G1 G2 of (Gk,hk) \ (if degree Gk = 0 then hk else 0))" definition resultant_impl where "resultant_impl f g = (if length (coeffs f) \ length (coeffs g) then resultant_impl_main f g else let res = resultant_impl_main g f in if even (degree f) \ even (degree g) then res else - res)" end locale div_exp_sound = div_exp_param + assumes div_exp: "\ x y n. (to_fract x)^n / (to_fract y)^(n-1) \ range to_fract \ to_fract (div_exp x y n) = (to_fract x)^n / (to_fract y)^(n-1)" definition basic_div_exp :: "'a :: idom_divide \ 'a \ nat \ 'a" where "basic_div_exp x y n = x^n div y^(n-1)" text \We have an instance for arbitrary integral domains.\ lemma basic_div_exp: "div_exp_sound basic_div_exp" by (unfold_locales, unfold basic_div_exp_def, rule sym, rule div_divide_to_fract, auto simp: hom_distribs) text \Lazard's optimization is only proven for factorial rings.\ lemma dichotomous_Lazard: "div_exp_sound (dichotomous_Lazard :: 'a :: factorial_ring_gcd \ _)" by (unfold_locales, rule dichotomous_Lazard) subsection \Soundness Proof for @{term "div_exp_param.resultant_impl div_exp = resultant"}\ abbreviation pdivmod :: "'a::field poly \ 'a poly \ 'a poly \ 'a poly" where "pdivmod p q \ (p div q, p mod q)" lemma even_sum_list: assumes "\ x. x \ set xs \ even (f x) = even (g x)" shows "even (sum_list (map f xs)) = even (sum_list (map g xs))" using assms by (induct xs, auto) lemma for_all_Suc: "P i \ (\ j \ Suc i. P j) = (\ j \ i. P j)" for P by (metis (full_types) Suc_le_eq less_le) (* part on pseudo_divmod *) lemma pseudo_mod_left_0[simp]: "pseudo_mod 0 x = 0" unfolding pseudo_mod_def pseudo_divmod_def by (cases "x = 0"; cases "length (coeffs x)", auto) lemma pseudo_mod_right_0[simp]: "pseudo_mod x 0 = x" unfolding pseudo_mod_def pseudo_divmod_def by simp lemma snd_pseudo_divmod_main_cong: assumes "a1 = b1" "a3 = b3" "a4 = b4" "a5 = b5" "a6 = b6" (* note that a2 = b2 is not required! *) shows "snd (pseudo_divmod_main a1 a2 a3 a4 a5 a6) = snd (pseudo_divmod_main b1 b2 b3 b4 b5 b6)" using assms snd_pseudo_divmod_main by metis lemma snd_pseudo_mod_smult_invar_right: shows "(snd (pseudo_divmod_main (x * lc) q r (smult x d) dr n)) = snd (pseudo_divmod_main lc q' (smult (x^n) r) d dr n)" proof(induct n arbitrary: q q' r dr) case (Suc n) let ?q = "smult (x * lc) q + monom (coeff r dr) n" let ?r = "smult (x * lc) r - (smult x (monom (coeff r dr) n * d))" let ?dr = "dr - 1" let ?rec_lhs = "pseudo_divmod_main (x * lc) ?q ?r (smult x d) ?dr n" let ?rec_rhs = "pseudo_divmod_main lc q' (smult (x^n) ?r) d ?dr n" have [simp]: "\ n. x ^ n * (x * lc) = lc * (x * x ^ n)" "\ n c. x ^ n * (x * c) = x * x ^ n * c" "\ n. x * x ^ n * lc = lc * (x * x ^ n)" by (auto simp: ac_simps) have "snd (pseudo_divmod_main (x * lc) q r (smult x d) dr (Suc n)) = snd ?rec_lhs" by (auto simp:Let_def) also have "\ = snd ?rec_rhs" using Suc by auto also have "\ = snd (pseudo_divmod_main lc q' (smult (x^Suc n) r) d dr (Suc n))" unfolding pseudo_divmod_main.simps Let_def proof(rule snd_pseudo_divmod_main_cong,goal_cases) case 2 show ?case by (auto simp:smult_add_right smult_diff_right smult_monom smult_monom_mult) qed auto finally show ?case by auto qed auto lemma snd_pseudo_mod_smult_invar_left: shows "snd (pseudo_divmod_main lc q (smult x r) d dr n) = smult x (snd (pseudo_divmod_main lc q' r d dr n))" proof(induct n arbitrary:x lc q q' r d dr) case (Suc n) have sm:"smult lc (smult x r) - monom (coeff (smult x r) dr) n * d =smult x (smult lc r - monom (coeff r dr) n * d) " by (auto simp:smult_diff_right smult_monom smult_monom_mult mult.commute[of lc x]) let ?q' = "smult lc q' + monom (coeff r dr) n" show ?case unfolding pseudo_divmod_main.simps Let_def Suc(1)[of lc _ _ _ _ _ ?q'] sm by auto qed auto lemma snd_pseudo_mod_smult_left[simp]: shows "snd (pseudo_divmod (smult (x::'a::idom) p) q) = (smult x (snd (pseudo_divmod p q)))" unfolding pseudo_divmod_def by (auto simp:snd_pseudo_mod_smult_invar_left[of _ _ _ _ _ _ _ 0] Polynomial.coeffs_smult) lemma pseudo_mod_smult_right: assumes "(x::'a::idom)\0" "q\0" shows "(pseudo_mod p (smult (x::'a::idom) q)) = (smult (x^(Suc (length (coeffs p)) - length (coeffs q))) (pseudo_mod p q))" unfolding pseudo_divmod_def pseudo_mod_def by (auto simp:snd_pseudo_mod_smult_invar_right[of _ _ _ _ _ _ _ 0] snd_pseudo_mod_smult_invar_left[of _ _ _ _ _ _ _ 0] Polynomial.coeffs_smult assms) lemma pseudo_mod_zero[simp]: "pseudo_mod 0 f = (0::'a :: {idom} poly)" "pseudo_mod f 0 = f" unfolding pseudo_mod_def snd_pseudo_mod_smult_left[of 0 _ f,simplified] unfolding pseudo_divmod_def by auto (* part on prod_list *) lemma prod_combine: assumes "j \ i" shows "f i * (\l\[j.. i. (-1)^(f i)) xs) = (-1)^(sum_list (map f xs))" by (induct xs, auto simp: power_add) lemma minus_1_power_even: "(- (1 :: 'b :: comm_ring_1))^ k = (if even k then 1 else (-1))" by auto lemma minus_1_even_eqI: assumes "even k = even l" shows "(- (1 :: 'b :: comm_ring_1))^k = (- 1)^l" unfolding minus_1_power_even assms by auto lemma (in comm_monoid_mult) prod_list_multf: "(\x\xs. f x * g x) = prod_list (map f xs) * prod_list (map g xs)" by (induct xs) (simp_all add: algebra_simps) lemma inverse_prod_list: "inverse (prod_list xs) = prod_list (map inverse (xs :: 'a :: field list))" by (induct xs, auto) (* part on pow_int, i.e., exponentiation with integer exponent *) definition pow_int :: "'a :: field \ int \ 'a" where "pow_int x e = (if e < 0 then 1 / (x ^ (nat (-e))) else x ^ (nat e))" lemma pow_int_0[simp]: "pow_int x 0 = 1" unfolding pow_int_def by auto lemma pow_int_1[simp]: "pow_int x 1 = x" unfolding pow_int_def by auto lemma exp_pow_int: "x ^ n = pow_int x n" unfolding pow_int_def by auto lemma pow_int_add: assumes x: "x \ 0" shows "pow_int x (a + b) = pow_int x a * pow_int x b" proof - have *: "\ a + b < 0 \ a < 0 \ nat b = nat (a + b) + nat (-a)" "\ a + b < 0 \ b < 0 \ nat a = nat (a + b) + nat (-b)" "a + b < 0 \ \ a < 0 \ nat (-b) = nat a + nat (-a -b) " "a + b < 0 \ \ b < 0 \ nat (-a) = nat b + nat (-a -b) " by auto have pow_eq: "l = m \ (x ^ l = x ^ m)" for l m by auto from x show ?thesis unfolding pow_int_def by (auto split: if_splits simp: power_add[symmetric] simp: * intro!: pow_eq, auto simp: power_add) qed lemma pow_int_mult: "pow_int (x * y) a = pow_int x a * pow_int y a" unfolding pow_int_def by (cases "a < 0", auto simp: power_mult_distrib) lemma pow_int_base_1[simp]: "pow_int 1 a = 1" unfolding pow_int_def by (cases "a < 0", auto) lemma pow_int_divide: "a / pow_int x b = a * pow_int x (-b)" unfolding pow_int_def by (cases b rule: linorder_cases[of _ 0], auto) lemma divide_prod_assoc: "x / (y * z :: 'a :: field) = x / y / z" by (simp add: field_simps) lemma minus_1_inverse_pow[simp]: "x / (-1)^n = (x :: 'a :: field) * (-1)^n" by (simp add: minus_1_power_even) (* part on subresultants *) definition subresultant_mat :: "nat \ 'a :: comm_ring_1 poly \ 'a poly \ 'a poly mat" where "subresultant_mat J F G = (let dg = degree G; df = degree F; f = coeff_int F; g = coeff_int G; n = (df - J) + (dg - J) in mat n n (\ (i,j). if j < dg - J then if i = n - 1 then monom 1 (dg - J - 1 - j) * F else [: f (df - int i + int j) :] else let jj = j - (dg - J) in if i = n - 1 then monom 1 (df - J - 1 - jj) * G else [: g (dg - int i + int jj) :]))" lemma subresultant_mat_dim[simp]: fixes j p q defines "S \ subresultant_mat j p q" shows "dim_row S = (degree p - j) + (degree q - j)" and "dim_col S = (degree p - j) + (degree q - j)" unfolding S_def subresultant_mat_def Let_def by auto definition subresultant'_mat :: "nat \ nat \ 'a :: comm_ring_1 poly \ 'a poly \ 'a mat" where "subresultant'_mat J l F G = (let \ = degree G; \ = degree F; f = coeff_int F; g = coeff_int G; n = (\ - J) + (\ - J) in mat n n (\ (i,j). if j < \ - J then if i = n - 1 then (f (l - int (\ - J - 1) + int j)) else (f (\ - int i + int j)) else let jj = j - (\ - J) in if i = n - 1 then (g (l - int (\ - J - 1) + int jj)) else (g (\ - int i + int jj))))" lemma subresultant_index_mat: fixes F G assumes i: "i < (degree F - J) + (degree G - J)" and j: "j < (degree F - J) + (degree G - J)" shows "subresultant_mat J F G $$ (i,j) = (if j < degree G - J then if i = (degree F - J) + (degree G - J) - 1 then monom 1 (degree G - J - 1 - j) * F else ([: coeff_int F ( degree F - int i + int j) :]) else let jj = j - (degree G - J) in if i = (degree F - J) + (degree G - J) - 1 then monom 1 ( degree F - J - 1 - jj) * G else ([: coeff_int G (degree G - int i + int jj) :]))" unfolding subresultant_mat_def Let_def unfolding index_mat(1)[OF i j] split by auto definition subresultant :: "nat \ 'a :: comm_ring_1 poly \ 'a poly \ 'a poly" where "subresultant J F G = det (subresultant_mat J F G)" lemma subresultant_smult_left: assumes "(c :: 'a :: {comm_ring_1, semiring_no_zero_divisors}) \ 0" shows "subresultant J (smult c f) g = smult (c ^ (degree g - J)) (subresultant J f g)" proof - let ?df = "degree f" let ?dg = "degree g" let ?n = "(?df - J) + (?dg - J)" let ?m = "?dg - J" let ?M = "mat ?n ?n (\ (i,j). if i = j then if i < ?m then [:c:] else 1 else 0)" from \c \ 0\ have deg: "degree (smult c f) = ?df" by simp let ?S = "subresultant_mat J f g" let ?cS = "subresultant_mat J (smult c f) g" have dim: "dim_row ?S = ?n" "dim_col ?S = ?n" "dim_row ?cS = ?n" "dim_col ?cS = ?n" using deg by auto hence C: "?S \ carrier_mat ?n ?n" "?cS \ carrier_mat ?n ?n" "?M \ carrier_mat ?n ?n" by auto have dim': "dim_row (?S * ?M) = ?n" "dim_col (?S * ?M) = ?n" using dim (1,2) by simp_all define S where "S = ?S" have "?cS = ?S * ?M" proof (rule eq_matI, unfold dim' dim) fix i j assume ij: "i < ?n" "j < ?n" have "(?S * ?M) $$ (i,j) = row ?S i \ col ?M j" by (rule index_mult_mat, insert ij dim, auto) also have "\ = (\k = 0.. = (\k = 0.. = S $$ (i,j) * ?M $$ (j,j) + sum (\ k. S $$ (i,k) * ?M $$ (k,j)) ({0.. = S $$ (i,j) * ?M $$ (j,j)" by (subst sum.neutral, insert ij, auto) also have "\ = ?cS $$ (i,j)" unfolding subresultant_index_mat[OF ij] S_def by (subst subresultant_index_mat, unfold deg, insert ij, auto) finally show "?cS $$ (i,j) = (?S * ?M) $$ (i,j)" by simp qed auto from arg_cong[OF this, of det] det_mult[OF C(1) C(3)] have "subresultant J (smult c f) g = subresultant J f g * det ?M" unfolding subresultant_def by auto also have "det ?M = [:c ^ ?m :]" proof (subst det_upper_triangular[OF _ C(3)]) show "upper_triangular ?M" by (rule upper_triangularI, auto) have "prod_list (diag_mat ?M) = (\k = 0.. = (\k = 0..k = ?m..k = 0..k = 0..k = 0..k = ?m..k = ?m.. ((if j < ?k then j + ?n else j - ?k) < ?n) = (\ (j < ?k))" for j by auto have "subresultant J f g = det ?A" unfolding subresultant_def by simp also have "\ = (-1)^(?k * ?n) * det (mat (?k + ?n) (?k + ?n) (\ (i,j). ?A $$ (i,(if j < ?k then j + ?n else j - ?k))))" (is "_ = _ * det ?B") by (rule det_swap_cols, auto simp: subresultant_mat_def Let_def) also have "?B = subresultant_mat J g f" unfolding subresultant_mat_def Let_def by (rule eq_matI, unfold dim_row_mat dim_col_mat nk index_mat split, subst index_mat, (auto)[2], unfold split, subst change, force, unfold if_conn, rule if_cong[OF refl if_cong if_cong], auto) also have "det \ = subresultant J g f" unfolding subresultant_def .. also have "(-1)^(?k * ?n) * \ = [: (-1)^(?k * ?n) :] * \ " by (unfold hom_distribs, simp) also have "\ = smult ((-1)^(?k * ?n)) (subresultant J g f)" by simp finally show ?thesis . qed lemma subresultant_smult_right:assumes "(c :: 'a :: {comm_ring_1, semiring_no_zero_divisors}) \ 0" shows "subresultant J f (smult c g) = smult (c ^ (degree f - J)) (subresultant J f g)" unfolding subresultant_swap[of _ f] subresultant_smult_left[OF assms] degree_smult_eq using assms by (simp add: ac_simps) lemma coeff_subresultant: "coeff (subresultant J F G) l = (if degree F - J + (degree G - J) = 0 \ l \ 0 then 0 else det (subresultant'_mat J l F G))" proof (cases "degree F - J + (degree G - J) = 0") case True show ?thesis unfolding subresultant_def subresultant_mat_def subresultant'_mat_def Let_def True by simp next case False let ?n = "degree F - J + (degree G - J)" define n where "n = ?n" from False have n: "n \ 0" unfolding n_def by auto hence id: "{0..i = 0..i = 0 ..< (n - 1). coeff_int (M $$ (i, p i)) 0) * coeff_int (M $$ (n - 1, p (n - 1))) l" unfolding id proof (rule coeff_int_prod_const, (auto)[2]) fix i assume "i \ {0 ..< n - 1}" with p have i: "i \ n - 1" and "i < n" "p i < n" by (auto simp: n_def) note id = subresultant_index_mat[OF this(2-3)[unfolded n_def], folded M_def n_def] show "degree (M $$ (i, p i)) = 0" unfolding id Let_def using i by (simp split: if_splits) qed also have "(\i = 0 ..< (n - 1). coeff_int (M $$ (i, p i)) 0) = (\i = 0 ..< (n - 1). L $$ (i, p i))" proof (rule prod.cong[OF refl]) fix i assume "i \ {0 ..< n - 1}" with p have i: "i \ n - 1" and ii: "i < n" "p i < n" by (auto simp: n_def) note id = subresultant_index_mat[OF this(2-3)[unfolded n_def], folded M_def n_def] note id' = L_def[unfolded subresultant'_mat_def Let_def, folded n_def] index_mat[OF ii] show "coeff_int (M $$ (i, p i)) 0 = L $$ (i, p i)" unfolding id id' split using i proof (simp add: if_splits Let_def) qed qed also have "coeff_int (M $$ (n - 1, p (n - 1))) l = (if p (n - 1) < degree G - J then coeff_int (monom 1 (degree G - J - 1 - p (n - 1)) * F) l else coeff_int (monom 1 (degree F - J - 1 - (p (n - 1) - (degree G - J))) * G) l)" using subresultant_index_mat[OF n1[unfolded n_def], folded M_def n_def, unfolded idn if_True Let_def] by simp also have "\ = (if p (n - 1) < degree G - J then coeff_int F (int l - int (degree G - J - 1 - p (n - 1))) else coeff_int G (int l - int (degree F - J - 1 - (p (n - 1) - (degree G - J)))))" unfolding coeff_int_monom_mult by simp also have "\ = (if p (n - 1) < degree G - J then coeff_int F (int l - int (degree G - J - 1) + p (n - 1)) else coeff_int G (int l - int (degree F - J - 1) + (p (n - 1) - (degree G - J))))" proof (cases "p (n - 1) < degree G - J") case True hence "int (degree G - J - 1 - p (n - 1)) = int (degree G - J - 1) - p (n - 1)" by simp hence id: "int l - int (degree G - J - 1 - p (n - 1)) = int l - int (degree G - J - 1) + p (n - 1)" by simp show ?thesis using True unfolding id by simp next case False from n1 False have "degree F - J - 1 \ p (n - 1) - (degree G - J)" unfolding n_def by linarith hence "int (degree F - J - 1 - (p (n - 1) - (degree G - J))) = int (degree F - J - 1) - (p (n - 1) - (degree G - J))" by linarith hence id: "int l - int (degree F - J - 1 - (p (n - 1) - (degree G - J))) = int l - int (degree F - J - 1) + (p (n - 1) - (degree G - J))" by simp show ?thesis unfolding id using False by simp qed also have "\ = L $$ (n - 1, p (n - 1))" unfolding L_def subresultant'_mat_def Let_def n_def[symmetric] using n1 by simp also have "(\i = 0.. = (\i = 0..i = 0..i = 0..p\{p. p permutes {0..i = 0.. = (\p\{p. p permutes {0..i = 0.. = det L" proof - have id: "dim_row (subresultant'_mat J l F G) = n" "dim_col (subresultant'_mat J l F G) = n" unfolding subresultant'_mat_def Let_def n_def by auto show ?thesis unfolding det_def L_def id by simp qed finally show ?thesis unfolding L_def coeff_int_def using False by auto qed lemma subresultant'_zero_ge: assumes "(degree f - J) + (degree g - J) \ 0" and "k \ degree f + (degree g - J)" shows "det (subresultant'_mat J k f g) = 0" (* last row is zero *) proof - obtain dg where dg: "degree g - J = dg" by simp obtain df where df: "degree f - J = df" by simp obtain ddf where ddf: "degree f = ddf" by simp note * = assms(2)[unfolded ddf dg] assms(1) define M where "M = (\ i j. if j < dg then coeff_int f (degree f - int i + int j) else coeff_int g (degree g - int i + int (j - dg)))" let ?M = "subresultant'_mat J k f g" have M: "det ?M = det (mat (df + dg) (df + dg) (\(i, j). if i = df + dg - 1 then if j < dg then coeff_int f (int k - int (dg - 1) + int j) else coeff_int g (int k - int (df - 1) + int (j - dg)) else M i j))" (is "_ = det ?N") unfolding subresultant'_mat_def Let_def M_def by (rule arg_cong[of _ _ det], rule eq_matI, auto simp: df dg) also have "?N = mat (df + dg) (df + dg) (\(i, j). if i = df + dg - 1 then 0 else M i j)" by (rule cong_mat[OF refl refl], unfold split, rule if_cong[OF refl _ refl], auto simp add: coeff_int_def df dg ddf intro!: coeff_eq_0, insert *(1), unfold ddf[symmetric] dg[symmetric] df[symmetric], linarith+) also have "\ = mat\<^sub>r (df + dg) (df + dg) (\i. if i = df + dg - 1 then 0\<^sub>v (df + dg) else vec (df + dg) (\ j. M i j))" by (rule eq_matI, auto) also have "det \ = 0" by (rule det_row_0, insert *, auto simp: df[symmetric] dg[symmetric] ddf[symmetric]) finally show ?thesis . qed lemma subresultant'_zero_lt: assumes J: "J \ degree f" "J \ degree g" "J < k" and k: "k < degree f + (degree g - J)" shows "det (subresultant'_mat J k f g) = 0" (* last row is identical to last row - k *) proof - obtain dg where dg: "dg = degree g - J" by simp obtain df where df: "df = degree f - J" by simp note * = assms[folded df dg] define M where "M = (\ i j. if j < dg then coeff_int f (degree f - int i + int j) else coeff_int g (degree g - int i + int (j - dg)))" define N where "N = (\ j. if j < dg then coeff_int f (int k - int (dg - 1) + int j) else coeff_int g (int k - int (df - 1) + int (j - dg)))" let ?M = "subresultant'_mat J k f g" have M: "?M = mat (df + dg) (df + dg) (\(i, j). if i = df + dg - 1 then N j else M i j)" unfolding subresultant'_mat_def Let_def by (rule eq_matI, auto simp: df dg M_def N_def) also have "\ = mat (df + dg) (df + dg) (\(i, j). if i = df + dg - 1 then N j else if i = degree f + dg - 1 - k then N j else M i j)" (is "_ = ?N") unfolding N_def by (rule cong_mat[OF refl refl], unfold split, rule if_cong[OF refl refl], unfold M_def N_def, insert J k, auto simp: df dg intro!: arg_cong[of _ _ "coeff_int _"]) finally have id: "?M = ?N" . have deg: "degree f + dg - 1 - k < df + dg" "df + dg - 1 < df + dg" using k J unfolding df dg by auto have id: "row ?M (degree f + dg - 1 - k) = row ?M (df + dg - 1)" unfolding arg_cong[OF id, of row] by (rule eq_vecI, insert deg, auto) show ?thesis by (rule det_identical_rows[OF _ _ _ _ id, of "df + dg"], insert deg assms, auto simp: subresultant'_mat_def Let_def df dg) qed lemma subresultant'_mat_sylvester_mat: "transpose_mat (subresultant'_mat 0 0 f g) = sylvester_mat f g" proof - obtain dg where dg: "degree g = dg" by simp obtain df where df: "degree f = df" by simp let ?M = "transpose_mat (subresultant'_mat 0 0 f g)" let ?n = "degree f + degree g" have dim: "dim_row ?M = ?n" "dim_col ?M = ?n" by (auto simp: subresultant'_mat_def Let_def) show ?thesis proof (rule eq_matI, unfold sylvester_mat_dim dim df dg, goal_cases) case ij: (1 i j) have "?M $$ (i,j) = (if i < dg then if j = df + dg - 1 then coeff_int f (- int (dg - 1) + int i) else coeff_int f (int df - int j + int i) else if j = df + dg - 1 then coeff_int g (- int (df - 1) + int (i - dg)) else coeff_int g (int dg - int j + int (i - dg)))" using ij unfolding subresultant'_mat_def Let_def by (simp add: if_splits df dg) also have "\ = (if i < dg then coeff_int f (int df - int j + int i) else coeff_int g (int dg - int j + int (i - dg)))" proof - have cong: "(b \ x = z) \ (\ b \ y = z) \ (if b then coeff_int f x else coeff_int f y) = coeff_int f z" for b x y z and f :: "'a poly" by auto show ?thesis by (rule if_cong[OF refl cong cong], insert ij, auto) qed also have "\ = sylvester_mat f g $$ (i,j)" proof - have *: "i \ j \ j - i \ df \ nat (int df - int j + int i) = df - (j - i)" for j i df by simp show ?thesis unfolding sylvester_index_mat[OF ij[folded df dg]] df dg proof (rule if_cong[OF refl]) assume i: "i < dg" have "int df - int j + int i < 0 \ \ j - i \ df" by auto thus "coeff_int f (int df - int j + int i) = (if i \ j \ j - i \ df then coeff f (df + i - j) else 0)" using i ij by (simp add: coeff_int_def *, intro impI coeff_eq_0[of f, unfolded df], linarith) next assume i: "\ i < dg" hence **: "i - dg \ j \ dg - (j + dg - i) = i - j" using ij by linarith have "int dg - int j + int (i - dg) < 0 \ \ j \ i" by auto thus "coeff_int g (int dg - int j + int (i - dg)) = (if i - dg \ j \ j \ i then coeff g (i - j) else 0)" using ij i by (simp add: coeff_int_def * **, intro impI coeff_eq_0[of g, unfolded dg], linarith) qed qed finally show ?case . qed auto qed lemma coeff_subresultant_0_0_resultant: "coeff (subresultant 0 f g) 0 = resultant f g" proof - let ?M = "transpose_mat (subresultant'_mat 0 0 f g)" have "det (subresultant'_mat 0 0 f g) = det ?M" by (subst det_transpose, auto simp: subresultant'_mat_def Let_def) also have "?M = sylvester_mat f g" by (rule subresultant'_mat_sylvester_mat) finally show ?thesis by (simp add: coeff_subresultant resultant_def) qed lemma subresultant_zero_ge: assumes "k \ degree f + (degree g - J)" and "(degree f - J) + (degree g - J) \ 0" shows "coeff (subresultant J f g) k = 0" unfolding coeff_subresultant by (subst subresultant'_zero_ge[OF assms(2,1)], simp) lemma subresultant_zero_lt: assumes "k < degree f + (degree g - J)" and "J \ degree f" "J \ degree g" "J < k" shows "coeff (subresultant J f g) k = 0" unfolding coeff_subresultant by (subst subresultant'_zero_lt[OF assms(2,3,4,1)], simp) lemma subresultant_resultant: "subresultant 0 f g = [: resultant f g :]" proof (cases "degree f + degree g = 0") case True thus ?thesis unfolding subresultant_def subresultant_mat_def resultant_def Let_def sylvester_mat_def sylvester_mat_sub_def by simp next case 0: False show ?thesis proof (rule poly_eqI) fix k show "coeff (subresultant 0 f g) k = coeff [:resultant f g:] k" proof (cases "k = 0") case True thus ?thesis using coeff_subresultant_0_0_resultant[of f g] by auto next case False hence "0 < k \ k < degree f + degree g \ k \ degree f + degree g" by auto thus ?thesis using subresultant_zero_ge[of f g 0 k] 0 subresultant_zero_lt[of k f g 0] 0 False by (cases k, auto) qed qed qed lemma (in inj_comm_ring_hom) subresultant_hom: "map_poly hom (subresultant J f g) = subresultant J (map_poly hom f) (map_poly hom g)" proof - note d = subresultant_mat_def Let_def interpret p: map_poly_inj_comm_ring_hom hom.. show ?thesis unfolding subresultant_def unfolding p.hom_det[symmetric] proof (rule arg_cong[of _ _ det]) show "p.mat_hom (subresultant_mat J f g) = subresultant_mat J (map_poly hom f) (map_poly hom g)" proof (rule eq_matI, goal_cases) case (1 i j) hence ij: "i < degree f - J + (degree g - J)" "j < degree f - J + (degree g - J)" unfolding d degree_map_poly by auto show ?case by (auto simp add: coeff_int_def d map_mat_def index_mat(1)[OF ij] hom_distribs) qed (auto simp: d) qed qed text \We now derive properties of the resultant via the connection to subresultants.\ lemma resultant_smult_left: assumes "(c :: 'a :: idom) \ 0" shows "resultant (smult c f) g = c ^ degree g * resultant f g" unfolding coeff_subresultant_0_0_resultant[symmetric] subresultant_smult_left[OF assms] coeff_smult by simp lemma resultant_smult_right: assumes "(c :: 'a :: idom) \ 0" shows "resultant f (smult c g) = c ^ degree f * resultant f g" unfolding coeff_subresultant_0_0_resultant[symmetric] subresultant_smult_right[OF assms] coeff_smult by simp lemma resultant_swap: "resultant f g = (-1)^(degree f * degree g) * (resultant g f)" unfolding coeff_subresultant_0_0_resultant[symmetric] unfolding arg_cong[OF subresultant_swap[of 0 f g], of "\ x. coeff x 0"] coeff_smult by simp text \The following equations are taken from Brown-Traub ``On Euclid's Algorithm and the Theory of Subresultant'' (BT)\ lemma fixes F B G H :: "'a :: idom poly" and J :: nat defines df: "df \ degree F" and dg: "dg \ degree G" and dh: "dh \ degree H" and db: "db \ degree B" defines n: "n \ (df - J) + (dg - J)" and f: "f \ coeff_int F" and b: "b \ coeff_int B" and g: "g \ coeff_int G" and h: "h \ coeff_int H" assumes FGH: "F + B * G = H" and dfg: "df \ dg" and choice: "dg > dh \ H = 0 \ F \ 0 \ G \ 0" shows BT_eq_18: "subresultant J F G = smult ((-1)^((df - J) * (dg - J))) (det (mat n n (\ (i,j). if j < df - J then if i = n - 1 then monom 1 ((df - J) - 1 - j) * G else [:g (int dg - int i + int j):] else if i = n - 1 then monom 1 ((dg - J) - 1 - (j - (df - J))) * H else [:h (int df - int i + int (j - (df - J))):])))" (is "_ = smult ?m1 ?right") and BT_eq_19: "dh \ J \ J < dg \ subresultant J F G = smult ( (-1)^((df - J) * (dg - J)) * lead_coeff G ^ (df - J) * coeff H J ^ (dg - J - 1)) H" (is "_ \ _ \ _ = smult (_ * ?G * ?H) H") and BT_lemma_1_12: "J < dh \ subresultant J F G = smult ( (-1)^((df - J) * (dg - J)) * lead_coeff G ^ (df - dh)) (subresultant J G H)" and BT_lemma_1_13': "J = dh \ dg > dh \ H \ 0 \ subresultant dh F G = smult ( (-1)^((df - dh) * (dg - dh)) * lead_coeff G ^ (df - dh) * lead_coeff H ^ (dg - dh - 1)) H" and BT_lemma_1_14: "dh < J \ J < dg - 1 \ subresultant J F G = 0" and BT_lemma_1_15': "J = dg - 1 \ dg > dh \ H \ 0 \ subresultant (dg - 1) F G = smult ( (-1)^(df - dg + 1) * lead_coeff G ^ (df - dg + 1)) H" proof - define dfj where "dfj = df - J" define dgj where "dgj = dg - J" note d = df dg dh db have F0: "F \ 0" using dfg choice df by auto have G0: "G \ 0" using choice dg by auto have dgh: "dg \ dh" using choice unfolding dh by auto have B0: "B \ 0" using FGH dfg dgh choice F0 G0 unfolding d by auto have dfh: "df \ dh" using dfg dgh by auto have "df = degree (B * G)" proof (cases "H = 0") case False with choice dfg have dfh: "df > dh" by auto show ?thesis using dfh[folded arg_cong[OF FGH, of degree, folded dh]] choice unfolding df by (metis \degree (F + B * G) < df\ degree_add_eq_left degree_add_eq_right df nat_neq_iff) next case True have "F = - B * G" using arg_cong[OF FGH[unfolded True], of "\ x. x - B * G"] by auto thus ?thesis using F0 G0 B0 unfolding df by simp qed hence dfbg: "df = db + dg" using degree_mult_eq[OF B0 G0] by (simp add: d) hence dbfg: "db = df - dg" by simp let ?dfj = "df - J" let ?dgj = "dg - J" have norm: "?dgj + ?dfj = ?dfj + ?dgj" by simp let ?bij = "\ i j. b (db - int i + int (j - dfj))" let ?M = "mat n n (\ (i,j). if i = j then 1 else if j < dfj then 0 else if i < j then [: ?bij i j :] else 0)" (* this is the matrix which contains all the column-operations that are required to transform the subresultant-matrix of F and G into the one of G and H ( the matrix depicted in equation (18) in BT *) let ?GF = "\ i j. if j < dfj then if i = n - 1 then monom 1 (dfj - 1 - j) * G else [:g (int dg - int i + int j):] else if i = n - 1 then monom 1 (dgj - 1 - (j - dfj)) * F else [:f (int df - int i + int (j - dfj)):]" let ?G_F = "mat n n (\ (i,j). ?GF i j)" let ?GH = "\ i j. if j < dfj then if i = n - 1 then monom 1 (dfj - 1 - j) * G else [:g (int dg - int i + int j):] else if i = n - 1 then monom 1 (dgj - 1 - (j - dfj)) * H else [:h (int df - int i + int (j - dfj)):]" let ?G_H = "mat n n (\ (i,j). ?GH i j)" have hfg: "h i = f i + coeff_int (B * G) i" for i unfolding FGH[symmetric] f g h unfolding coeff_int_def by simp have dM1: "det ?M = 1" by (subst det_upper_triangular, (auto)[2], subst prod_list_diag_prod, auto) have "subresultant J F G = smult ?m1 (subresultant J G F)" unfolding subresultant_swap[of _ F] d by simp also have "subresultant J G F = det ?G_F" unfolding subresultant_def n norm subresultant_mat_def g f Let_def d[symmetric] dfj_def dgj_def by simp also have "\ = det (?G_F * ?M)" by (subst det_mult[of _ n], unfold dM1, auto) also have "?G_F * ?M = ?G_H" proof (rule eq_matI, unfold dim_col_mat dim_row_mat) fix i j assume i: "i < n" and j: "j < n" have "(?G_F * ?M) $$ (i,j) = row (?G_F) i \ col ?M j" using i j by simp also have "\ = ?GH i j" proof (cases "j < dfj") case True have id: "col ?M j = unit_vec n j" by (rule eq_vecI, insert True i j, auto) show ?thesis unfolding id using True i j by simp next case False define d where "d = j - dfj" from False have jd: "j = d + dfj" unfolding d_def by auto hence idj: "{0 ..< j} = {0 ..< dfj} \ {dfj ..< dfj + d}" by auto let ?H = "(if i = n - 1 then monom 1 (dgj - Suc d) * H else [:h (int df - int i + int d):])" have idr: "?GH i j = ?H" unfolding d_def using jd by auto let ?bi = "\ i. b (db - int i + int d)" let ?m = "\ i. if i = j then 1 else if i < j then [:?bij i j:] else 0" let ?P = "\ k. (?GF i k * ?m k)" let ?Q = "\ k. ?GF i k * [: ?bi k :]" let ?G = "\ k. if i = n - 1 then monom 1 (dfj - 1 - k) * G else [:g (int dg - int i + int k):]" let ?Gb = "\ k. ?G k * [:?bi k:]" let ?off = "- (int db - int dfj + 1 + int d)" have off0: "?off \ 0" using False dfg j unfolding dfj_def d_def dbfg n by simp from nat_0_le[OF this] obtain off where off: "int off = ?off" by blast have "int off \ int dfj" unfolding off by auto hence "off \ dfj" by simp hence split1: "{0 ..< dfj} = {0 ..< off} \ {off ..< dfj}" by auto have "int off + Suc db \ dfj" unfolding off by auto hence split2: "{off ..< dfj} = {off .. off + db} \ {off + Suc db ..< dfj} " by auto let ?g_b = "\k. (if i = n - 1 then monom 1 k * G else [:g (int dg - int i + int (dfj - Suc k)):]) * [:b (k - int off):]" let ?gb = "\k. (if i = n - 1 then monom 1 (k + off) * G else [:g (int dg - int i + int (dfj - Suc k - off)):]) * [:coeff B k:]" let ?F = "\ k. if i = n - 1 then monom 1 (dgj - 1 - (k - dfj)) * F else [:f (int df - int i + int (k - dfj)):]" let ?Fb = "\ k. ?F k * [:?bi k:]" let ?Pj = "if i = n - 1 then monom 1 (dgj - Suc d) * F else [:f (int df - int i + int d):]" from False have id: "col ?M j = vec n ?m" using j i by (intro eq_vecI, auto) have "row ?G_F i \ col ?M j = sum ?P {0 ..< n}" using i j unfolding id by (simp add: scalar_prod_def) also have "{0 ..< n} = {0 ..< j} \ {j} \ {Suc j ..< n}" using j by auto also have "sum ?P \ = sum ?P {0 ..< j} + ?P j + sum ?P {Suc j ..< n}" by (simp add: sum.union_disjoint) also have "sum ?P {Suc j ..< n} = 0" by (rule sum.neutral, auto) also have "?P j = ?Pj" unfolding d_def using jd by simp also have "sum ?P {0 ..< j} = sum ?Q {0 ..< j}" by (rule sum.cong[OF refl], unfold d_def, insert jd, auto) also have "sum ?Q {0 ..< j} = sum ?Q {0 ..< dfj} + sum ?Q {dfj ..< dfj+d}" unfolding idj by (simp add: sum.union_disjoint) also have "sum ?Q {0 ..< dfj} = sum ?Gb {0 ..< dfj}" by (rule sum.cong, auto) also have "sum ?Q {dfj ..< dfj+d} = sum ?Fb {dfj ..< dfj+d}" by (rule sum.cong, auto) also have "\ = 0" proof (rule sum.neutral, intro ballI) fix k assume k: "k \ {dfj ..< dfj+d}" hence k: "db + d < k" using k j False unfolding n db[symmetric] dfbg dfj_def d_def by auto let ?k = "(int db - int k + int d)" have "?k < 0" using k by auto hence "b ?k = 0" unfolding b by (intro coeff_int_eq_0, auto) thus "?Fb k = 0" by simp qed also have "sum ?Gb {0 ..< dfj} = sum ?g_b {0 ..< dfj}" proof (rule sum.reindex_cong[of "\ k. dfj - Suc k"], (auto simp: inj_on_def off)[2], goal_cases) case (1 k) hence "k = dfj - (Suc (dfj - Suc k))" and "(dfj - Suc k) \ {0.. = sum ?g_b {0 ..< off} + sum ?g_b {off ..< dfj}" unfolding split1 by (simp add: sum.union_disjoint) also have "sum ?g_b {0 ..< off} = 0" by (rule sum.neutral, intro ballI, auto simp: b coeff_int_def) also have "sum ?g_b {off ..< dfj} = sum ?g_b {off .. off + db} + sum ?g_b {off + Suc db ..< dfj}" unfolding split2 by (rule sum.union_disjoint, auto) also have "sum ?g_b {off + Suc db ..< dfj} = 0" proof (rule sum.neutral, intro ballI, goal_cases) case (1 k) hence "b (int k - int off) = 0" unfolding b db by (intro coeff_int_eq_0, auto) thus ?case by simp qed also have "sum ?g_b {off .. off + db} = sum ?gb {0 .. db}" using sum.atLeastAtMost_shift_bounds [of ?g_b 0 off db] by (auto intro: sum.cong simp add: b ac_simps) finally have id: "row ?G_F i \ col ?M j - ?H = ?Pj + sum ?gb {0 .. db} - ?H" (is "_ = ?E") by (simp add: ac_simps) define E where "E = ?E" let ?b = "coeff B" have Bsum: "(\k = 0..db. monom (?b k) k) = B" unfolding db using atMost_atLeast0 poly_as_sum_of_monoms by auto have "E = 0" proof (cases "i = n - 1") case i_n: False hence id: "(i = n - 1) = False" by simp with i have i: "i < n - 1" by auto let ?ii = "int df - int i + int d" have "?thesis = ([:f ?ii:] + (\k = 0..db. [:g (int dg - int i + int (dfj - Suc k - off)):] * [:?b k:]) - [:h ?ii:] = 0)" (is "_ = (?e = 0)") unfolding E_def id if_False by simp also have "?e = [: f ?ii + (\k = 0..db. g (int dg - int i + int (dfj - Suc k - off)) * ?b k) - h ?ii:]" (is "_ = [: ?e :]") proof (rule poly_eqI, goal_cases) case (1 n) show ?case unfolding coeff_diff coeff_add coeff_sum coeff_const by (cases n, auto simp: ac_simps) qed also have "[: ?e :] = 0 \ ?e = 0" by simp also have "?e = (\k = 0..db. g (int dg - int i + int (dfj - Suc k - off)) * ?b k) - coeff_int (B * G) ?ii" unfolding hfg by simp also have "(B * G) = (\k = 0..db. monom (?b k) k) * G" unfolding Bsum by simp also have "\ = (\k = 0..db. monom (?b k) k * G)" by (rule sum_distrib_right) also have "coeff_int \ ?ii = (\k = 0..db. g (?ii - k) * ?b k)" unfolding coeff_int_sum coeff_int_monom_mult g by (simp add: ac_simps) also have "\ = (\k = 0..db. g (int dg - int i + int (dfj - Suc k - off)) * ?b k)" proof (rule sum.cong[OF refl], goal_cases) case (1 k) hence "k \ db" by simp hence id: "int dg - int i + int (dfj - Suc k - off) = ?ii - k" using False i j off dfg unfolding dbfg d_def dfj_def n by linarith show ?case unfolding id .. qed finally show ?thesis by simp next case True let ?jj = "dgj - Suc d" have zero: "int off - (dgj - Suc d) = 0" using dfg False j unfolding off dbfg dfj_def d_def dgj_def n by linarith from True have "E = monom 1 ?jj * F + (\k = 0.. db. monom 1 (k + off) * G * [: ?b k :]) - monom 1 ?jj * H" (is "_ = ?A + ?sum - ?mon") unfolding id E_def by simp also have "?mon = monom 1 ?jj * F + monom 1 ?jj * (B * G)" unfolding FGH[symmetric] by (simp add: ring_distribs) also have "?A + ?sum - \ = ?sum - (monom 1 ?jj * G) * B" (is "_ = _ - ?GB * B") by simp also have "?sum = (\k = 0..db. (monom 1 ?jj * G) * (monom 1 (k + off - ?jj) * [: ?b k :]))" proof (rule sum.cong[OF refl], goal_cases) case (1 k) let ?one = "1 :: 'a" have "int off \ int ?jj" using j False i True unfolding off d_def dfj_def dgj_def dfbg n by linarith hence "k + off = ?jj + (k + off - ?jj)" by linarith hence id: "monom ?one (k + off) = monom (1 * 1) (?jj + (k + off - ?jj))" by simp show ?case unfolding id[folded mult_monom] by (simp add: ac_simps) qed also have "\ = (monom 1 ?jj * G) * (\k = 0..db. monom 1 (k + off - ?jj) * [:?b k:])" (is "_ = _ * ?sum") unfolding sum_distrib_left .. also have "\ - (monom 1 ?jj * G) * B = (monom 1 ?jj * G) * (?sum - B)" by (simp add: ring_distribs) also have "?sum = (\k = 0..db. monom 1 k * [:?b k:])" by (rule sum.cong[OF refl], insert zero, auto) also have "\ = (\k = 0..db. monom (?b k) k)" by (rule sum.cong[OF refl], rule poly_eqI, auto) also have "\ = B" unfolding Bsum .. finally show ?thesis by simp qed from id[folded E_def, unfolded this] show ?thesis using False unfolding d_def by simp qed also have "\ = ?G_H $$ (i,j)" using i j by simp finally show "(?G_F * ?M) $$ (i,j) = ?G_H $$ (i,j)" . qed auto finally show eq_18: "subresultant J F G = smult ?m1 (det ?G_H)" unfolding dfj_def dgj_def . { fix i j assume ij: "i < j" and j: "j < n" with dgh have "int dg - int i + int j > int dg" by auto hence "g (int dg - int i + int j) = 0" unfolding g dg by (intro coeff_int_eq_0, auto) } note g0 = this { assume *: "dh \ J" "J < dg" have n_dfj: "n > dfj" using * unfolding n dfj_def by auto note eq_18 also have "det ?G_H = prod_list (diag_mat ?G_H)" proof (rule det_lower_triangular[of n]) fix i j assume ij: "i < j" and j: "j < n" from ij j have if_e: "i = n - 1 \ False" by auto have "?G_H $$ (i,j) = ?GH i j" using ij j by auto also have "\ = 0" proof (cases "j < dfj") case True with True g0[OF ij j] show ?thesis unfolding if_e by simp next case False have "h (int df - int i + int (j - dfj)) = 0" unfolding h by (rule coeff_int_eq_0, insert False * ij j dfg, unfold dfj_def dh[symmetric], auto) with False show ?thesis unfolding if_e by auto qed finally show "?G_H $$ (i,j) = 0" . qed auto also have "\ = (\i = 0.. {dfj ..< n}" unfolding n dfj_def by auto also have "prod (\ i. ?GH i i) \ = prod (\ i. ?GH i i) {0 ..< dfj} * prod (\ i. ?GH i i) {dfj ..< n}" by (simp add: prod.union_disjoint) also have "prod (\ i. ?GH i i) {0 ..< dfj} = prod (\ i. [: lead_coeff G :]) {0 ..< dfj}" proof - show ?thesis by (rule prod.cong[OF refl], insert n_dfj, auto simp: g coeff_int_def dg) qed also have "\ = [: (lead_coeff G)^dfj :]" by (simp add: poly_const_pow) also have "{dfj ..< n} = {dfj ..< n-1} \ {n - 1}" using n_dfj by auto also have "prod (\ i. ?GH i i) \ = prod (\ i. ?GH i i) {dfj ..< n-1} * ?GH (n - 1) (n - 1)" by (simp add: prod.union_disjoint) also have "?GH (n - 1) (n - 1) = H" proof - have "dgj - 1 - (n - 1 - dfj) = 0" using n_dfj unfolding dgj_def dfj_def n by auto with n_dfj show ?thesis by auto qed also have "prod (\ i. ?GH i i) {dfj ..< n-1} = prod (\ i. [:h (int df - dfj):]) {dfj ..< n-1}" by (rule prod.cong[OF refl], auto intro!: arg_cong[of _ _ h]) also have "\ = [: h (int df - dfj) ^ (n - 1 - dfj) :]" unfolding prod_constant by (simp add: poly_const_pow) also have "n - 1 - dfj = dg - J - 1" unfolding n dfj_def by simp also have "int df - dfj = J" using * dfg unfolding dfj_def by auto also have "h J = coeff H J" unfolding h coeff_int_def by simp finally show "subresultant J F G = smult (?m1 * ?G * ?H) H" by (simp add: dfj_def ac_simps) } note eq_19 = this { assume J: "J < dh" define dhj where "dhj = dh - J" have n_add: "n = (df - dh) + (dhj + dgj)" unfolding dhj_def dgj_def n using J dfg dgh by auto let ?split = "split_block ?G_H (df - dh) (df - dh)" have dim: "dim_row ?G_H = (df - dh) + (dhj + dgj)" "dim_col ?G_H = (df - dh) + (dhj + dgj)" unfolding n_add by auto obtain UL UR LL LR where spl: "?split = (UL,UR,LL,LR)" by (cases ?split, auto) note spl' = spl[unfolded split_block_def Let_def, simplified] let ?LR = "subresultant_mat J G H" have "LR = mat (dgj + dhj) (dgj + dhj) (\ (i,j). ?GH (i + (df - dh)) (j + (df - dh)))" using spl' by (auto simp: n_add) also have "\ = ?LR" unfolding subresultant_mat_def Let_def dhj_def dgj_def d[symmetric] proof (rule eq_matI, unfold dim_row_mat dim_col_mat index_mat split dfj_def, goal_cases) case (1 i j) hence id1: "(j + (df - dh) < df - J) = (j < dh - J)" using dgh dfg J by auto have id2: "(i + (df - dh) = n - 1) = (i = dg - J + (dh - J) - 1)" unfolding n_add dhj_def dgj_def using dgh dfg J by auto have id3: "(df - J - 1 - (j + (df - dh))) = (dh - J - 1 - j)" and id4: "(int dg - int (i + (df - dh)) + int (j + (df - dh))) = (int dg - int i + int j)" and id5: "(dg - J - 1 - (j + (df - dh) - (df - J))) = (dg - J - 1 - (j - (dh - J)))" and id6: "(int df - int (i + (df - dh)) + int (j + (df - dh) - (df - J))) = (int dh - int i + int (j - (dh - J)))" using dgh dfg J by auto show ?case unfolding g[symmetric] h[symmetric] id3 id4 id5 id6 by (rule if_cong[OF id1 if_cong[OF id2 refl refl] if_cong[OF id2 refl refl]]) qed auto finally have "LR = ?LR" . note spl = spl[unfolded this] let ?UR = "0\<^sub>m (df - dh) (dgj + dhj)" have "UR = mat (df - dh) (dgj + dhj) (\ (i,j). ?GH i (j + (df - dh)))" using spl' by (auto simp: n_add) also have "\ = ?UR" proof (rule eq_matI, unfold dim_row_mat dim_col_mat index_mat split dfj_def index_zero_mat, goal_cases) case (1 i j) hence in1: "i \ n - 1" using J unfolding dgj_def dhj_def n_add by auto { assume "j + (df - dh) < df - J" hence "dg < int dg - int i + int (j + (df - dh))" using 1 J unfolding dgj_def dhj_def by auto hence "g \ = 0" unfolding dg g by (intro coeff_int_eq_0, auto) } note g = this { assume "\ (j + (df - dh) < df - J)" hence "dh < int df - int i + int (j + (df - dh) - (df - J))" using 1 J unfolding dgj_def dhj_def by auto hence "h \ = 0" unfolding dh h by (intro coeff_int_eq_0, auto) } note h = this show ?case using in1 g h by auto qed auto finally have "UR = ?UR" . note spl = spl[unfolded this] let ?G = "\ (i,j). if i = j then [:lead_coeff G:] else if i < j then 0 else ?GH i j" let ?UL = "mat (df - dh) (df - dh) ?G" have "UL = mat (df - dh) (df - dh) (\ (i,j). ?GH i j)" using spl' by (auto simp: n_add) also have "\ = ?UL" proof (rule eq_matI, unfold dim_row_mat dim_col_mat index_mat split, goal_cases) case (1 i j) { assume "i = j" hence "int dg - int i + int j = dg" using 1 by auto hence "g (int dg - int i + int j) = lead_coeff G" unfolding g dg coeff_int_def by simp } note eq = this { assume "i < j" hence "dg < int dg - int i + int j" using 1 by auto hence "g (int dg - int i + int j) = 0" unfolding g dg by (intro coeff_int_eq_0, auto) } note lt = this from 1 have *: "j < dfj" "i \ n - 1" using J unfolding n_add dhj_def dgj_def dfj_def by auto hence "?GH i j = [:g (int dg - int i + int j):]" by simp also have "\ = (if i = j then [: lead_coeff G :] else if i < j then 0 else ?GH i j)" using eq lt * by auto finally show ?case by simp qed auto finally have "UL = ?UL" . note spl = spl[unfolded this] from split_block[OF spl dim] have GH: "?G_H = four_block_mat ?UL ?UR LL ?LR" and C: "?UL \ carrier_mat (df - dh) (df - dh)" "?UR \ carrier_mat (df - dh) (dhj + dgj)" "LL \ carrier_mat (dhj + dgj) (df - dh)" "?LR \ carrier_mat (dhj + dgj) (dhj + dgj)" by auto from arg_cong[OF GH, of det] have "det ?G_H = det (four_block_mat ?UL ?UR LL ?LR)" unfolding GH[symmetric] .. also have "\ = det ?UL * det ?LR" by (rule det_four_block_mat_upper_right_zero[OF _ refl], insert C, auto simp: ac_simps) also have "det ?LR = subresultant J G H" unfolding subresultant_def by simp also have "det ?UL = prod_list (diag_mat ?UL)" by (rule det_lower_triangular[of "df - dh"], auto) also have "\ = (\i = 0..< (df - dh). [: lead_coeff G :])" unfolding prod_list_diag_prod by simp also have "\ = [: lead_coeff G ^ (df - dh) :]" by (simp add: poly_const_pow) finally have det: "det ?G_H = [:lead_coeff G ^ (df - dh):] * subresultant J G H" by auto show "subresultant J F G = smult (?m1 * lead_coeff G ^ (df - dh)) (subresultant J G H)" unfolding eq_18 det by simp } { assume J: "dh < J" "J < dg - 1" hence "dh \ J" "J < dg" by auto from eq_19[OF this] have "subresultant J F G = smult ((- 1) ^ ((df - J) * (dg - J)) * lead_coeff G ^ (df - J) * coeff H J ^ (dg - J - 1)) H" by simp also have "coeff H J = 0" by (rule coeff_eq_0, insert J, auto simp: dh) also have "\ ^ (dg - J - 1) = 0" using J by auto finally show "subresultant J F G = 0" by simp } { assume J: "J = dh" and "dg > dh \ H \ 0" with choice have dgh: "dg > dh" by auto show "subresultant dh F G = smult ( (-1)^((df - dh) * (dg - dh)) * lead_coeff G ^ (df - dh) * lead_coeff H ^ (dg - dh - 1)) H" unfolding eq_19[unfolded J, OF le_refl dgh] unfolding dh by simp } { assume J: "J = dg - 1" and "dg > dh \ H \ 0" with choice have dgh: "dg > dh" by auto have *: "dh \ dg - 1" "dg - 1 < dg" using dgh by auto have **: "df - (dg - 1) = df - dg + 1" "dg - (dg - 1) - 1 = 0" "dg - (dg - 1) = 1" using dfg dgh by linarith+ show "subresultant (dg - 1) F G = smult ( (-1)^(df - dg + 1) * lead_coeff G ^ (df - dg + 1)) H" unfolding eq_19[unfolded J, OF *] unfolding ** by simp } qed lemmas BT_lemma_1_13 = BT_lemma_1_13'[OF _ _ _ refl] lemmas BT_lemma_1_15 = BT_lemma_1_15'[OF _ _ _ refl] lemma subresultant_product: fixes F :: "'a :: idom poly" assumes "F = B * G" and FG: "degree F \ degree G" shows "subresultant J F G = (if J < degree G then 0 else if J < degree F then smult (lead_coeff G ^ (degree F - J - 1)) G else 1)" proof (cases "J < degree G") case J: True from assms have eq: "F + (-B) * G = 0" by auto from J have lt: "degree 0 < degree G \ b" for b by auto from BT_lemma_1_13[OF eq FG lt lt] have "subresultant 0 F G = 0" using J by auto with BT_lemma_1_14[OF eq FG lt, of J] have 00: "J = 0 \ J < degree G - 1 \ subresultant J F G = 0" by auto from BT_lemma_1_15[OF eq FG lt lt] J have 01: "subresultant (degree G - 1) F G = 0" by simp from J have "(J = 0 \ J < degree G - 1) \ J = degree G - 1" by linarith with 00 01 have "subresultant J F G = 0" by auto thus ?thesis using J by simp next case J: False hence dg: "degree G - J = 0" by simp let ?n = "degree F - J" have *: "(j :: nat) < 0 \ False" "j - 0 = j" for j by auto let ?M = "mat ?n ?n (\(i, j). if i = ?n - 1 then monom 1 (?n - 1 - j) * G else [:coeff_int G (int (degree G) - int i + int j):])" have "subresultant J F G = det ?M" unfolding subresultant_def subresultant_mat_def Let_def dg * by auto also have "det ?M = prod_list (diag_mat ?M)" by (rule det_lower_triangular[of ?n], auto intro: coeff_int_eq_0) also have "\ = (\i = 0..< ?n. ?M $$ (i,i))" unfolding prod_list_diag_prod by simp also have "\ = (\i = 0..< ?n. if i = ?n - 1 then G else [: lead_coeff G :])" by (rule prod.cong[OF refl], auto simp: coeff_int_def) also have "\ = (if J < degree F then smult (lead_coeff G ^ (?n - 1)) G else 1)" proof (cases "J < degree F") case True hence id: "{ 0 ..< ?n} = { 0 ..< ?n - 1} \ {?n - 1}" by auto have "(\i = 0..< ?n. if i = ?n - 1 then G else [: lead_coeff G :]) = (\i = 0 ..< ?n - 1. if i = ?n - 1 then G else [: lead_coeff G :]) * G" (is "_ = ?P * G") unfolding id by (subst prod.union_disjoint, auto) also have "?P = (\i = 0 ..< ?n - 1. [: lead_coeff G :])" by (rule prod.cong, auto) also have "\ = [: lead_coeff G ^ (?n - 1) :]" by (simp add: poly_const_pow) finally show ?thesis by auto qed auto finally have "subresultant J F G = (if J < degree F then smult (lead_coeff G ^ (degree F - J - 1)) G else 1)" . thus ?thesis using J by simp qed lemma resultant_pseudo_mod_0: assumes "pseudo_mod f g = (0 :: 'a :: idom_divide poly)" and dfg: "degree f \ degree g" and f: "f \ 0" and g: "g \ 0" shows "resultant f g = (if degree g = 0 then lead_coeff g^degree f else 0)" proof - let ?df = "degree f" let ?dg = "degree g" obtain d r where pd: "pseudo_divmod f g = (d,r)" by force from pd have r: "r = pseudo_mod f g" unfolding pseudo_mod_def by simp with assms pd have pd: "pseudo_divmod f g = (d,0)" by auto from pseudo_divmod[OF g pd] g obtain a q where prod: "smult a f = g * q" and a: "a \ 0" "a = lead_coeff g ^ (Suc ?df - ?dg)" by auto from a dfg have dfg: "degree g \ degree (smult a f)" by auto have g0: "degree g = 0 \ coeff g 0 = 0 \ g = 0" using leading_coeff_0_iff by fastforce from prod have "smult a f = q * g" by simp from arg_cong[OF subresultant_product[OF this dfg, of 0, unfolded subresultant_resultant resultant_smult_left[OF a(1)]], of "\ x. coeff x 0"] show ?thesis using a g0 by (cases "degree f", auto) qed locale primitive_remainder_sequence = fixes F :: "nat \ 'a :: idom_divide poly" and n :: "nat \ nat" and \ :: "nat \ nat" and f :: "nat \ 'a" and k :: nat and \ :: "nat \ 'a" assumes f: "\ i. f i = lead_coeff (F i)" and n: "\ i. n i = degree (F i)" and \: "\ i. \ i = n i - n (Suc i)" and n12: "n 1 \ n 2" and F12: "F 1 \ 0" "F 2 \ 0" and F0: "\ i. i \ 0 \ F i = 0 \ i > k" and \0: "\ i. \ i \ 0" and pmod: "\ i. i \ 3 \ i \ Suc k \ smult (\ i) (F i) = pseudo_mod (F (i - 2)) (F (i - 1))" begin lemma f10: "f 1 \ 0" and f20: "f 2 \ 0" unfolding f using F12 by auto lemma f0: "i \ 0 \ f i = 0 \ i > k" using F0[of i] unfolding f by auto lemma n_gt: assumes "2 \ i" "i < k" shows "n i > n (Suc i)" proof - from assms have "3 \ Suc i" "Suc i \ Suc k" by auto note pmod = pmod[OF this] from assms F0 have "F (Suc i - 1) \ 0" "F (Suc i) \ 0" by auto from pseudo_mod(2)[OF this(1), of "F (Suc i - 2)", folded pmod] this(2) show ?thesis unfolding n using \0 by auto qed lemma n_ge: assumes "1 \ i" "i < k" shows "n i \ n (Suc i)" using n12 n_gt[OF _ assms(2)] assms(1) by (cases "i = 1", auto simp: numeral_2_eq_2) lemma n_ge_trans: assumes "1 \ i" "i \ j" "j \ k" shows "n i \ n j" proof - from assms(2) have "j = i + (j - i)" by simp then obtain jj where j: "j = i + jj" by blast from assms(3)[unfolded j] show ?thesis unfolding j proof (induct jj) case (Suc j) from Suc(2) have "i + j \ k" by simp from Suc(1)[OF this] have IH: "n (i + j) \ n i" . have "n (Suc (i + j)) \ n (i + j)" by (rule n_ge, insert assms(1) Suc(2), auto) with IH show ?case by auto qed auto qed lemma delta_gt: assumes "2 \ i" "i < k" shows "\ i > 0" using n_gt[OF assms] unfolding \ by auto lemma k2:"2 \ k" by (metis le_cases linorder_not_le F0 F12(2) zero_order(2)) lemma k0: "k \ 0" using k2 by auto lemma ni2:"3 \ i \ i \ k \ n i \ n 2" by (metis Suc_numeral \ delta_gt k2 le_imp_less_Suc le_less n_ge_trans not_le one_le_numeral semiring_norm(5) zero_less_diff) end locale subresultant_prs_locale = primitive_remainder_sequence F n \ f k \ for F :: "nat \ 'a :: idom_divide fract poly" and n :: "nat \ nat" and \ :: "nat \ nat" and f :: "nat \ 'a fract" and k :: nat and \ :: "nat \ 'a fract" + fixes G1 G2 :: "'a poly" assumes F1: "F 1 = map_poly to_fract G1" and F2: "F 2 = map_poly to_fract G2" begin definition "\ i = (f (i - 1))^(Suc (\ (i - 2)))" lemma \0: "i > 1 \ \ i = 0 \ (i - 1) > k" unfolding \_def using f0[of "i - 1"] by auto lemma \_char: assumes "3 \ i" "i < k + 2" shows "\ i = (f (i - 1)) ^ (Suc (length (coeffs (F (i - 2)))) - length (coeffs (F (i - 1))))" proof (cases "i = 3") case True have triv:"Suc (Suc 0) = 2" by auto have l:"length (coeffs (F 2)) \ 0" "length (coeffs (F 1)) \ 0" using F12 by auto hence "length (coeffs (F 2)) \ length (coeffs (F (Suc 0)))" using n12 unfolding n degree_eq_length_coeffs One_nat_def by linarith hence "Suc (length (coeffs (F 1)) - 1 - (length (coeffs (F 2)) - 1)) = (Suc (length (coeffs (F 1))) - length (coeffs (F (3 - 1))))" using l by simp thus ?thesis unfolding True \_def n \ degree_eq_length_coeffs by (simp add:triv) next case False hence assms:"2 \ i - 2" "i - 2 < k" using assms by auto have i:"i - 2 \ 0" "i - 1 \ 0" using assms by auto hence [simp]: "Suc (i - 2) = i - 1" by auto from assms(2) F0[OF i(2)] have "F (i - 1) \ 0" by auto then have "length (coeffs (F (i - 1))) > 0" by (cases "F (i - 1)") auto with delta_gt[unfolded \ n degree_eq_length_coeffs,OF assms] have * : "Suc (\ (i - 2)) = Suc (length (coeffs (F (i - 2)))) - (length (coeffs (F (Suc (i - 2)))))" by (auto simp:\ n degree_eq_length_coeffs) show ?thesis unfolding \_def * by simp qed definition Q :: "nat \ 'a fract poly" where "Q i \ smult (\ i) (fst (pdivmod (F (i - 2)) (F (i - 1))))" lemma beta_F_as_sum: assumes "3 \ i" "i \ Suc k" shows "smult (\ i) (F i) = smult (\ i) (F (i - 2)) + - Q i * F (i - 1)" (is ?t1) proof - have ik:"i < k + 2" using assms by auto have f0:"F (i - 1) = 0 \ False" "F (i - Suc 0) = 0 \ False" using F0[of "i - 1"] assms by auto hence f0_b:"(inverse (coeff (F (i - 1)) (degree (F (i - 1))))) \ 0" "F (i - 1) \ 0" by auto have i:"i - 2 \ 0" "Suc (i - 2) = i - 1" "(k < i - 2) \ False" using assms by auto have "F (i - 2) \ 0" using F0[of "i - 2"] assms by auto let ?c = "(inverse (f (i - 1)) ^ (Suc (length (coeffs (F (i - 2)))) - length (coeffs (F (i - 1)))))" have inv:"inverse (\ i) = ?c" unfolding \_char[OF assms(1) ik] power_inverse by auto have alpha0:"\ i \ 0" unfolding \_def f using f0 by auto have \_inv[simp]:"\ i * inverse (\ i) = 1" using field_class.field_inverse[OF alpha0] mult.commute by metis with field_class.field_inverse[OF alpha0,unfolded inv] have c_times_Q:"smult ?c (Q i) = fst (pdivmod (F (i - 2)) (F (i - 1)))" unfolding Q_def by auto have "pdivmod (F (i - 2)) (F (i - 1)) = (smult ?c (Q i), smult ?c (smult (\ i) (F i)))" unfolding c_times_Q unfolding pdivmod_via_pseudo_divmod pmod[OF assms] f n c_times_Q pseudo_mod_smult_right[OF f0_b, of "F (i - 2)",symmetric] f0 if_False Let_def unfolding pseudo_mod_def by (auto split:prod.split) - from this[folded pdivmod_pdivmodrel] - have pr:"F (i - 2) = smult ?c (Q i) * F (i - 1) + smult ?c ( smult (\ i) (F i))" - by (auto simp: eucl_rel_poly_iff) - hence "F (i - 2) = smult (inverse (\ i)) (Q i) * F (i - 1) + from this [symmetric] + have pr: \F (i - 2) = smult ?c (Q i) * F (i - 1) + smult ?c ( smult (\ i) (F i))\ + by (simp only: prod_eq_iff fst_conv snd_conv div_mult_mod_eq) + then have "F (i - 2) = smult (inverse (\ i)) (Q i) * F (i - 1) + smult (inverse (\ i)) ( smult (\ i) (F i))" (is "?l = ?r" is "_ = ?t + _") unfolding inv. hence eq:"smult (\ i) (?l - ?t) = smult (\ i) (?r - ?t)" by auto have " smult (\ i) (F (i - 2)) - Q i * (F (i - 1)) = smult (\ i) (?l - ?t)" unfolding smult_diff_right by auto also have "\ = smult (\ i) (?r - ?t)" unfolding eq.. also have "\ = smult (\ i) (F i)" by (auto simp:mult.assoc[symmetric]) finally show ?t1 by auto qed lemma assumes "3 \ i" "i \ k" shows BT_lemma_2_21: "j < n i \ smult (\ i ^ (n (i - 1) - j)) (subresultant j (F (i - 2)) (F (i - 1))) = smult ((- 1) ^ ((n (i - 2) - j) * (n (i - 1) - j)) * (f (i - 1)) ^ (\ (i - 2) + \ (i - 1)) * (\ i) ^ (n (i - 1) - j)) (subresultant j (F (i - 1)) (F i))" (is "_ \ ?eq_21") and BT_lemma_2_22: "smult (\ i ^ (\ (i - 1))) (subresultant (n i) (F (i - 2)) (F (i - 1))) = smult ((- 1) ^ ((\ (i - 2) + \ (i - 1)) * \ (i - 1)) * f (i - 1) ^ (\ (i - 2) + \ (i - 1)) * f i ^ (\ (i - 1) - 1) * (\ i) ^ \ (i - 1)) (F i)" (is "?eq_22") and BT_lemma_2_23: "n i < j \ j < n (i - 1) - 1 \ subresultant j (F (i - 2)) (F (i - 1)) = 0" (is "_ \ _ \ ?eq_23") and BT_lemma_2_24: "smult (\ i) (subresultant (n (i - 1) - 1) (F (i - 2)) (F (i - 1))) = smult ((- 1) ^ (\ (i - 2) + 1) * f (i - 1) ^ (\ (i - 2) + 1) * \ i) (F i)" (is "?eq_24") proof - from assms have ik:"i \ Suc k" by auto note beta_F_as_sum = beta_F_as_sum[OF assms(1) ik, symmetric] have s[simp]:"Suc (i - 2) = i - 1" "Suc (i - 1) = i" using assms by auto have \0:"\ i \ 0" using assms f0[of "i - 1"] unfolding \_def f by auto hence \0pow:"\ x. \ i ^ x \ 0" by auto have df:"degree (F (i - 1)) \ degree (smult (\ i) (F (i - 2)))" "degree (smult (\ i) (F i)) < degree (F (i - 1)) \ b" for b using n_ge[of "i-2"] n_gt[of "i-1"] assms \0 \0 unfolding n by auto have degree_smult_eq:"\ c f. (c::_::{idom_divide}) \ 0 \ degree (smult c f) = degree f" by auto have n_lt:"n i < n (i - 1)" using n_gt[of "i-1"] assms unfolding n by auto from semiring_normalization_rules(30) mult.commute have *:"\ x y q. (x * (y::'a fract)) ^ q = y ^ q * x ^ q" by metis have "n (i - 1) - n i > 0" using n_lt by auto hence **:"\ i ^ (n (i - 1) - n i - 1) * \ i = \ i ^ (n (i - 1) - n i)" by (subst power_minus_mult) auto have "max (n (i - 2)) (n (i - 1)) = n (i - 2)" using n_ge[of "i-2"] assms unfolding max_def by auto with diff_add_assoc[OF n_ge[of "i-1"],symmetric] assms have ns : "n (i - 2) - n (i - 1) + (n (i - 1) - n i) = n (i - 2) - n i" by (auto simp:nat_minus_add_max) { assume "j < n i" hence j:"j < degree (smult (\ i) (F i))" using \0 unfolding n by auto from BT_lemma_1_12[OF beta_F_as_sum df j] show ?eq_21 unfolding subresultant_smult_right[OF \0] subresultant_smult_left[OF \0] degree_smult_eq[OF \0] degree_smult_eq[OF \0] n[symmetric] f[symmetric] \ s ns using f n by auto} { from BT_lemma_1_13[OF beta_F_as_sum df df(2)] show ?eq_22 unfolding subresultant_smult_left[OF \0] lead_coeff_smult smult_smult degree_smult_eq[OF \0] degree_smult_eq[OF \0] n[symmetric] f[symmetric] \ s ns by (metis (no_types, lifting) * ** coeff_smult f mult.assoc n)} { assume "n i < j" "j < n (i - 1) - 1" hence j:"degree (smult (\ i) (F i)) < j" "j < degree (F (i - 1)) - 1" using \0 unfolding n by auto from BT_lemma_1_14[OF beta_F_as_sum df j] show ?eq_23 unfolding subresultant_smult_left[OF \0] smult_eq_0_iff using \0pow by auto} { have ***: "n (i - 1) - (n (i - 1) - 1) = 1" using n_lt by auto from BT_lemma_1_15[OF beta_F_as_sum df df(2)] show ?eq_24 unfolding subresultant_smult_left[OF \0] *** degree_smult_eq[OF \0] n[symmetric] f \ by (auto simp:mult.commute)} qed lemma BT_eq_30: "3 \ i \ i \ k + 1 \ j < n (i - 1) \ smult (\l\[3.. l ^ (n (l - 1) - j)) (subresultant j (F 1) (F 2)) = smult (\l\[3.. l ^ (n (l - 1) - j) * f (l - 1) ^ (\ (l - 2) + \ (l - 1)) * (- 1) ^ ((n (l - 2) - j) * (n (l - 1) - j))) (subresultant j (F (i - 2)) (F (i - 1)))" proof (induct "i - 3" arbitrary:i) case (Suc x) from Suc.hyps(2) Suc.prems(1-2) have prems:"x = (i - 1) - 3" "3 \ i - 1" "i - 1 \ k + 1" "2 \ i - 1 - 1" "i - 1 - 1 < k" "i - 1 \ k" by auto from prems(2) have inset:"i - 1 \ set [3.. c d e x. smult c d = e \ smult (x * c) d = smult x e" by auto have **:"\ c d e x. smult c d = e \ smult c (smult x d) = smult x e" by (auto simp:mult.commute) show ?case unfolding prod_list_map_remove1[OF inset(1),unfolded r1] *[OF Suc.hyps(1)[OF prems(1-3) j]] **[OF BT_lemma_2_21[OF prems(2,6) Suc.prems(3)]] by (auto simp: numeral_2_eq_2 ac_simps) qed auto lemma nonzero_alphaprod: assumes "i \ k + 1" shows "(\l\[3.. l ^ (p l)) \ 0" unfolding prod_list_zero_iff using assms by (auto simp: \0) lemma BT_eq_30': assumes i: "3 \ i" "i \ k + 1" "j < n (i - 1)" shows "subresultant j (F 1) (F 2) = smult ((- 1) ^ (\l\[3..l\[3.. l / \ l) ^ (n (l - 1) - j)) * (\l\[3.. (l - 2) + \ (l - 1)))) (subresultant j (F (i - 2)) (F (i - 1)))" (is "_ = smult (?mm * ?b * ?f) _") proof - let ?a = "\l\[3.. l ^ (n (l - 1) - j)" let ?d = "\l\[3.. l ^ (n (l - 1) - j) * f (l - 1) ^ (\ (l - 2) + \ (l - 1)) * (- 1) ^ ((n (l - 2) - j) * (n (l - 1) - j))" let ?m = "\l\[3.. 0" by (rule nonzero_alphaprod, rule i) with arg_cong[OF BT_eq_30[OF i], of "smult (inverse ?a)", unfolded smult_smult] have "subresultant j (F 1) (F 2) = smult (inverse ?a * ?d) (subresultant j (F (i - 2)) (F (i - 1)))" by simp also have "inverse ?a * ?d = ?b * ?f * ?m" unfolding prod_list_multf inverse_prod_list map_map o_def power_inverse[symmetric] power_mult_distrib divide_inverse_commute by simp also have "?m = ?mm" unfolding prod_list_minus_1_exp by simp finally show ?thesis by (simp add: ac_simps) qed text \For defining the subresultant PRS, we mainly follow Brown's ``The Subresultant PRS Algorithm'' (B).\ definition "R j = (if j = n 2 then sdiv_poly (smult ((lead_coeff G2)^(\ 1)) G2) (lead_coeff G2) else subresultant j G1 G2)" abbreviation "ff i \ to_fract (i :: 'a)" abbreviation "ffp \ map_poly ff" sublocale map_poly_hom: map_poly_inj_idom_hom to_fract.. (* for \ and \ we only take additions, so that no negative number-problems occur *) definition "\ i = (\l\[3.. i = (\l\[3.. i = (-1)^(\ i) * pow_int ( f (i - 1)) (1 - int (\ (i - 1))) * (\l\[3.. l / \ l)^(n (l - 1) - n (i - 1) + 1) * (f (l - 1))^(\ (l - 2) + \ (l - 1)))" definition "\ i = (-1)^(\ i) * pow_int (f i) (int (\ (i - 1)) - 1) * (\l\[3.. l / \ l)^(n (l - 1) - n i) * (f (l - 1))^(\ (l - 2) + \ (l - 1)))" (* is eq 29 in BT *) lemma fundamental_theorem_eq_4: assumes i: "3 \ i" "i \ k" shows "ffp (R (n (i - 1) - 1)) = smult (\ i) (F i)" proof - have "n (i - 1) \ n 2" by (rule n_ge_trans, insert i, auto) with n_gt[of "i - 1"] i have "n (i - 1) - 1 < n 2" and lt: "n (i - 1) - 1 < n (i - 1)" by linarith+ hence "R (n (i - 1) - 1) = subresultant (n (i - 1) - 1) G1 G2" unfolding R_def by auto from arg_cong[OF this, of ffp, unfolded to_fract_hom.subresultant_hom, folded F1 F2] have id1: "ffp (R (n (i - 1) - 1)) = subresultant (n (i - 1) - 1) (F 1) (F 2)" . note eq_24 = BT_lemma_2_24[OF i] let ?o = "(- 1) :: 'a fract" let ?m1 = "(\ (i - 2) + 1)" let ?d1 = "f (i - 1) ^ (\ (i - 2) + 1) * \ i" let ?c1 = "?o ^ ?m1 * ?d1" let ?c0 = "\ i" have "?c0 \ 0" using \0[of i] i by auto with arg_cong[OF eq_24, of "smult (inverse ?c0)"] have id2: "subresultant (n (i - 1) - 1) (F (i - 2)) (F (i - 1)) = smult (inverse ?c0 * ?c1) (F i)" by (auto intro: poly_eqI) from i have "3 \ i" "i \ k + 1" by auto note id3 = BT_eq_30'[OF this lt] let ?f = "\ l. f (l - 1) ^ (\ (l - 2) + \ (l - 1))" let ?b = "\ l. (\ l / \ l) ^ (n (l - 1) - (n (i - 1) - 1))" let ?b' = "\ l. (\ l / \ l) ^ (n (l - 1) - n (i - 1) + 1)" let ?m = "\ l. (n (l - 2) - (n (i - 1) - 1)) * (n (l - 1) - (n (i - 1) - 1))" let ?m' = "\ l. (n (l - 2) + n (i - 1) + 1) * (n (l - 1) + n (i - 1) + 1)" let ?m2 = "(\l\[3.. i = ?o ^ (?m1 + ?m2) * (inverse ?c0 * ?d1 * ?b2 * ?f2)" proof - have id: "\ i = (-1)^(\ i) * (?f1 * (\l\[3..l\[3.._def prod_list_multf by simp have cong: "even m1 = even m2 \ c1 = c2 \ ?o^m1 * c1 = ?o^m2 * c2" for m1 m2 c1 c2 unfolding minus_1_power_even by auto show ?thesis unfolding id proof (rule cong) from n_gt[of "i - 1"] i have n1: "n (i - 1) \ 0" by linarith { fix l assume "2 \ l" "l \ i" hence l: "l \ 2" "l - 1 \ i - 1" "l \ k" using i by auto from n_ge_trans[OF _ l(2)] l i have n2: "n (i - 1) \ n (l - 1)" by auto from n1 n2 have id: "n (l - 1) - (n (i - 1) - 1) = n (l - 1) - n (i - 1) + 1" by auto have "even (n (l - 1) - (n (i - 1) - 1)) = even (n (l - 1) + n (i - 1) + 1)" unfolding id using n2 by auto note id n2 this } note diff = this have f0: "f (i - 1) \ 0" using f0[of "i - 1"] i by auto have "(\l\[3..l\[3.. = ?b2 * ?b i" using i by auto finally have "?f1 * (\l\[3..l\[3.. i * inverse ?c0" using n1 by (simp add: divide_inverse) also have "?f1 * ?f i = f (i - 1) ^ (\ (i - 2) + 1)" unfolding exp_pow_int pow_int_add[OF f0, symmetric] by simp finally show "?f1 * (\l\[3..l\[3.. i) = even ((\l\[3.._def using i by simp also have "\ = (even (\l\[3..l\[3.. 2" "l \ i" and l1: "l - 1 \ 2" "l - 1 \ i" by auto have l2: "l - 2 = l - 1 - 1" by simp show ?case using diff(3) [OF l] diff(3) [OF l1] l2 by auto qed also have "even (?m' i) = even ?m1" proof - from i have id: "Suc (i - 1 - 1) = i - 1" "i - 2 = i - 1 - 1" by auto have "even ?m1 = even (n (i - 2) + n (i - 1) + 1)" unfolding \ id using diff[of "i - 1"] i by auto also have "\ = even (?m' i)" by auto finally show ?thesis by simp qed also have "(even ?m2 = even ?m1) = even (?m2 + ?m1)" unfolding even_add by simp also have "?m2 + ?m1 = ?m1 + ?m2" by simp finally show "even (\ i) = even (?m1 + ?m2)" . qed qed show ?thesis unfolding id1 id3 id2 smult_smult id4 by (simp add: ac_simps power_add) qed (* equation 28 in BT *) lemma fundamental_theorem_eq_5: assumes i: "3 \ i" "i \ k" "n i < j" "j < n (i - 1) - 1" shows "R j = 0" proof - from BT_lemma_2_23[OF i] have id1: "subresultant j (F (i - 2)) (F (i - 1)) = 0" . have "n (i - 1) \ n 2" by (rule n_ge_trans, insert i, auto) with n_gt[of "i - 1"] i have "n (i - 1) - 1 < n 2" and lt: "j < n (i - 1)" by linarith+ with i have "R j = subresultant j G1 G2" unfolding R_def by auto from arg_cong[OF this, of ffp, unfolded to_fract_hom.subresultant_hom, folded F1 F2] have id2: "ffp (R j) = subresultant j (F 1) (F 2)" . from i have "3 \ i" "i \ k + 1" by auto note eq_30 = BT_eq_30[OF this lt] let ?c3 = "\l\[3.. l ^ (n (l - 1) - j)" let ?c2 = "\l\[3.. l ^ (n (l - 1) - j) * f (l - 1) ^ (\ (l - 2) + \ (l - 1)) * (- 1) ^ ((n (l - 2) - j) * (n (l - 1) - j))" have "?c3 \ 0" by (rule nonzero_alphaprod, insert i, auto) with arg_cong[OF eq_30, of "smult (inverse ?c3)"] have id3: "subresultant j (F 1) (F 2) = smult (inverse ?c3 * ?c2) (subresultant j (F (i - 2)) (F (i - 1)))" by (auto intro: poly_eqI) have "ffp (R j) = 0" unfolding id1 id2 id3 by simp thus ?thesis by simp qed (* equation 27 in BT *) lemma fundamental_theorem_eq_6: assumes "3 \ i" "i \ k" shows "ffp (R (n i)) = smult (\ i) (F i)" (is "?lhs=?rhs") proof - from assms have i1:"1 \ i" by auto from assms have nlt:"i \ k + 1" "n i < n (i - 1)" using n_gt[of "i - 1"] by auto from assms have \nz:"\ i ^ \ (i - 1) \ 0" using \0 by auto have *:"\ a f b. a \ 0 \ smult a f = b \ f = smult (inverse (a::'a fract)) b" by auto have **:"\ f g xs c. c * prod_list (map f xs) * prod_list (map g xs) = c * (\x\xs. f x * (g:: _ \ (_ :: comm_monoid_mult)) x)" by (auto simp:ac_simps prod_list_multf) have ***:"\ c. \ i ^ \ (i - Suc 0) * (inverse (\ i ^ \ (i - Suc 0)) * c) = (\ i / \ i) ^ \ (i - 1) * c" by (auto simp:inverse_eq_divide power_divide) have ****:"int (n (i - Suc 0) - n i) - 1 = int (n (i - 1) - Suc (n i))" using assms nlt by auto from assms n_ge[of "i-2"] nlt n_ge[of i] have nge:"n (i - Suc 0) \ n (i - 2)" "n i < n (i - Suc 0)" "n i < n (i - 1)" "Suc (i - 2) = i - 1" by (cases i,auto simp:numeral_2_eq_2 numeral_3_eq_3) have *****:"(- 1 :: 'a fract) ^ ((n (i - Suc 0) - n i) * (n (i - Suc 0) - n i + (n (i - 2) - n (Suc (i - 2))))) = (- 1) ^ ((n i + n (i - Suc 0)) * (n i + n (i - 2)))" "(- 1 :: 'a fract) ^ (\l\[3..l\[3.. n (x - Suc 0)" "n (x - 2) \ n i" by auto with 1 show ?case by auto qed have "ffp (R (n i)) = subresultant (n i) (F 1) (F 2)" unfolding R_def F1 F2 by (auto simp: to_fract_hom.subresultant_hom ni2[OF assms]) also have "\ = smult ((- 1) ^ (\l\[3..x\[3.. x / \ x) ^ (n (x - 1) - n i) * f (x - 1) ^ (\ (x - 1) + \ (x - 2))) * (((\ i / \ i) ^ \ (i - 1)) * f (i - 1) ^ (\ (i - 1) + \ (i - 2))) * ((- 1) ^ ((\ (i - 2) + \ (i - 1)) * \ (i - 1)) * f i ^ (\ (i - 1) - 1) )) (F i)" unfolding BT_eq_30'[OF assms(1) nlt] ** *[OF \nz BT_lemma_2_22[OF assms]] smult_smult by (auto simp:ac_simps *** ) also have "\ = ?rhs" unfolding \_def \_def using prod_combine[OF assms(1)] \ assms by (auto simp:ac_simps exp_pow_int[symmetric] power_add ***** ****) finally show ?thesis. qed (* equation 26 in BT *) lemma fundamental_theorem_eq_7: assumes j: "j < n k" shows "R j = 0" proof - let ?P = "pseudo_divmod (F (k - 1)) (F k)" from F0[of k] k2 have Fk: "F k \ 0" by auto from pmod[of "Suc k"] k2 F0[of "Suc k"] have "pseudo_mod (F (k - 1)) (F k) = 0" by auto then obtain Q where "?P = (Q,0)" unfolding pseudo_mod_def by (cases ?P, auto) from pseudo_divmod(1)[OF Fk this] Fk obtain c where id: "smult c (F (k - 1)) = F k * Q" and c: "c \ 0" by auto from id have id: "smult c (F (k - 1)) = Q * F k" by auto from n_ge[unfolded n, of "k - 1"] k2 c have "degree (F k) \ degree (smult c (F (k - 1)))" by auto from subresultant_product[OF id this, unfolded subresultant_smult_left[OF c], of j] j have *:"subresultant j (F (k + 1 - 2)) (F (k + 1 - 1)) = 0" using c unfolding n by simp from assms have **:"j \ n 2" by (meson k2 n_ge_trans not_le one_le_numeral order_refl) from k2 assms have "3 \ k + 1" "k + 1 \ k + 1" "j < n (k + 1 - 1)" by auto from BT_eq_30[OF this,unfolded *] nonzero_alphaprod[OF le_refl] ** F1 F2 show ?thesis by (auto simp:R_def F0 to_fract_hom.subresultant_hom[symmetric]) qed definition "G i = R (n (i - 1) - 1)" definition "H i = R (n i)" lemma gamma_delta_beta_3: "\ 3 = (- 1) ^ (\ 1 + 1) * \ 3" proof - have "\ 3 = (- 1) ^ \ 3 * pow_int (f 2) (1 - int (\ 2)) * (\ 3 / (f 2 ^ Suc (\ 1)) * f 2 ^ (\ 1 + \ 2))" unfolding \_def \ \_def by (simp add: \) also have "f 2 ^ (\ 1 + \ 2) = pow_int (f 2) (int (\ 1 + \ 2))" unfolding pow_int_def nat_int by auto also have "int (\ 1 + \ 2) = int (Suc (\ 1)) + (int (\ 2) - 1)" by simp also have "pow_int (f 2) \ = pow_int (f 2) (Suc (\ 1)) * pow_int (f 2) (int (\ 2) - 1)" by (rule pow_int_add, insert f20, auto) also have "pow_int (f 2) (Suc (\ 1)) = f 2 ^ (Suc (\ 1))" unfolding pow_int_def nat_int by simp also have "\ 3 / (f 2 ^ Suc (\ 1)) * (f 2 ^ Suc (\ 1) * pow_int (f 2) (int (\ 2) - 1)) = (\ 3 / (f 2 ^ Suc (\ 1)) * f 2 ^ Suc (\ 1) * pow_int (f 2) (int (\ 2) - 1))" by simp also have "\ 3 / (f 2 ^ Suc (\ 1)) * f 2 ^ Suc (\ 1) = \ 3" using f20 by auto finally have "\ 3 = ((- 1) ^ \ 3 * \ 3) * (pow_int (f 2) (1 - int (\ 2)) * pow_int (f 2) (int (\ 2) - 1))" by simp also have "pow_int (f 2) (1 - int (\ 2)) * pow_int (f 2) (int (\ 2) - 1) = 1" by (subst pow_int_add[symmetric], insert f20, auto) finally have "\ 3 = (- 1) ^ \ 3 * \ 3" by simp also have "\ 3 = (n 1 + n 2 + 1) * (n 2 + n 2 + 1)" unfolding \_def by simp also have "(- (1 :: 'a fract)) ^ \ = (- 1) ^ (n 1 - n 2 + 1)" by (rule minus_1_even_eqI, insert n12, auto) also have "\ = (- 1)^(\ 1 + 1)" unfolding \ by (simp add: numeral_2_eq_2) finally show "\ 3 = (- 1) ^ (\ 1 + 1) * \ 3" . qed fun h :: "nat \ 'a fract" where "h i = (if (i \ 1) then 1 else if i = 2 then (f 2 ^ \ 1) else (f i ^ \ (i - 1) / (h (i - 1) ^ (\ (i - 1) - 1))))" lemma smult_inverse_sdiv_poly: assumes ffp: "p \ range ffp" and p: "p = smult (inverse x) q" and p': "p' = sdiv_poly q' x'" and xx: "x = ff x'" and qq: "q = ffp q'" shows "p = ffp p'" proof (rule poly_eqI) fix i have "coeff p i = coeff q i / x" unfolding p by (simp add: field_simps) also have "\ = ff (coeff q' i) / ff x'" unfolding qq xx by simp finally have cpi: "coeff p i = ff (coeff q' i) / ff x'" . from ffp obtain r where pr: "p = ffp r" by auto from arg_cong[OF this, of "\ p. coeff p i", unfolded cpi] have "ff (coeff q' i) / ff x' \ range ff" by auto hence id: "ff (coeff q' i) / ff x' = ff (coeff q' i div x')" by (rule div_divide_to_fract, auto) show "coeff p i = coeff (ffp p') i" unfolding cpi id p' by (simp add: sdiv_poly_def coeff_map_poly) qed end locale subresultant_prs_locale2 = subresultant_prs_locale F n \ f k \ G1 G2 for F :: "nat \ 'a :: idom_divide fract poly" and n :: "nat \ nat" and \ :: "nat \ nat" and f :: "nat \ 'a fract" and k :: nat and \ :: "nat \ 'a fract" and G1 G2 :: "'a poly" + assumes \3: "\ 3 = (-1)^(\ 1 + 1)" and \i: "\ i. 4 \ i \ i \ Suc k \ \ i = (-1)^(\ (i - 2) + 1) * f (i - 2) * h (i - 2) ^ (\ (i - 2))" begin lemma B_eq_17_main: "2 \ i \ i \ k \ h i = (-1) ^ (n 1 + n i + i + 1) / f i * (\l\[3..< Suc (Suc i)]. (\ l / \ l)) \ h i \ 0" proof (induct i rule: less_induct) case (less i) from less(2-) have fi0: "f i \ 0" using f0[of i] by simp have 1: "(- 1) \ (0 :: 'a fract)" by simp show ?case (is "h i = ?r i \ _") proof (cases "i = 2") case True have f20: "f 2 \ 0" using f20 by auto have hi: "h i = f 2 ^ \ 1" unfolding True h.simps[of 2] by simp have id: "int (\ 1) = int (n 1) - int (n 2)" using n12 unfolding \ numeral_2_eq_2 by simp have "?r i = (- 1) ^ (1 + n 1 + n 2) * ((f 2 ^ Suc (\ 1)) / (\ 3)) / pow_int (f 2) 1" unfolding True \_def by simp also have "\ 3 = (- 1) ^ (\ 1 + 1)" by (rule \3) also have "f 2 ^ Suc (\ 1) / \ = \ * f 2 ^ Suc (\ 1)" by simp finally have "?r i = ((- 1) ^ (1 + n 1 + n 2) * ((- 1) ^ (\ 1 + 1))) * pow_int (f 2) (int (Suc (\ 1)) + (-1))" (is "_ = ?a * _") unfolding pow_int_divide exp_pow_int power_add pow_int_add[OF f20] by (simp add: ac_simps pow_int_add) also have "?a = (-1)^(1 + n 1 + n 2 + \ 1 + 1)" unfolding power_add by simp also have "\ = (-1)^0" by (rule minus_1_even_eqI, insert n12, auto simp: \ numeral_2_eq_2, presburger) finally have ri: "?r i = pow_int (f 2) (int (\ 1))" by simp show ?thesis unfolding ri hi exp_pow_int[symmetric] using f20 by simp next case False hence i: "i \ 3" and ii: "i - 1 < i" "2 \ i - 1" "i - 1 \ k" using less(2-) by auto from i less(2-) have cc: "4 \ Suc i" "Suc i \ Suc k" by auto define P where "P = (\l\[3..< Suc i]. \ l / \ l)" define Q where "Q = P * pow_int (h (i - 1)) (- int (\ (i - 1)))" define R where "R = f i ^ \ (i - 1)" define S where "S = pow_int (f (i - 1)) (- 1)" note IH = less(1)[OF ii] hence hi0: "h (i - 1) \ 0" by auto have hii: "h i = f i ^ \ (i - 1) / h (i - 1) ^ (\ (i - 1) - 1)" unfolding h.simps[of i] using i by simp also have "\ = f i ^ \ (i - 1) * pow_int (h (i - 1)) (- int (\ (i - 1) - 1))" unfolding exp_pow_int pow_int_divide by simp also have "int (\ (i - 1) - 1) = int (\ (i - 1)) - 1" proof - have "\ (i - 1) > 0" unfolding \[of "i - 1"] using n_gt[OF ii(2)] less(2-) by auto thus ?thesis by simp qed also have "- (int (\ (i - 1)) - 1) = 1 + (- int (\ (i - 1)))" by simp finally have hi: "h i = (- 1) ^ (n 1 + n (i - 1) + i) * (R * Q * S)" unfolding pow_int_add[OF hi0] P_def Q_def pow_int_divide[symmetric] R_def S_def using IH i by (simp add: ac_simps) from i have id: "[3.. (Suc i) / \ (Suc i)" unfolding pow_int_divide[symmetric] P_def id Fract_conv_to_fract by simp also have "\ (Suc i) = (- 1) ^ (\ (i - 1) + 1) * f (i - 1) * h (i - 1) ^ \ (i - 1)" using \i[OF cc] by simp also have "\ (Suc i) = f i ^ Suc (\ (i - 1))" unfolding \_def by simp finally have "?r i = (- 1) ^ (n 1 + n i + i + 1) * pow_int (f i) (- 1) * P * (f i ^ Suc (\ (i - 1))) / (- 1) ^ (\ (i - 1) + 1) * pow_int (f (i - 1)) (- 1) / h (i - 1) ^ \ (i - 1)" (is "_ = ?a1 * ?fi1 * P * ?fi2 / ?a2 * ?b / ?c") unfolding exp_pow_int pow_int_divide[symmetric] by simp also have "\ = (?a1 / ?a2) * (?fi1 * ?fi2) * (P / ?c) * ?b" by (simp add: ac_simps) also have "?a1 / ?a2 = (- 1) ^ (n 1 + n i + i + 1 + \ (i - 1) + 1)" by (simp add: power_add) also have "\ = (-1) ^ (n 1 + n i + i + \ (i - 1))" by (rule minus_1_even_eqI, auto) also have "n 1 + n i + i + \ (i - 1) = n 1 + n (i - 1) + i" unfolding \ using i less(2-) n_ge[of "i - 1"] by simp also have "?fi1 * ?fi2 = pow_int (f i) (-1 + int (Suc (\ (i - 1))))" unfolding exp_pow_int pow_int_add[OF fi0] by simp also have "\ = pow_int (f i) (int (\ (i - 1)))" by simp also have "P / ?c = Q" unfolding Q_def exp_pow_int pow_int_divide by simp also have "?b = S" unfolding S_def by simp finally have ri: "?r i = (-1)^(n 1 + n (i - 1) + i) * (R * Q * S)" by (simp add: exp_pow_int R_def) have id: "h i = ?r i" unfolding hi ri .. show ?thesis by (rule conjI[OF id], unfold hii, insert IH fi0, auto) qed qed lemma B_eq_17: "2 \ i \ i \ k \ h i = (-1) ^ (n 1 + n i + i + 1) / f i * (\l\[3..< Suc (Suc i)]. (\ l / \ l))" using B_eq_17_main by blast lemma B_theorem_2: "3 \ i \ i \ Suc k \ \ i = 1" proof (induct i rule: less_induct) case (less i) show ?case proof (cases "i = 3") case True show ?thesis unfolding True unfolding gamma_delta_beta_3 \3 by simp next case False with less(2-) have i: "i \ 4" and ii: "i - 1 < i" "3 \ i - 1" "i - 1 \ Suc k" and iii: "4 \ i" "i \ Suc k" and iv: "2 \ i - 2" "i - 2 \ k" by auto from less(1)[OF ii] have IH: "\ (i - 1) = 1" . define L where "L = [3..< i]" have id: "[3.. l. \ l / \ l)" define A where "A = (\ l. \ l / \ l)" define Q where "Q = (\ l. f (l - 1) ^ (\ (l - 2) + \ (l - 1)))" define R where "R = (\ i l. B l ^ (n (l - 1) - n (i - 1) + 1))" define P where "P = (\ i l. R i l * Q l)" have fi0: "f (i - 1) \ 0" using f0[of "i - 1"] less(2-) by auto have fi0': "f (i - 2) \ 0" using f0[of "i - 2"] less(2-) by auto { fix j assume "j \ set L" hence "j \ 3" "j < i" unfolding L_def by auto with less(3) have j: "j - 1 \ 0" "j - 1 < k" by auto hence Q: "Q j \ 0" unfolding Q_def using f0[of "j - 1"] by auto from j \0 \0[of j] have 0: "\ j \ 0" "\ j \ 0" by auto hence "B j \ 0" "A j \ 0" unfolding B_def A_def by auto note Q this } note L0 = this let ?exp = "\ (i - 2)" have "\ i = \ i / \ (i - 1)" unfolding IH by simp also have "\ = (- 1) ^ \ i * pow_int (f (i - 1)) (1 - int (\ (i - 1))) * (\l\L. P i l) * P i i / ((- 1) ^ \ (i - 1) * pow_int (f (i - 2)) (1 - int (\ (i - 2))) * (\l\L. P (i - 1) l))" (is "_ = ?a1 * ?f1 * ?L1 * P i i / (?a2 * ?f2 * ?L2)") unfolding \_def id P_def Q_def R_def B_def by (simp add: numeral_2_eq_2) also have "\ = (?a1 * ?a2) * (?f1 * P i i) / ?f2 * (?L1 / ?L2)" unfolding divide_prod_assoc by simp also have "?a1 * ?a2 = (-1)^(\ i + \ (i - 1))" (is "_ = ?a") unfolding power_add by simp also have "?L1 / ?L2 = (\l\L. R i l) / (\l\L. R (i - 1) l) * ((\l\L. Q l) / (\l\L. Q l))" unfolding P_def prod_list_multf divide_prod_assoc by simp also have "\ = (\l\L. R i l) / (\l\L. R (i - 1) l)" (is "_ = ?L1 / ?L2") proof - have "(\l\L. Q l) \ 0" unfolding prod_list_zero_iff using L0 by auto thus ?thesis by simp qed also have "?f1 * P i i = (?f1 * pow_int (f (i - 1)) (int ?exp + int (\ (i - 1)))) * R i i" unfolding P_def Q_def exp_pow_int by simp also have "?f1 * pow_int (f (i - 1)) (int ?exp + \ (i - 1)) = pow_int (f (i - 1)) (1 + int ?exp)" (is "_ = ?f1") unfolding pow_int_add[OF fi0, symmetric] by simp also have "R i i = \ i / \ i" unfolding B_def R_def Fract_conv_to_fract by simp also have "\ i = f (i - 1) ^ Suc ?exp" unfolding \_def by simp also have "\ i / \ = \ i * pow_int (f (i - 1)) (- 1 - ?exp)" (is "_ = ?\ * ?f12") unfolding exp_pow_int pow_int_divide by simp finally have "\ i = (?a * (?f1 * ?f12)) * ?\ / ?f2 * (?L1 / ?L2)" by simp also have "?a * (?f1 * ?f12) = ?a" unfolding pow_int_add[OF fi0, symmetric] by simp also have "?L1 / ?L2 = pow_int (\l\L. A l) (- ?exp)" proof - have id: "i - 1 - 1 = i - 2" by simp have "set L \ {l. 3 \ l \ l \ k \ l < i}" unfolding L_def using less(3) by auto thus ?thesis unfolding R_def id proof (induct L) case (Cons l L) from Cons(2) have l: "3 \ l" "l \ k" "l < i" and L: "set L \ {l. 3 \ l \ l \ k \ l < i}" by auto note IH = Cons(1)[OF L] from l \0 \0[of l] have 0: "\ l \ 0" "\ l \ 0" by auto hence B0: "B l \ 0" unfolding B_def by auto have "(\l\l # L. B l ^ (n (l - 1) - n (i - 1) + 1)) / (\l\l # L. B l ^ (n (l - 1) - n (i - 2) + 1)) = (B l ^ (n (l - 1) - n (i - 1) + 1) * (\l\L. B l ^ (n (l - 1) - n (i - 1) + 1))) / (B l ^ (n (l - 1) - n (i - 2) + 1) * (\l\L. B l ^ (n (l - 1) - n (i - 2) + 1)))" (is "_ = (?l1 * ?L1) / (?l2 * ?L2)") by simp also have "\ = (?l1 / ?l2) * (?L1 / ?L2)" by simp also have "?L1 / ?L2 = pow_int (prod_list (map A L)) (- int (\ (i - 2)))" by (rule IH) also have "?l1 / ?l2 = pow_int (B l) (int (n (l - 1) - n (i - 1)) - int (n (l - 1) - n (i - 2)))" unfolding exp_pow_int pow_int_divide pow_int_add[OF B0, symmetric] by simp also have "int (n (l - 1) - n (i - 1)) - int (n (l - 1) - n (i - 2)) = int ?exp" proof - have "n (l - 1) \ n (i - 2)" "n (l - 1) \ n (i - 1)" "n (i - 2) \ n (i - 1)" using i l less(3) by (intro n_ge_trans, auto)+ hence id: "int (n (l - 1) - n (i - 1)) = int (n (l - 1)) - int (n (i - 1))" "int (n (l - 1) - n (i - 2)) = int (n (l - 1)) - int (n (i - 2))" "int (n (i - 2) - n (i - 1)) = int (n (i - 2)) - int (n (i - 1))" by simp_all have id2: "int ?exp = int (n (i - 2) - n (i - 1))" unfolding \ using i by (cases i; cases "i - 1", auto) show ?thesis unfolding id2 unfolding id by simp qed also have "pow_int (B l) \ = pow_int (inverse (B l)) (- \)" unfolding pow_int_def by (cases "int (\ (i - 2))" rule: linorder_cases, auto simp: field_simps) also have "inverse (B l) = A l" unfolding B_def A_def by simp also have "pow_int (A l) (- int ?exp) * pow_int (prod_list (map A L)) (- int ?exp) = pow_int (prod_list (map A (l # L))) (- int ?exp)" by (simp add: pow_int_mult) finally show ?case . qed simp qed also have "\ i = (- 1) ^ (?exp + 1) * f (i - 2) * h (i - 2) ^ ?exp" unfolding \i[OF iii] .. finally have "\ i = (((- 1) ^ (\ i + \ (i - 1)) * (- 1) ^ (?exp + 1))) * (pow_int (f (i - 2)) 1 * pow_int (f (i - 2)) (int ?exp - 1)) * h (i - 2) ^ ?exp / (\l\L. A l) ^ ?exp" (is "_ = ?a * ?f1 * ?H / ?L") unfolding pow_int_divide exp_pow_int by simp also have "?f1 = pow_int (f (i - 2)) (int ?exp)" (is "_ = ?f1") unfolding pow_int_add[OF fi0', symmetric] by simp also have "h (i - 2) = (- 1) ^ (n 1 + n (i - 2) + (i - 2) + 1) / f (i - 2) * (\l\L. A l)" (is "_ = ?a2 / ?f2 * ?L") unfolding B_eq_17[OF iv] A_def id L_def by simp also have "((- (1 :: 'a fract)) ^ (\ i + \ (i - 1)) * (- 1) ^ (?exp + 1)) = ((- 1) ^ (\ i + \ (i - 1) + ?exp + 1))" (is "_ = ?a1") by (simp add: power_add) finally have "\ i = ?a1 * ?f1 * (?a2 / ?f2 * ?L) ^ ?exp / ?L ^ ?exp" by simp also have "\ = (?a1 * ?a2^?exp) * (?f1 / ?f2 ^ ?exp) * (?L^?exp / ?L ^ ?exp)" unfolding power_mult_distrib power_divide by auto also have " ?L ^ ?exp / ?L ^ ?exp = 1" proof - have "?L \ 0" unfolding prod_list_zero_iff using L0 by auto thus ?thesis by simp qed also have "?f1 / ?f2 ^ ?exp = 1" unfolding exp_pow_int pow_int_divide pow_int_add[OF fi0', symmetric] by simp also have "?a2^?exp = (- 1) ^ ((n 1 + n (i - 2) + (i - 2) + 1) * ?exp)" by (rule semiring_normalization_rules) also have "?a1 * \ = (- 1) ^ (\ i + \ (i - 1) + ?exp + 1 + (n 1 + n (i - 2) + (i - 2) + 1) * ?exp)" (is "_ = _ ^ ?e") by (simp add: power_add) also have "\ = (-1)^0" proof - define e where "e = ?e" have *: "?e = (2 * ?exp + \ i + \ (i - 1) + 1 + (n 1 + n (i - 2) + (i - 2)) * ?exp)" by simp define A where "A = (\ i l. (n (l - 2) + n (i - 1) + 1) * (n (l - 1) + n (i - 1) + 1))" define B where "B = (\ i. (n (i - 1) + 1) * (n (i - 1) + 1))" define C where "C = (\ l. (n (l - 1) + n (l - 2) + n (l - 1) * n (l - 2)))" define D where "D = (\ l. n (l - 1) + n (l - 2))" define m2 where "m2 = n (i - 2)" define m1 where "m1 = n (i - 1)" define m0 where "m0 = n 1" define i3 where "i3 = i - 3" have m12: "m2 \ m1" unfolding m2_def m1_def using n_ge[of "i - 2"] i less(3) by (cases i, auto) have idd: "Suc (i - 2) = i - 1" "i - 1 - 1 = i - 2" using i by auto have id4: "i - 2 = Suc i3" unfolding i3_def using i by auto from i have "3 < i" by auto hence "\ k. sum_list (map D L) = n 1 + n (i - 2) + 2 * k" unfolding L_def proof (induct i rule: less_induct) case (less i) show ?case proof (cases "i = 4") case True thus ?thesis by (simp add: D_def) next case False obtain ii where i: "i = Suc ii" and ii: "ii < i" "3 < ii" using False less(2) by (cases i, auto) from less(1)[OF ii] obtain k where IH: "sum_list (map D [3 ..< ii]) = n 1 + n (ii - 2) + 2 * k" by auto have "map D [3 ..< i] = map D [3 ..< ii] @ [D ii]" unfolding i using ii by auto hence "sum_list (map D [3.. = n 1 + n (ii - 1) + 2 * (n (ii - 2) + k)" unfolding D_def by simp also have "n (ii - 1) = n (i - 2)" unfolding i by simp finally show ?thesis by blast qed qed then obtain kk where DL: "sum_list (map D L) = n 1 + n (i - 2) + 2 * kk" .. let ?l = "i - 3" have len: "length L = i - 3" unfolding L_def using i by auto have A: "A i l = B i + D l * n (i - 1) + C l" for i l unfolding A_def B_def C_def D_def ring_distribs by simp have id2: "[3.. = even ((1 + (n 1 + n (i - 2) + (i - 2)) * ?exp) + (\ i + \ (i - 1)))" (is "_ = even (?g + ?j)") unfolding * by (simp add: ac_simps) also have "?j = (\l\L @ [i]. A i l) + (\l\L. A (i - 1) l)" unfolding \_def id A_def by simp also have "\ = 2 * (\l\L. C l) + (Suc ?l) * B i + (\l\L @ [i]. D l * n (i - 1)) + C i + ?l * B (i - 1) + (\l\L. D l * n (i - 1 - 1))" unfolding A sum_list_addf by (simp add: sum_list_triv len) also have "\ = ((Suc ?l * B i + C i + ?l * B (i - 1) + D i * n (i - 1)) + ((\l\L. D l) * (n (i - 1) + n (i - 2)) + 2 * (\l\L. C l)))" (is "_ = ?i + ?j") unfolding sum_list_mult_const by (simp add: ring_distribs numeral_2_eq_2) also have "?j = (n 1 + n (i - 2)) * (n (i - 1) + n (i - 2)) + 2 * (kk * (n (i - 1) + n (i - 2)) + (\l\L. C l))" (is "_ = ?h + 2 * ?f") unfolding DL by (simp add: ring_distribs) finally have "even e = even (?g + ?i + ?h + 2 * ?f)" by presburger also have "\ = even (?g + ?i + ?h)" by presburger also have "?g + ?i + ?h = i3 * (m2 - m1 + m1 * m1 + m2 * m2) + (m2 - m1 + m1 + m2) * (m0 + m2) + (m1 + m2 + (m2 - m1)) + 2 * (m1 * m2 + m1 * m1 + 1 + i3 + m1 * Suc i3 + m2 * i3)" unfolding idd B_def D_def C_def \ m1_def[symmetric] m2_def[symmetric] m0_def[symmetric] unfolding i3_def[symmetric] id4 by (simp add: ring_distribs) also have "(m1 + m2 + (m2 - m1)) = 2 * m2" using m12 by simp also have "(m2 - m1 + m1 + m2) * (m0 + m2) = 2 * (m2 * (m0 + m2))" using m12 by simp finally obtain l1 l2 l3 where "even e = even (i3 * (m2 - m1 + m1 * m1 + m2 * m2) + 2 * l1 + 2 * l2 + 2 * l3)" by blast also have "\ = even (i3 * (m2 - m1 + m1 * m1 + m2 * m2))" by simp also have "\ = even (i3 * (2 * m1 + (m2 - m1 + m1 * m1 + m2 * m2)))" by simp also have "2 * m1 + (m2 - m1 + m1 * m1 + m2 * m2) = m1 + m2 + m1 * m1 + m2 * m2" using m12 by simp also have "even (i3 * \)" by auto finally have "even e" . thus ?thesis unfolding e_def by (intro minus_1_even_eqI, auto) qed finally show "\ i = 1" by simp qed qed context fixes i :: nat assumes i: "3 \ i" "i \ k" begin lemma B_theorem_3_b: "\ i * f i = ff (lead_coeff (H i))" using arg_cong[OF fundamental_theorem_eq_6[folded H_def, OF i], of lead_coeff] unfolding f[of i] lead_coeff_smult by simp lemma B_theorem_3_main: "\ i * f i / \ (i + 1) = (-1)^(n 1 + n i + i + 1) / f i * (\l\[3..< Suc (Suc i)]. (\ l / \ l))" proof (cases "f i = 0") case True thus ?thesis by simp next case False note ff0 = this from i(1) have "Suc (Suc i) > 3" by auto hence id: "[3 ..< Suc (i + 1)] = [3 ..< Suc i] @ [Suc i]" "[3 ..< Suc (Suc i)] = [3 ..< Suc i] @ [Suc i]" by auto have cong: "\ a b c d. a = c \ b = d \ a * b = c * (d :: 'a fract)" by auto define AB where "AB = (\ l. \ l / \ l)" define ABP where "ABP = (\ l. AB l ^ (n (l - 1) - n i) * f (l - 1) ^ (\ (l - 2) + \ (l - 1)))" define PR where "PR = (\l\[3..l\[3.. i * f i / \ (i + 1) = ( ((- 1) ^ \ i * (- 1) ^ \ (i + 1)) * (pow_int (f i) (int (\ (i - 1)) - 1) * PR * f i / pow_int (f i) (1 - int (\ i)) / ((\l\[3.. (i - 1) + \ i))))" unfolding id prod_list.append map_append \_def \_def divide_prod_assoc by (simp add: field_simps power_add AB_def ABP_def PR_def) also have "(- 1 :: 'a fract) ^ \ i * (- 1) ^ \ (i + 1) = (- 1) ^ (\ i + \ (i + 1))" unfolding power_add by (auto simp: field_simps) also have "\ = (- 1) ^ (n 1 + n i + i + 1)" proof (cases "i = 2") case True show ?thesis unfolding \_def \_def True by (auto, rule minus_1_even_eqI, auto) next case False define a where "a = (\ l. n (l - 2) + n i)" define b where "b = (\ l. n (l - 1) + n i)" define c where "c = (\l\[3..l\[3.. i + \ (i + 1)) = ((\l\[3.._def \_def id a_def b_def sum_list_addf by simp also have "(\l\[3..l\[3.. = (\l\[3.. = 2 * c + (\l\[3..l\[3..l\[3..l\[3..l\[3..l\3 # [4.. = n 1 + (\l\[(Suc 3)..l\[(Suc 3)..l\[3..l\[3..l\[3.. i + \ (i + 1) = 2 * d + n (i - 1) + n 1 + length [3.. i + \ (i + 1) = 2 * (d + n (i - 1) + e) + n 1 + (i - 2) + n i + 1" by simp show ?thesis by (rule minus_1_even_eqI, unfold id, insert i, auto) qed also have "(\l\[3.. (i - 1)) - 1) * PR * f i / pow_int (f i) (1 - int (\ i)) / (PR * PR2 * AB (Suc i) * f i ^ (\ (i - 1) + \ i))) = ((pow_int (f i) (int (\ (i - 1)) - 1) * pow_int (f i) 1 * pow_int (f i) (int (\ i) - 1) / pow_int (f i) (int (\ (i - 1) + \ i)))) * (PR / PR / (PR2 * AB (Suc i)))" (is "\ = ?x * ?y") unfolding exp_pow_int[symmetric] by (simp add: pow_int_divide ac_simps) also have "?x = pow_int (f i) (-1)" unfolding pow_int_divide pow_int_add[OF ff0, symmetric] by simp also have "\ = 1 / (f i)" unfolding pow_int_def by simp also have "PR / PR = 1" proof - have "PR \ 0" unfolding PR_def prod_list_zero_iff set_map proof assume "0 \ ABP ` set [3 ..< Suc i]" then obtain j where j: "3 \ j" "j < Suc i" and 0: "ABP j = 0" by auto with i have jk: "j \ k" and j1: "j - 1 \ 0" "j - 1 < k" by auto hence 1: "\ j \ 0" "f (j - 1) \ 0" using \0 f0 by auto with 0 have "AB j = 0" unfolding ABP_def by simp from this[unfolded AB_def] 1(1) \0[of j] show False by auto qed thus ?thesis by simp qed also have "PR2 * AB (Suc i) = (\l\[3.. = inverse \" by (simp add: inverse_eq_divide) also have "\ = (\l\[3.. l / \ l)" unfolding AB_def inverse_prod_list map_map o_def by (auto cong: map_cong) finally show ?thesis by simp qed lemma B_theorem_3: "h i = \ i * f i" "h i = ff (lead_coeff (H i))" proof - have "\ i * f i = \ i * f i / \ (i + 1)" using B_theorem_2[of "i + 1"] i by auto also have "\ = (- 1) ^ (n 1 + n i + i + 1) / f i * (\l\[3.. l / \ l)" by (rule B_theorem_3_main) also have "\ = h i" using B_eq_17[of i] i by simp finally show "h i = \ i * f i" .. thus "h i = ff (lead_coeff (H i))" using B_theorem_3_b by auto qed end lemma h0: "i \ k \ h i \ 0" proof (induct i) case (Suc i) thus ?case unfolding h.simps[of "Suc i"] using f0 by (auto simp del: h.simps) qed auto lemma deg_G12: "degree G1 \ degree G2" using n12 unfolding n F1 F2 by auto lemma R0: shows "R 0 = [: resultant G1 G2 :]" proof(cases "n 2 = 0") case True hence d:"degree G2 = 0" unfolding n F2 by auto from degree0_coeffs[OF d] F2 F12 obtain a where G2: "G2 = [:a:]" and a: "a \ 0" by auto have "sdiv_poly [:a * a ^ degree G1:] a = [:a ^ degree G1:]" using a unfolding sdiv_poly_def by auto note dp = this show ?thesis using G2 F12 unfolding R_def \ n F1 F2 Suc_1 by (auto split:if_splits simp:mult.commute dp) next case False from False n12 have d:"degree G2 \ 0" "degree G2 \ degree G1" unfolding n F2 F1 by auto from False have "R 0 = subresultant 0 G1 G2" unfolding R_def by simp also have "\ = [: resultant G1 G2 :]" unfolding subresultant_resultant by simp finally show ?thesis . qed context fixes div_exp :: "'a \ 'a \ nat \ 'a" assumes div_exp_sound: "div_exp_sound div_exp" begin interpretation div_exp_sound div_exp by (rule div_exp_sound) lemma subresultant_prs_main: assumes "subresultant_prs_main Gi_1 Gi hi_1 = (Gk, hk)" and "F i = ffp Gi" and "F (i - 1) = ffp Gi_1" and "h (i - 1) = ff hi_1" and "i \ 3" "i \ k" shows "F k = ffp Gk \ h k = ff hk \ (\ j. i \ j \ j \ k \ F j \ range ffp \ \ (Suc j) \ range ff)" proof - obtain m where m: "m = k - i" by auto show ?thesis using m assms proof (induct m arbitrary: Gi_1 Gi hi_1 i rule: less_induct) case (less m Gi_1 Gi hi_1 i) note IH = less(1) note m = less(2) note res = less(3) note id = less(4-6) note i = less(7-8) let ?pmod = "pseudo_mod Gi_1 Gi" let ?ni = "degree Gi" let ?ni_1 = "degree Gi_1" let ?gi = "lead_coeff Gi" let ?gi_1 = "lead_coeff Gi_1" let ?d1 = "?ni_1 - ?ni" obtain hi where hi: "hi = div_exp ?gi hi_1 ?d1" by auto obtain divisor where div: "divisor = (-1) ^ (?d1 + 1) * ?gi_1 * (hi_1 ^ ?d1)" by auto obtain G1_p1 where G1_p1: "G1_p1 = sdiv_poly ?pmod divisor" by auto note res = res[unfolded subresultant_prs_main.simps[of Gi_1] Let_def, folded hi, folded div, folded G1_p1] have h_i: "h i = f i ^ \ (i - 1) / h (i - 1) ^ (\ (i - 1) - 1)" unfolding h.simps[of i] using i by simp have hi_ff: "h i \ range ff" using B_theorem_3[OF _ i(2)] i by auto have d1: "\ (i - 1) = ?d1" unfolding \ n using id(1,2) using i by simp have fi: "f i = ff ?gi" unfolding f id by simp have fi1: "f (i - 1) = ff ?gi_1" unfolding f id by simp have eq': "h i = ff (lead_coeff Gi) ^ \ (i - 1) / ff hi_1 ^ (\ (i - 1) - 1)" unfolding h_i fi id .. have idh: "h i = ff hi" using hi_ff h_i fi id unfolding hi d1[symmetric] by (subst div_exp[of ?gi "\ (i - 1)" hi_1], unfold eq'[symmetric], insert assms, blast+) have "\ (Suc i) = (- 1) ^ (\ (i - 1) + 1) * f (i - 1) * h (i - 1) ^ \ (i - 1)" using \i[of "Suc i"] i by auto also have "\ = ff ((- 1) ^ (\ (i - 1) + 1) * lead_coeff Gi_1 * hi_1 ^ \ (i - 1))" unfolding id f by (simp add: hom_distribs) also have "\ \ range ff" by blast finally have beta: "\ (Suc i) \ range ff" . have pm: "pseudo_mod (F (i - 1)) (F i) = ffp ?pmod" unfolding to_fract_hom.pseudo_mod_hom[symmetric] id by simp have eq: "(?pmod = 0) = (i = k)" using pm i pmod[of "Suc i"] F0[of "Suc i"] i \0[of "Suc i"] by auto show ?case proof (cases "i = k") case True with res eq have res: "Gk = Gi" "hk = hi" by auto with pmod have "F k = ffp Gk \ h k = ff hk" unfolding res idh[symmetric] id[symmetric] True by auto thus ?thesis using beta unfolding True by auto next case False with res eq have res: "subresultant_prs_main Gi G1_p1 hi = (Gk, hk)" by auto from m False i have m: "m - 1 < m" "m - 1 = k - Suc i" by auto have si: "Suc i - 1 = i" and ii: "3 \ Suc i" "Suc i \ k" and iii: "3 \ Suc i" "Suc i \ Suc k" using False i by auto have *: "(\j\Suc i. j \ k \ F j \ range ffp \ \ (Suc j) \ range ff) = (\j\i. j \ k \ F j \ range ffp \ \ (Suc j) \ range ff)" by (rule for_all_Suc, insert id(1) beta, auto) show ?thesis proof (rule IH[OF m res, unfolded si, OF _ id(1) idh ii, unfolded *]) have F_ffp: "F (Suc i) \ range ffp" using fundamental_theorem_eq_4[OF ii, symmetric] B_theorem_2[OF iii] by auto from pmod[OF iii] have "smult (\ (Suc i)) (F (Suc i)) = pseudo_mod (F (i - 1)) (F i)" by simp from arg_cong[OF this, of "\ x. smult (inverse (\ (Suc i))) x"] have Fsi: "F (Suc i) = smult (inverse (\ (Suc i))) (pseudo_mod (F (i - 1)) (F i))" using \0[of "Suc i"] by auto show "F (Suc i) = ffp G1_p1" proof (rule smult_inverse_sdiv_poly[OF F_ffp Fsi G1_p1 _ pm]) from i ii have iv: "4 \ Suc i" "Suc i \ Suc k" by auto have *: "Suc i - 2 = i - 1" by auto show "\ (Suc i) = ff divisor" unfolding \i[OF iv] div d1 * fi1 using id by (simp add: hom_distribs) qed qed qed qed qed lemma subresultant_prs: assumes res: "subresultant_prs G1 G2 = (Gk, hk)" shows "F k = ffp Gk \ h k = ff hk \ (i \ 0 \ F i \ range ffp) \ (3 \ i \ i \ Suc k \ \ i \ range ff)" proof - let ?pmod = "pseudo_mod G1 G2" have pm: "pseudo_mod (F 1) (F 2) = ffp ?pmod" unfolding to_fract_hom.pseudo_mod_hom[symmetric] F1 F2 by simp let ?g2 = "lead_coeff G2" let ?n2 = "degree G2" obtain d1 where d1: "d1 = degree G1 - ?n2" by auto obtain h2 where h2: "h2 = ?g2 ^ d1" by auto have "(?pmod = 0) = (pseudo_mod (F 1) (F 2) = 0)" using pm by auto also have "\ = (k < 3)" using k2 pmod[of 3] F0[of 3] \0[of 3] by auto finally have eq: "?pmod = 0 \ k = 2" using k2 by linarith note res = res[unfolded subresultant_prs_def Let_def eq, folded d1, folded h2] have idh2: "h 2 = ff h2" unfolding h2 d1 h.simps[of 2] \ n F1 using F2 by (simp add: numeral_2_eq_2 f hom_distribs) have main: "F k = ffp Gk \ h k = ff hk \ (i \ 3 \ i \ k \ F i \ range ffp \ \ (Suc i) \ range ff)" for i proof (cases "k = 2") case True with res have "Gk = G2" "hk = h2" by auto thus ?thesis using True idh2 F2 by auto next case False hence "(k = 2) = False" by simp note res = res[unfolded this if_False] have F_2: "F (3 - 1) = ffp G2" using F2 by simp have h2: "h (3 - 1) = ff h2" using idh2 by simp have n2: "degree G2 = n (3 - 1)" unfolding n using F2 by simp from False k2 have k3: "3 \ k" by auto have "F k = ffp Gk \ h k = ff hk \ (\j\3. j \ k \ F j \ range ffp \ \ (Suc j) \ range ff)" proof (rule subresultant_prs_main[OF res _ F_2 h2 le_refl k3]) let ?pow = "(- 1) ^ (\ 1 + 1) :: 'a fract" from pmod[of 3] k3 have "smult (\ 3) (F 3) = pseudo_mod (F 1) (F 2)" by simp also have "\ = pseudo_mod (ffp G1) (ffp G2)" using F1 F2 by auto also have "\ = ffp (pseudo_mod G1 G2)" unfolding to_fract_hom.pseudo_mod_hom by simp also have "\ 3 = (- 1) ^ (\ 1 + 1)" unfolding \3 by simp finally have "smult ((- 1) ^ (\ 1 + 1)) (F 3) = ffp (pseudo_mod G1 G2)" by simp also have "smult ((- 1) ^ (\ 1 + 1)) (F 3) = [: ?pow :] * F 3" by simp also have "[: ?pow :] = (- 1) ^ (\ 1 + 1)" by (unfold hom_distribs, simp) finally have "(- 1) ^ (\ 1 + 1) * F 3 = ffp (pseudo_mod G1 G2)" by simp from arg_cong[OF this, of "\ i. (- 1) ^ (\ 1 + 1) * i"] have "F 3 = (- 1) ^ (\ 1 + 1) * ffp (pseudo_mod G1 G2)" by simp also have "\ 1 = d1" unfolding \ n d1 using F1 F2 by (simp add: numeral_2_eq_2) finally show F3: "F 3 = ffp ((- 1) ^ (d1 + 1) * pseudo_mod G1 G2)" by (simp add: hom_distribs) qed thus ?thesis by auto qed show ?thesis proof (intro conjI impI) assume "i \ 0" then consider (12) "i = 1 \ i = 2" | (i3) "i \ 3 \ i \ k" | (ik) "i > k" by linarith thus "F i \ range ffp" proof cases case 12 thus ?thesis using F1 F2 by auto next case i3 thus ?thesis using main by auto next case ik hence "F i = 0" using F0 by auto thus ?thesis by simp qed next assume "3 \ i" and "i \ Suc k" then consider (3) "i = 3" | (4) "3 \ i - 1" "i - 1 \ k" by linarith thus "\ i \ range ff" proof (cases) case 3 have "\ i = ff ((- 1) ^ (\ 1 + 1))" unfolding 3 \3 by (auto simp: hom_distribs) thus ?thesis by blast next case 4 with main[of "i - 1"] show ?thesis by auto qed qed (insert main, auto) qed lemma resultant_impl_main: "resultant_impl_main G1 G2 = resultant G1 G2" proof - from F0[of 2] F12(2) have k2: "k \ 2" by auto obtain Gk hk where sub: "subresultant_prs G1 G2 = (Gk, hk)" by force from subresultant_prs[OF this] have *: "F k = ffp Gk" "h k = ff hk" by auto have "resultant_impl_main G1 G2 = (if degree (F k) = 0 then hk else 0)" unfolding resultant_impl_main_def sub split * using F2 F12 by auto also have "\ = resultant G1 G2" proof (cases "n k = 0") case False with fundamental_theorem_eq_7[of 0] show ?thesis unfolding n[of k] * R0 by auto next case True from H_def[of k, unfolded True] have R: "R 0 = H k" by simp show ?thesis proof (cases "k = 2") case False with k2 have k3: "k \ 3" by auto from B_theorem_3[OF k3] R0 R have "h k = ff (resultant G1 G2)" by simp from this[folded *] * have "hk = resultant G1 G2" by simp with True show ?thesis unfolding n by auto next case 2: True have id: "(if degree (F k) = 0 then hk else 0) = hk" using True unfolding n by simp from F0[of 3, unfolded 2] have "F 3 = 0" by simp with pmod[of 3, unfolded 2] \0[of 3] have "pseudo_mod (F 1) (F 2) = 0" by auto hence pm: "pseudo_mod G1 G2 = 0" unfolding F1 F2 to_fract_hom.pseudo_mod_hom by simp from subresultant_prs_def[of G1 G2, unfolded sub Let_def this] have id: "Gk = G2" "hk = lead_coeff G2 ^ (degree G1 - degree G2)" by auto from F12 F1 F2 have "G1 \ 0" "G2 \ 0" by auto from resultant_pseudo_mod_0[OF pm deg_G12 this] have res: "resultant G1 G2 = (if degree G2 = 0 then lead_coeff G2 ^ degree G1 else 0)" by simp from True[unfolded 2 n F2] have "degree G2 = 0" by simp thus ?thesis unfolding res 2 F2 id by simp qed qed finally show ?thesis . qed end end text \At this point, we have soundness of the resultant-implementation, provided that we can instantiate the locale by constructing suitable values of F, b, h, etc. Now we show the existence of suitable locale parameters by constructively computing them.\ context fixes G1 G2 :: "'a :: idom_divide poly" begin private function F and b and h where "F i = (if i = (0 :: nat) then 1 else if i = 1 then map_poly to_fract G1 else if i = 2 then map_poly to_fract G2 else (let G = pseudo_mod (F (i - 2)) (F (i - 1)) in if F (i - 1) = 0 \ G = 0 then 0 else smult (inverse (b i)) G))" | "b i = (if i \ 2 then 1 else if i = 3 then (- 1) ^ (degree (F 1) - degree (F 2) + 1) else if F (i - 2) = 0 then 1 else (- 1) ^ (degree (F (i - 2)) - degree (F (i - 1)) + 1) * lead_coeff (F (i - 2)) * h (i - 2) ^ (degree (F (i - 2)) - degree (F (i - 1))))" | "h i = (if (i \ 1) then 1 else if i = 2 then (lead_coeff (F 2) ^ (degree (F 1) - degree (F 2))) else if F i = 0 then 1 else (lead_coeff (F i) ^ (degree (F (i - 1)) - degree (F i)) / (h (i - 1) ^ ((degree (F (i - 1)) - degree (F i)) - 1))))" by pat_completeness auto termination proof show "wf (measure (case_sum (\ fi. 3 * fi +1) (case_sum (\ bi. 3 * bi) (\ hi. 3 * hi + 2))))" by simp qed (auto simp: termination_simp) declare h.simps[simp del] b.simps[simp del] F.simps[simp del] private lemma Fb0: assumes base: "G1 \ 0" "G2 \ 0" shows "(F i = 0 \ F (Suc i) = 0) \ b i \ 0 \ h i \ 0" proof (induct i rule: less_induct) case (less i) note * [simp] = F.simps[of i] b.simps[of i] h.simps[of i] consider (0) "i = 0" | (1) "i = 1" | (2) "i \ 2" by linarith thus ?case proof cases case 0 show ?thesis unfolding * unfolding 0 by simp next case 1 show ?thesis unfolding * unfolding 1 using assms by simp next case 2 have F: "F i = 0 \ F (Suc i) = 0" unfolding F.simps[of "Suc i"] using 2 by simp from assms have F2: "F 2 \ 0" unfolding F.simps[of 2] by simp from 2 have "i - 1 < i" "i - 2 < i" by auto note IH = less[OF this(1)] less[OF this(2)] hence b: "b (i - 1) \ 0" and h: "h (i - 1) \ 0" "h (i - 2) \ 0" by auto from h have hi: "h i \ 0" unfolding h.simps[of i] using 2 F2 by auto have bi: "b i \ 0" unfolding b.simps[of i] using h(2) by auto show ?thesis using hi bi F by blast qed qed private definition "k = (LEAST i. F (Suc i) = 0)" private lemma k_exists: "\ i. F (Suc i) = 0" proof - obtain n i where "i \ 3" "length (coeffs (F (Suc i))) = n" by blast thus ?thesis proof (induct n arbitrary: i rule: less_induct) case (less n i) let ?ii = "Suc (Suc i)" let ?i = "Suc i" from less(2) have i: "?i \ 3" by auto let ?mod = "pseudo_mod (F (?ii - 2)) (F ?i)" have Fi: "F ?ii = (if F ?i = 0 \ ?mod = 0 then 0 else smult (inverse (b ?ii)) ?mod)" unfolding F.simps[of ?ii] using i by auto show ?case proof (cases "F ?ii = 0") case False hence Fi: "F ?ii = smult (inverse (b ?ii)) ?mod" and mod: "?mod \ 0" and Fi1: "F ?i \ 0" unfolding Fi by auto from pseudo_mod[OF Fi1, of "F (?ii - 2)"] mod have "degree ?mod < degree (F ?i)" by simp hence deg: "degree (F ?ii) < degree (F ?i)" unfolding Fi by auto hence "length (coeffs (F ?ii)) < length (coeffs (F ?i))" unfolding degree_eq_length_coeffs by auto from less(1)[OF _ i refl, folded less(3), OF this] show ?thesis by auto qed blast qed qed private lemma k: "F (Suc k) = 0" "i < k \ F (Suc i) \ 0" proof - show "F (Suc k) = 0" unfolding k_def using k_exists by (rule LeastI2_ex) assume "i < k" from not_less_Least[OF this[unfolded k_def]] show "F (Suc i) \ 0" . qed lemma enter_subresultant_prs: assumes len: "length (coeffs G1) \ length (coeffs G2)" and G2: "G2 \ 0" shows "\ F n d f k b. subresultant_prs_locale2 F n d f k b G1 G2" proof (intro exI) from G2 len have G1: "G1 \ 0" by auto from len have deg_le: "degree (F 2) \ degree (F 1)" by (simp add: F.simps degree_eq_length_coeffs) from G2 G1 have F1: "F 1 \ 0" and F2: "F 2 \ 0" by (auto simp: F.simps) note Fb0 = Fb0[OF G1 G2] interpret s: subresultant_prs_locale F "\ i. degree (F i)" "\ i. degree (F i) - degree (F (Suc i))" "\ i. lead_coeff (F i)" k b G1 G2 proof (unfold_locales, rule refl, rule refl, rule refl, rule deg_le, rule F1, rule F2) from k(1) F1 have k0: "k \ 0" by (cases k, auto) show Fk: "(F i = 0) = (k < i)" for i proof assume "F i = 0" with k(2)[of "i - 1"] have "\ (i - 1 < k)" by (cases i, auto simp: F.simps) thus "i > k" using k0 by auto next assume "i > k" then obtain j l where i: "i = j + l" and "j = Suc k" and "l = i - Suc k" and Fj: "F j = 0" using k(1) by auto with F1 F2 k0 have j2: "j \ 2" by auto show "F i = 0" unfolding i proof (induct l) case (Suc l) thus ?case unfolding F.simps[of "j + Suc l"] using j2 by auto qed (auto simp: Fj) qed show b: "b i \ 0" for i using Fb0 by blast show "F 1 = map_poly to_fract G1" unfolding F.simps[of 1] by simp show "F 2 = map_poly to_fract G2" unfolding F.simps[of 2] by simp fix i let ?mod = "pseudo_mod (F (i - 2)) (F (i - 1))" assume i: "3 \ i" "i \ Suc k" from Fk[of "i - 1"] i have "F (i - 1) \ 0" by auto with i have Fi: "F i = (if ?mod = 0 then 0 else smult (inverse (b i)) ?mod)" unfolding F.simps[of i] Let_def by simp show "smult (b i) (F i) = ?mod" proof (cases "?mod = 0") case True thus ?thesis unfolding Fi by simp next case False with Fi have Fi: "F i = smult (inverse (b i)) ?mod" by simp from arg_cong[OF this, of "smult (b i)"] b[of i] show ?thesis by simp qed qed note s.h.simps[simp del] show "subresultant_prs_locale2 F (\ i. degree (F i)) (\ i. degree (F i) - degree (F (Suc i))) (\ i. lead_coeff (F i)) k b G1 G2" proof show "b 3 = (- 1) ^ (degree (F 1) - degree (F (Suc 1)) + 1)" unfolding b.simps numeral_2_eq_2 by simp fix i assume i: "4 \ i" "i \ Suc k" with s.F0[of "i - 2"] have "F (i - 2) \ 0" by auto hence bi: "b i = (- 1) ^ (degree (F (i - 2)) - degree (F (i - 1)) + 1) * lead_coeff (F (i - 2)) * h (i - 2) ^ (degree (F (i - 2)) - degree (F (i - 1)))" unfolding b.simps using i by auto have "i < k \ s.h i = h i" for i proof (induct i) case 0 thus ?case by (simp add: h.simps s.h.simps) next case (Suc i) from Suc(2) s.F0[of "Suc i"] have "F (Suc i) \ 0" by auto with Suc show ?case unfolding h.simps[of "Suc i"] s.h.simps[of "Suc i"] numeral_2_eq_2 by simp qed hence sh: "s.h (i - 2) = h (i - 2)" using i by simp from i have *: "Suc (i - 2) = i - 1" by auto show "b i = (- 1) ^ (degree (F (i - 2)) - degree (F (Suc (i - 2))) + 1) * lead_coeff (F (i - 2)) * s.h (i - 2) ^ (degree (F (i - 2)) - degree (F (Suc (i - 2))))" unfolding sh bi * .. qed qed end text \Now we obtain the soundness lemma outside the locale.\ context div_exp_sound begin lemma resultant_impl_main: assumes len: "length (coeffs G1) \ length (coeffs G2)" shows "resultant_impl_main G1 G2 = resultant G1 G2" proof (cases "G2 = 0") case G2: False from enter_subresultant_prs[OF len G2] obtain F n d f k b where "subresultant_prs_locale2 F n d f k b G1 G2" by auto interpret subresultant_prs_locale2 F n d f k b G1 G2 by fact show ?thesis by (rule resultant_impl_main, standard) next case G2: True show ?thesis unfolding G2 resultant_impl_main_def using resultant_const(2)[of G1 0] by simp qed theorem resultant_impl: "resultant_impl = resultant" proof (intro ext) fix f g :: "'a poly" show "resultant_impl f g = resultant f g" proof (cases "length (coeffs f) \ length (coeffs g)") case True thus ?thesis unfolding resultant_impl_def resultant_impl_main[OF True] by auto next case False hence "length (coeffs g) \ length (coeffs f)" by auto from resultant_impl_main[OF this] show ?thesis unfolding resultant_impl_def resultant_swap[of f g] using False by (auto simp: Let_def) qed qed end subsection \Code Equations\ text \In the following code-equations, we only compute the required values, e.g., $h_k$ is not required if $n_k > 0$, we compute $(-1)^{\ldots} * \ldots$ via a case-analysis, and we perform special cases for $\delta_i = 1$, which is the most frequent case.\ context div_exp_param begin partial_function(tailrec) subresultant_prs_main_impl where "subresultant_prs_main_impl f Gi_1 Gi ni_1 d1_1 hi_2 = (let gi_1 = lead_coeff Gi_1; ni = degree Gi; hi_1 = (if d1_1 = 1 then gi_1 else div_exp gi_1 hi_2 d1_1); d1 = ni_1 - ni; pmod = pseudo_mod Gi_1 Gi in (if pmod = 0 then f (Gi, (if d1 = 1 then lead_coeff Gi else div_exp (lead_coeff Gi) hi_1 d1)) else let gi = lead_coeff Gi; divisor = (-1) ^ (d1 + 1) * gi_1 * (hi_1 ^ d1) ; Gi_p1 = sdiv_poly pmod divisor in subresultant_prs_main_impl f Gi Gi_p1 ni d1 hi_1))" definition subresultant_prs_impl where "subresultant_prs_impl f G1 G2 = (let pmod = pseudo_mod G1 G2; n2 = degree G2; delta_1 = (degree G1 - n2); g2 = lead_coeff G2; h2 = g2 ^ delta_1 in if pmod = 0 then f (G2,h2) else let G3 = (- 1) ^ (delta_1 + 1) * pmod; g3 = lead_coeff G3; n3 = degree G3; d2 = n2 - n3; pmod = pseudo_mod G2 G3 in if pmod = 0 then f (G3, if d2 = 1 then g3 else div_exp g3 h2 d2) else let divisor = (- 1) ^ (d2 + 1) * g2 * h2 ^ d2; G4 = sdiv_poly pmod divisor in subresultant_prs_main_impl f G3 G4 n3 d2 h2 )" end context div_exp_sound begin lemma div_exp_1: "div_exp g h (Suc 0) = g" using div_exp[of g "Suc 0" h] by simp lemma subresultant_prs_impl: "subresultant_prs_impl f G1 G2 = f (subresultant_prs G1 G2)" proof - define h2 where "h2 = lead_coeff G2 ^ (degree G1 - degree G2)" define G3 where "G3 = ((- 1) ^ (degree G1 - degree G2 + 1) * pseudo_mod G1 G2)" define G4 where "G4 = sdiv_poly (pseudo_mod G2 G3) ((- 1) ^ (degree G2 - degree G3 + 1) * lead_coeff G2 * h2 ^ (degree G2 - degree G3))" define d2 where "d2 = degree G2 - degree G3" have dl1: "(if d = 1 then (g :: 'a) else div_exp g h d) = div_exp g h d" for d g h by (cases "d = 1", auto simp: div_exp_1) show ?thesis unfolding subresultant_prs_impl_def subresultant_prs_def Let_def subresultant_prs_main.simps[of G2] if_distrib[of f] dl1 proof (rule if_cong[OF refl refl if_cong[OF refl refl]], unfold h2_def[symmetric], unfold G3_def[symmetric], unfold G4_def[symmetric], unfold d2_def[symmetric]) note simp = subresultant_prs_main_impl.simps[of f] subresultant_prs_main.simps show "subresultant_prs_main_impl f G3 G4 (degree G3) d2 h2 = f (subresultant_prs_main G3 G4 (div_exp (lead_coeff G3) h2 d2))" proof (induct G4 arbitrary: G3 d2 h2 rule: wf_induct[OF wf_measure[of degree]]) case (1 G4 G3 d2 h2) let ?M = "pseudo_mod G3 G4" show ?case proof (cases "?M = 0") case True thus ?thesis unfolding simp[of G3] Let_def dl1 by simp next case False hence id: "(?M = 0) = False" by auto let ?c = "((- 1) ^ (degree G3 - degree G4 + 1) * lead_coeff G3 * (div_exp (lead_coeff G3) h2 d2) ^ (degree G3 - degree G4))" let ?N = "sdiv_poly ?M ?c" show ?thesis proof (cases "G4 = 0") case G4: False have "degree ?N \ degree ?M" unfolding sdiv_poly_def by (rule degree_map_poly_le) also have "\ < degree G4" using pseudo_mod[OF G4, of G3] False by auto finally show ?thesis unfolding simp[of G3] Let_def id if_False dl1 by (intro 1(1)[rule_format], auto) next case 0: True with False have "G3 \ 0" by auto show ?thesis unfolding 0 unfolding simp[of G3] Let_def unfolding dl1 simp[of 0] by simp qed qed qed qed qed definition "resultant_impl_rec = subresultant_prs_main_impl (\ (Gk,hk). if degree Gk = 0 then hk else 0)" definition "resultant_impl_start = subresultant_prs_impl (\ (Gk,hk). if degree Gk = 0 then hk else 0)" lemma resultant_impl_start_code: "resultant_impl_start G1 G2 = (let pmod = pseudo_mod G1 G2; n2 = degree G2; n1 = degree G1; g2 = lead_coeff G2; d1 = n1 - n2 in if pmod = 0 then if n2 = 0 then if d1 = 0 then 1 else if d1 = 1 then g2 else g2 ^ d1 else 0 else let G3 = if even d1 then - pmod else pmod; n3 = degree G3; pmod = pseudo_mod G2 G3 in if pmod = 0 then if n3 = 0 then let d2 = n2 - n3; g3 = lead_coeff G3 in (if d2 = 1 then g3 else div_exp g3 (if d1 = 1 then g2 else g2 ^ d1) d2) else 0 else let h2 = (if d1 = 1 then g2 else g2 ^ d1); d2 = n2 - n3; divisor = (if d2 = 1 then g2 * h2 else if even d2 then - g2 * h2 ^ d2 else g2 * h2 ^ d2); G4 = sdiv_poly pmod divisor in resultant_impl_rec G3 G4 n3 d2 h2)" proof - obtain d1 where d1: "degree G1 - degree G2 = d1" by auto have id1: "(if even d1 then - pmod else pmod) = (-1)^ (d1 + 1) * (pmod :: 'a poly)" for pmod by simp have id3: "(if d2 = 1 then g2 * h2 else if even d2 then - g2 * h2 ^ d2 else g2 * h2 ^ d2) = ((- 1) ^ (d2 + 1) * g2 * h2 ^ d2)" for d2 and g2 h2 :: 'a by auto show ?thesis unfolding resultant_impl_start_def subresultant_prs_impl_def resultant_impl_rec_def[symmetric] Let_def split unfolding d1 unfolding id1 unfolding id3 by (rule if_cong[OF refl if_cong if_cong], auto simp: power2_eq_square) qed lemma resultant_impl_rec_code: "resultant_impl_rec Gi_1 Gi ni_1 d1_1 hi_2 = ( let ni = degree Gi; pmod = pseudo_mod Gi_1 Gi in if pmod = 0 then if ni = 0 then let d1 = ni_1 - ni; gi = lead_coeff Gi in if d1 = 1 then gi else let gi_1 = lead_coeff Gi_1; hi_1 = (if d1_1 = 1 then gi_1 else div_exp gi_1 hi_2 d1_1) in div_exp gi hi_1 d1 else 0 else let d1 = ni_1 - ni; gi_1 = lead_coeff Gi_1; hi_1 = (if d1_1 = 1 then gi_1 else div_exp gi_1 hi_2 d1_1); divisor = if d1 = 1 then gi_1 * hi_1 else if even d1 then - gi_1 * hi_1 ^ d1 else gi_1 * hi_1 ^ d1; Gi_p1 = sdiv_poly pmod divisor in resultant_impl_rec Gi Gi_p1 ni d1 hi_1)" unfolding resultant_impl_rec_def subresultant_prs_main_impl.simps[of _ Gi_1] split Let_def unfolding resultant_impl_rec_def[symmetric] by (rule if_cong[OF _ if_cong _], auto) lemma resultant_impl_main_code: "resultant_impl_main G1 G2 = (if G2 = 0 then if degree G1 = 0 then 1 else 0 else resultant_impl_start G1 G2)" unfolding resultant_impl_main_def resultant_impl_start_def subresultant_prs_impl by simp lemma resultant_impl_code: "resultant_impl f g = (if length (coeffs f) \ length (coeffs g) then resultant_impl_main f g else let res = resultant_impl_main g f in if even (degree f) \ even (degree g) then res else - res)" unfolding resultant_impl_def resultant_impl_def .. lemma resultant_code: "resultant = resultant_impl" using resultant_impl by fastforce lemmas resultant_code_lemmas = resultant_impl_code resultant_impl_main_code resultant_impl_start_code resultant_impl_rec_code end global_interpretation div_exp_Lazard: div_exp_sound "dichotomous_Lazard :: 'a :: factorial_ring_gcd \ _" defines resultant_impl_Lazard = div_exp_Lazard.resultant_impl and resultant_impl_main_Lazard = div_exp_Lazard.resultant_impl_main and resultant_impl_start_Lazard = div_exp_Lazard.resultant_impl_start and resultant_impl_rec_Lazard = div_exp_Lazard.resultant_impl_rec by (rule dichotomous_Lazard) declare div_exp_Lazard.resultant_code_lemmas[code] text \As default use Lazard-implementation, which implements resultants on factorial rings.\ declare div_exp_Lazard.resultant_code[code] text \We also provide a second implementation without Lazard's optimization, which works on integral domains.\ global_interpretation div_exp_basic: div_exp_sound basic_div_exp defines resultant_impl_basic = div_exp_basic.resultant_impl and resultant_impl_main_basic = div_exp_basic.resultant_impl_main and resultant_impl_start_basic = div_exp_basic.resultant_impl_start and resultant_impl_rec_basic = div_exp_basic.resultant_impl_rec by (rule basic_div_exp) declare div_exp_basic.resultant_code_lemmas[code] thm div_exp_basic.resultant_code end \ No newline at end of file diff --git a/thys/Winding_Number_Eval/Missing_Algebraic.thy b/thys/Winding_Number_Eval/Missing_Algebraic.thy --- a/thys/Winding_Number_Eval/Missing_Algebraic.thy +++ b/thys/Winding_Number_Eval/Missing_Algebraic.thy @@ -1,371 +1,335 @@ (* Author: Wenda Li *) section \Some useful lemmas in algebra\ theory Missing_Algebraic imports "HOL-Computational_Algebra.Polynomial_Factorial" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" "HOL-Complex_Analysis.Complex_Analysis" Missing_Topology "Budan_Fourier.BF_Misc" begin subsection \Misc\ lemma poly_holomorphic_on[simp]: "(poly p) holomorphic_on s" apply (rule holomorphic_onI) apply (unfold field_differentiable_def) apply (rule_tac x="poly (pderiv p) x" in exI) by (simp add:has_field_derivative_at_within) lemma order_zorder: fixes p::"complex poly" and z::complex assumes "p\0" shows "order z p = nat (zorder (poly p) z)" proof - define n where "n=nat (zorder (poly p) z)" define h where "h=zor_poly (poly p) z" have "\w. poly p w \ 0" using assms poly_all_0_iff_0 by auto then obtain r where "0 < r" "cball z r \ UNIV" and h_holo: "h holomorphic_on cball z r" and poly_prod:"(\w\cball z r. poly p w = h w * (w - z) ^ n \ h w \ 0)" using zorder_exist_zero[of "poly p" UNIV z,folded h_def] poly_holomorphic_on unfolding n_def by auto then have "h holomorphic_on ball z r" and "(\w\ball z r. poly p w = h w * (w - z) ^ n)" and "h z\0" by auto then have "order z p = n" using \p\0\ proof (induct n arbitrary:p h) case 0 then have "poly p z=h z" using \r>0\ by auto then have "poly p z\0" using \h z\0\ by auto then show ?case using order_root by blast next case (Suc n) define sn where "sn=Suc n" define h' where "h'\ \w. deriv h w * (w-z)+ sn * h w" have "(poly p has_field_derivative poly (pderiv p) w) (at w)" for w using poly_DERIV[of p w] . moreover have "(poly p has_field_derivative (h' w)*(w-z)^n ) (at w)" when "w\ball z r" for w proof (subst DERIV_cong_ev[of w w "poly p" "\w. h w * (w - z) ^ Suc n" ],simp_all) show "\\<^sub>F x in nhds w. poly p x = h x * ((x - z) * (x - z) ^ n)" unfolding eventually_nhds using Suc(3) \w\ball z r\ apply (intro exI[where x="ball z r"]) by auto next have "(h has_field_derivative deriv h w) (at w)" using \h holomorphic_on ball z r\ \w\ball z r\ holomorphic_on_imp_differentiable_at by (simp add: holomorphic_derivI) then have "((\w. h w * ((w - z) ^ sn)) has_field_derivative h' w * (w - z) ^ (sn - 1)) (at w)" unfolding h'_def apply (auto intro!: derivative_eq_intros simp add:field_simps) by (auto simp add:field_simps sn_def) then show "((\w. h w * ((w - z) * (w - z) ^ n)) has_field_derivative h' w * (w - z) ^ n) (at w)" unfolding sn_def by auto qed ultimately have "\w\ball z r. poly (pderiv p) w = h' w * (w - z) ^ n" using DERIV_unique by blast moreover have "h' holomorphic_on ball z r" unfolding h'_def using \h holomorphic_on ball z r\ by (auto intro!: holomorphic_intros) moreover have "h' z\0" unfolding h'_def sn_def using \h z \ 0\ of_nat_neq_0 by auto moreover have "pderiv p \ 0" proof assume "pderiv p = 0" obtain c where "p=[:c:]" using \pderiv p = 0\ using pderiv_iszero by blast then have "c=0" using Suc(3)[rule_format,of z] \r>0\ by auto then show False using \p\0\ using \p=[:c:]\ by auto qed ultimately have "order z (pderiv p) = n" by (auto elim: Suc.hyps) moreover have "order z p \ 0" using Suc(3)[rule_format,of z] \r>0\ order_root \p\0\ by auto ultimately show ?case using order_pderiv[OF \pderiv p \ 0\] by auto qed then show ?thesis unfolding n_def . qed lemma pcompose_pCons_0:"pcompose p [:a:] = [:poly p a:]" apply (induct p) by (auto simp add:pcompose_pCons algebra_simps) lemma pcompose_coeff_0: "coeff (pcompose p q) 0 = poly p (coeff q 0)" apply (induct p) by (auto simp add:pcompose_pCons coeff_mult) lemma poly_field_differentiable_at[simp]: "poly p field_differentiable (at x within s)" apply (unfold field_differentiable_def) apply (rule_tac x="poly (pderiv p) x" in exI) by (simp add:has_field_derivative_at_within) lemma deriv_pderiv: "deriv (poly p) = poly (pderiv p)" apply (rule ext) apply (rule DERIV_imp_deriv) using poly_DERIV . lemma lead_coeff_map_poly_nz: assumes "f (lead_coeff p) \0" "f 0=0" shows "lead_coeff (map_poly f p) = f (lead_coeff p) " proof - have "lead_coeff (Poly (map f (coeffs p))) = f (lead_coeff p)" by (metis (mono_tags, lifting) antisym assms(1) assms(2) coeff_0_degree_minus_1 coeff_map_poly degree_Poly degree_eq_length_coeffs le_degree length_map map_poly_def) then show ?thesis by (simp add: map_poly_def) qed lemma filterlim_poly_at_infinity: fixes p::"'a::real_normed_field poly" assumes "degree p>0" shows "filterlim (poly p) at_infinity at_infinity" using assms proof (induct p) case 0 then show ?case by auto next case (pCons a p) have ?case when "degree p=0" proof - obtain c where c_def:"p=[:c:]" using \degree p = 0\ degree_eq_zeroE by blast then have "c\0" using \0 < degree (pCons a p)\ by auto then show ?thesis unfolding c_def apply (auto intro!:tendsto_add_filterlim_at_infinity) apply (subst mult.commute) by (auto intro!:tendsto_mult_filterlim_at_infinity filterlim_ident) qed moreover have ?case when "degree p\0" proof - have "filterlim (poly p) at_infinity at_infinity" using that by (auto intro:pCons) then show ?thesis by (auto intro!:tendsto_add_filterlim_at_infinity filterlim_at_infinity_times filterlim_ident) qed ultimately show ?case by auto qed lemma poly_divide_tendsto_aux: fixes p::"'a::real_normed_field poly" shows "((\x. poly p x/x^(degree p)) \ lead_coeff p) at_infinity" proof (induct p) case 0 then show ?case by (auto intro:tendsto_eq_intros) next case (pCons a p) have ?case when "p=0" using that by auto moreover have ?case when "p\0" proof - define g where "g=(\x. a/(x*x^degree p))" define f where "f=(\x. poly p x/x^degree p)" have "\\<^sub>Fx in at_infinity. poly (pCons a p) x / x ^ degree (pCons a p) = g x + f x" proof (rule eventually_at_infinityI[of 1]) fix x::'a assume "norm x\1" then have "x\0" by auto then show "poly (pCons a p) x / x ^ degree (pCons a p) = g x + f x" using that unfolding g_def f_def by (auto simp add:field_simps) qed moreover have "((\x. g x+f x) \ lead_coeff (pCons a p)) at_infinity" proof - have "(g \ 0) at_infinity" unfolding g_def using filterlim_poly_at_infinity[of "monom 1 (Suc (degree p))"] apply (auto intro!:tendsto_intros tendsto_divide_0 simp add: degree_monom_eq) apply (subst filterlim_cong[where g="poly (monom 1 (Suc (degree p)))"]) by (auto simp add:poly_monom) moreover have "(f \ lead_coeff (pCons a p)) at_infinity" using pCons \p\0\ unfolding f_def by auto ultimately show ?thesis by (auto intro:tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed ultimately show ?case by auto qed lemma filterlim_power_at_infinity: assumes "n\0" shows "filterlim (\x::'a::real_normed_field. x^n) at_infinity at_infinity" using filterlim_poly_at_infinity[of "monom 1 n"] assms apply (subst filterlim_cong[where g="poly (monom 1 n)"]) by (auto simp add:poly_monom degree_monom_eq) lemma poly_divide_tendsto_0_at_infinity: fixes p::"'a::real_normed_field poly" assumes "degree p > degree q" shows "((\x. poly q x / poly p x) \ 0 ) at_infinity" proof - define pp where "pp=(\x. x^(degree p) / poly p x)" define qq where "qq=(\x. poly q x/x^(degree q))" define dd where "dd=(\x::'a. 1/x^(degree p - degree q))" have "\\<^sub>Fx in at_infinity. poly q x / poly p x = qq x * pp x * dd x" proof (rule eventually_at_infinityI[of 1]) fix x::'a assume "norm x\1" then have "x\0" by auto then show "poly q x / poly p x = qq x * pp x * dd x" unfolding qq_def pp_def dd_def using assms by (auto simp add:field_simps divide_simps power_diff) qed moreover have "((\x. qq x * pp x * dd x) \ 0) at_infinity" proof - have "(qq \ lead_coeff q) at_infinity" unfolding qq_def using poly_divide_tendsto_aux[of q] . moreover have "(pp \ 1/lead_coeff p) at_infinity" proof - have "p\0" using assms by auto then show ?thesis unfolding pp_def using poly_divide_tendsto_aux[of p] apply (drule_tac tendsto_inverse) by (auto simp add:inverse_eq_divide) qed moreover have "(dd \ 0) at_infinity" unfolding dd_def apply (rule tendsto_divide_0) by (auto intro!: filterlim_power_at_infinity simp add:assms) ultimately show ?thesis by (auto intro:tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed lemma lead_coeff_list_def: "lead_coeff p= (if coeffs p=[] then 0 else last (coeffs p))" by (simp add: last_coeffs_eq_coeff_degree) lemma poly_linepath_comp: fixes a::"'a::{real_normed_vector,comm_semiring_0,real_algebra_1}" shows "poly p o (linepath a b) = poly (p \\<^sub>p [:a, b-a:]) o of_real" apply rule by (auto simp add:poly_pcompose linepath_def scaleR_conv_of_real algebra_simps) lemma poly_eventually_not_zero: fixes p::"real poly" assumes "p\0" shows "eventually (\x. poly p x\0) at_infinity" proof (rule eventually_at_infinityI[of "Max (norm ` {x. poly p x=0}) + 1"]) fix x::real assume asm:"Max (norm ` {x. poly p x=0}) + 1 \ norm x" have False when "poly p x=0" proof - define S where "S=norm `{x. poly p x = 0}" have "norm x\S" using that unfolding S_def by auto moreover have "finite S" using \p\0\ poly_roots_finite unfolding S_def by blast ultimately have "norm x\Max S" by simp moreover have "Max S + 1 \ norm x" using asm unfolding S_def by simp ultimately show False by argo qed then show "poly p x \ 0" by auto qed subsection \More about @{term degree}\ -lemma degree_div_less: - fixes x y::"'a::field poly" - assumes "degree x\0" "degree y\0" - shows "degree (x div y) < degree x" -proof - - have "x\0" "y\0" using assms by auto - define q r where "q=x div y" and "r=x mod y" - have *:"eucl_rel_poly x y (q, r)" unfolding q_def r_def - by (simp add: eucl_rel_poly) - then have "r = 0 \ degree r < degree y" using \y\0\ unfolding eucl_rel_poly_iff by auto - moreover have ?thesis when "r=0" - proof - - have "x = q * y" using * that unfolding eucl_rel_poly_iff by auto - then have "q\0" using \x\0\ \y\0\ by auto - from degree_mult_eq[OF this \y\0\] \x = q * y\ - have "degree x = degree q +degree y" by auto - then show ?thesis unfolding q_def using assms by auto - qed - moreover have ?thesis when "degree rdegree x") - case True - then have "q=0" unfolding q_def using div_poly_less by auto - then show ?thesis unfolding q_def using assms(1) by auto - next - case False - then have "degree x>degree r" using that by auto - then have "degree x = degree (x-r)" using degree_add_eq_right[of "-r" x] by auto - have "x-r = q*y" using * unfolding eucl_rel_poly_iff by auto - then have "q\0" using \degree r < degree x\ by auto - have "degree x = degree q +degree y" - using degree_mult_eq[OF \q\0\ \y\0\] \x-r = q*y\ \degree x = degree (x-r)\ by auto - then show ?thesis unfolding q_def using assms by auto - qed - ultimately show ?thesis by auto -qed - lemma map_poly_degree_eq: assumes "f (lead_coeff p) \0" shows "degree (map_poly f p) = degree p" using assms unfolding map_poly_def degree_eq_length_coeffs coeffs_Poly lead_coeff_list_def by (metis (full_types) last_conv_nth_default length_map no_trailing_unfold nth_default_coeffs_eq nth_default_map_eq strip_while_idem) lemma map_poly_degree_less: assumes "f (lead_coeff p) =0" "degree p\0" shows "degree (map_poly f p) < degree p" proof - have "length (coeffs p) >1" using \degree p\0\ by (simp add: degree_eq_length_coeffs) then obtain xs x where xs_def:"coeffs p=xs@[x]" "length xs>0" by (metis One_nat_def add.commute add_diff_cancel_left' append_Nil assms(2) degree_eq_length_coeffs length_greater_0_conv list.size(3) list.size(4) not_less_zero rev_exhaust) have "f x=0" using assms(1) by (simp add: lead_coeff_list_def xs_def(1)) have "degree (map_poly f p) = length (strip_while ((=) 0) (map f (xs@[x]))) - 1" unfolding map_poly_def degree_eq_length_coeffs coeffs_Poly by (subst xs_def,auto) also have "... = length (strip_while ((=) 0) (map f xs)) - 1" using \f x=0\ by simp also have "... \ length xs -1" using length_strip_while_le by (metis diff_le_mono length_map) also have "... < length (xs@[x]) - 1" using xs_def(2) by auto also have "... = degree p" unfolding degree_eq_length_coeffs xs_def by simp finally show ?thesis . qed lemma map_poly_degree_leq[simp]: shows "degree (map_poly f p) \ degree p" unfolding map_poly_def degree_eq_length_coeffs by (metis coeffs_Poly diff_le_mono length_map length_strip_while_le) subsection \roots / zeros of a univariate function\ definition roots_within::"('a \ 'b::zero) \ 'a set \ 'a set" where "roots_within f s = {x\s. f x = 0}" abbreviation roots::"('a \ 'b::zero) \ 'a set" where "roots f \ roots_within f UNIV" subsection \The argument principle specialised to polynomials.\ lemma argument_principle_poly: assumes "p\0" and valid:"valid_path g" and loop: "pathfinish g = pathstart g" and no_proots:"path_image g \ - proots p" shows "contour_integral g (\x. deriv (poly p) x / poly p x) = 2 * of_real pi * \ * (\x\proots p. winding_number g x * of_nat (order x p))" proof - have "contour_integral g (\x. deriv (poly p) x / poly p x) = 2 * of_real pi * \ * (\x | poly p x = 0. winding_number g x * of_int (zorder (poly p) x))" apply (rule argument_principle[of UNIV "poly p" "{}" "\_. 1" g,simplified,OF _ valid loop]) using no_proots[unfolded proots_def] by (auto simp add:poly_roots_finite[OF \p\0\] ) also have "... = 2 * of_real pi * \ * (\x\proots p. winding_number g x * of_nat (order x p))" proof - have "nat (zorder (poly p) x) = order x p" when "x\proots p" for x using order_zorder[OF \p\0\] that unfolding proots_def by auto then show ?thesis unfolding proots_def apply (auto intro!:comm_monoid_add_class.sum.cong) by (metis assms(1) nat_eq_iff2 of_nat_nat order_root) qed finally show ?thesis . qed end