diff --git a/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy b/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy --- a/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy +++ b/thys/Complex_Bounded_Operators/Complex_Bounded_Linear_Function.thy @@ -1,2981 +1,2981 @@ section \\Complex_Bounded_Linear_Function\ -- Complex bounded linear functions (bounded operators)\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Complex_Bounded_Linear_Function imports Complex_Inner_Product One_Dimensional_Spaces Banach_Steinhaus.Banach_Steinhaus "HOL-Types_To_Sets.Types_To_Sets" Complex_Bounded_Linear_Function0 begin subsection \Misc basic facts and declarations\ notation cblinfun_apply (infixr "*\<^sub>V" 70) lemma id_cblinfun_apply[simp]: "id_cblinfun *\<^sub>V \ = \" apply transfer by simp lemma isCont_cblinfun_apply[simp]: "isCont ((*\<^sub>V) A) \" apply transfer by (simp add: clinear_continuous_at) declare cblinfun.scaleC_left[simp] lemma cblinfun_apply_clinear[simp]: \clinear (cblinfun_apply A)\ using bounded_clinear.axioms(1) cblinfun_apply by blast lemma cblinfun_cinner_eqI: fixes A B :: \'a::chilbert_space \\<^sub>C\<^sub>L 'a\ assumes \\\. cinner \ (A *\<^sub>V \) = cinner \ (B *\<^sub>V \)\ shows \A = B\ proof - define C where \C = A - B\ have C0[simp]: \cinner \ (C \) = 0\ for \ by (simp add: C_def assms cblinfun.diff_left cinner_diff_right) { fix f g \ have \0 = cinner (f + \ *\<^sub>C g) (C *\<^sub>V (f + \ *\<^sub>C g))\ by (simp add: cinner_diff_right minus_cblinfun.rep_eq) also have \\ = \ *\<^sub>C cinner f (C g) + cnj \ *\<^sub>C cinner g (C f)\ by (smt (z3) C0 add.commute add.right_neutral cblinfun.add_right cblinfun.scaleC_right cblinfun_cinner_right.rep_eq cinner_add_left cinner_scaleC_left complex_scaleC_def) finally have \\ *\<^sub>C cinner f (C g) = - cnj \ *\<^sub>C cinner g (C f)\ by (simp add: eq_neg_iff_add_eq_0) } then have \cinner f (C g) = 0\ for f g by (metis complex_cnj_i complex_cnj_one complex_vector.scale_cancel_right complex_vector.scale_left_imp_eq equation_minus_iff i_squared mult_eq_0_iff one_neq_neg_one) then have \C g = 0\ for g using cinner_eq_zero_iff by blast then have \C = 0\ by (simp add: cblinfun_eqI) then show \A = B\ using C_def by auto qed lemma id_cblinfun_not_0[simp]: \(id_cblinfun :: 'a::{complex_normed_vector, not_singleton} \\<^sub>C\<^sub>L _) \ 0\ by (metis (full_types) Extra_General.UNIV_not_singleton cblinfun.zero_left cblinfun_id_cblinfun_apply ex_norm1 norm_zero one_neq_zero) lemma cblinfun_norm_geqI: assumes \norm (f *\<^sub>V x) / norm x \ K\ shows \norm f \ K\ using assms apply transfer by (smt (z3) bounded_clinear.bounded_linear le_onorm) (* This lemma is proven in Complex_Bounded_Linear_Function0 but we add the [simp] only here because we try to keep Complex_Bounded_Linear_Function0 as close to Bounded_Linear_Function as possible. *) declare scaleC_conv_of_complex[simp] lemma cblinfun_eq_0_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = 0" shows \F *\<^sub>V x = 0\ apply (rule complex_vector.linear_eq_0_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_on_span: fixes S::\'a::complex_normed_vector set\ assumes "x \ cspan S" and "\s. s\S \ F *\<^sub>V s = G *\<^sub>V s" shows \F *\<^sub>V x = G *\<^sub>V x\ apply (rule complex_vector.linear_eq_on_span[where f=F]) using bounded_clinear.axioms(1) cblinfun_apply assms by auto lemma cblinfun_eq_0_on_UNIV_span: fixes basis::\'a::complex_normed_vector set\ assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = 0" shows \F = 0\ by (metis cblinfun_eq_0_on_span UNIV_I assms cblinfun.zero_left cblinfun_eqI) lemma cblinfun_eq_on_UNIV_span: fixes basis::"'a::complex_normed_vector set" and \::"'a \ 'b::complex_normed_vector" assumes "cspan basis = UNIV" and "\s. s\basis \ F *\<^sub>V s = G *\<^sub>V s" shows \F = G\ proof- have "F - G = 0" apply (rule cblinfun_eq_0_on_UNIV_span[where basis=basis]) using assms by (auto simp add: cblinfun.diff_left) thus ?thesis by simp qed lemma cblinfun_eq_on_canonical_basis: fixes f g::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = g *\<^sub>V u" shows "f = g" apply (rule cblinfun_eq_on_UNIV_span[where basis=basis]) using assms is_generator_set is_cindependent_set by auto lemma cblinfun_eq_0_on_canonical_basis: fixes f ::"'a::{basis_enum,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector" defines "basis == set (canonical_basis::'a list)" assumes "\u. u \ basis \ f *\<^sub>V u = 0" shows "f = 0" by (simp add: assms cblinfun_eq_on_canonical_basis) lemma cinner_canonical_basis_eq_0: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = 0" shows "F = 0" proof- have "F *\<^sub>V u = 0" if "u\basisA" for u proof- have "\v. v\basisB \ \v, F *\<^sub>V u\ = 0" by (simp add: assms(3) that) moreover have "(\v. v\basisB \ \v, x\ = 0) \ x = 0" for x proof- assume r1: "\v. v\basisB \ \v, x\ = 0" have "\v, x\ = 0" for v proof- have "cspan basisB = UNIV" using basisB_def is_generator_set by auto hence "v \ cspan basisB" by (smt iso_tuple_UNIV_I) hence "\t s. v = (\a\t. s a *\<^sub>C a) \ finite t \ t \ basisB" using complex_vector.span_explicit by (smt mem_Collect_eq) then obtain t s where b1: "v = (\a\t. s a *\<^sub>C a)" and b2: "finite t" and b3: "t \ basisB" by blast have "\v, x\ = \(\a\t. s a *\<^sub>C a), x\" by (simp add: b1) also have "\ = (\a\t. \s a *\<^sub>C a, x\)" using cinner_sum_left by blast also have "\ = (\a\t. cnj (s a) * \a, x\)" by auto also have "\ = 0" using b3 r1 subsetD by force finally show ?thesis by simp qed thus ?thesis by (simp add: \\v. \v, x\ = 0\ cinner_extensionality) qed ultimately show ?thesis by simp qed thus ?thesis using basisA_def cblinfun_eq_0_on_canonical_basis by auto qed lemma cinner_canonical_basis_eq: defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \v, F *\<^sub>V u\ = \v, G *\<^sub>V u\" shows "F = G" proof- define H where "H = F - G" have "\u v. u\basisA \ v\basisB \ \v, H *\<^sub>V u\ = 0" unfolding H_def by (simp add: assms(3) cinner_diff_right minus_cblinfun.rep_eq) hence "H = 0" by (simp add: basisA_def basisB_def cinner_canonical_basis_eq_0) thus ?thesis unfolding H_def by simp qed lemma cinner_canonical_basis_eq': defines "basisA == set (canonical_basis::'a::onb_enum list)" and "basisB == set (canonical_basis::'b::onb_enum list)" assumes "\u v. u\basisA \ v\basisB \ \F *\<^sub>V u, v\ = \G *\<^sub>V u, v\" shows "F = G" using cinner_canonical_basis_eq assms by (metis cinner_commute') lemma cblinfun_norm_approx_witness: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ > 0\ shows \\\. norm (A *\<^sub>V \) \ norm A - \ \ norm \ = 1\ proof (transfer fixing: \) fix A :: \'a \ 'b\ assume [simp]: \bounded_clinear A\ have \\y\{norm (A x) |x. norm x = 1}. y > \ {norm (A x) |x. norm x = 1} - \\ apply (rule Sup_real_close) using assms by (auto simp: ex_norm1 bounded_clinear.bounded_linear bdd_above_norm_f) also have \\ {norm (A x) |x. norm x = 1} = onorm A\ by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear onorm_sphere) finally show \\\. onorm A - \ \ norm (A \) \ norm \ = 1\ by force qed lemma cblinfun_norm_approx_witness_mult: fixes A :: \'a::{not_singleton,complex_normed_vector} \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\ < 1\ shows \\\. norm (A *\<^sub>V \) \ norm A * \ \ norm \ = 1\ proof (cases \norm A = 0\) case True then show ?thesis apply auto by (simp add: ex_norm1) next case False then have \(1 - \) * norm A > 0\ using assms by fastforce then obtain \ where geq: \norm (A *\<^sub>V \) \ norm A - ((1 - \) * norm A)\ and \norm \ = 1\ using cblinfun_norm_approx_witness by blast have \norm A * \ = norm A - (1 - \) * norm A\ by (simp add: mult.commute right_diff_distrib') also have \\ \ norm (A *\<^sub>V \)\ by (rule geq) finally show ?thesis using \norm \ = 1\ by auto qed lemma cblinfun_to_CARD_1_0[simp]: \(A :: _ \\<^sub>C\<^sub>L _::CARD_1) = 0\ apply (rule cblinfun_eqI) by auto lemma cblinfun_from_CARD_1_0[simp]: \(A :: _::CARD_1 \\<^sub>C\<^sub>L _) = 0\ apply (rule cblinfun_eqI) apply (subst CARD_1_vec_0) by auto lemma cblinfun_cspan_UNIV: fixes basis :: \('a::{complex_normed_vector,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_normed_vector) set\ and basisA :: \'a set\ and basisB :: \'b set\ assumes \cspan basisA = UNIV\ and \cspan basisB = UNIV\ assumes basis: \\a b. a\basisA \ b\basisB \ \F\basis. \a'\basisA. F *\<^sub>V a' = (if a'=a then b else 0)\ shows \cspan basis = UNIV\ proof - obtain basisA' where \basisA' \ basisA\ and \cindependent basisA'\ and \cspan basisA' = UNIV\ by (metis assms(1) complex_vector.maximal_independent_subset complex_vector.span_eq top_greatest) then have [simp]: \finite basisA'\ by (simp add: cindependent_cfinite_dim_finite) have basis': \\a b. a\basisA' \ b\basisB \ \F\basis. \a'\basisA'. F *\<^sub>V a' = (if a'=a then b else 0)\ using basis \basisA' \ basisA\ by fastforce obtain F where F: \F a b \ basis \ F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using basis' by metis then have F_apply: \F a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ \b\basisB\ \a'\basisA'\ for a b a' using that by auto have F_basis: \F a b \ basis\ if \a\basisA'\ \b\basisB\ for a b using that F by auto have b_span: \\G\cspan {F a b|b. b\basisB}. \a'\basisA'. G *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ for a b proof - from \cspan basisB = UNIV\ obtain r t where \finite t\ and \t \ basisB\ and b_lincom: \b = (\a\t. r a *\<^sub>C a)\ unfolding complex_vector.span_alt apply atomize_elim by blast define G where \G = (\i\t. r i *\<^sub>C F a i)\ have \G \ cspan {F a b|b. b\basisB}\ using \finite t\ \t \ basisB\ unfolding G_def by (smt (verit, ccfv_threshold) complex_vector.span_base complex_vector.span_scale complex_vector.span_sum mem_Collect_eq subset_eq) moreover have \G *\<^sub>V a' = (if a'=a then b else 0)\ if \a'\basisA'\ for a' apply (cases \a'=a\) using \t \ basisB\ \a\basisA'\ \a'\basisA'\ by (auto simp: b_lincom G_def cblinfun.sum_left F_apply intro!: sum.neutral sum.cong) ultimately show ?thesis by blast qed have a_span: \cspan (\a\basisA'. cspan {F a b|b. b\basisB}) = UNIV\ proof (intro equalityI subset_UNIV subsetI, rename_tac H) fix H obtain G where G: \G a b \ cspan {F a b|b. b\basisB} \ G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' apply atomize_elim apply (intro choice allI) using b_span by blast then have G_cspan: \G a b \ cspan {F a b|b. b\basisB}\ if \a\basisA'\ for a b using that by auto from G have G: \G a b *\<^sub>V a' = (if a'=a then b else 0)\ if \a\basisA'\ and \a'\basisA'\ for a b a' using that by auto define H' where \H' = (\a\basisA'. G a (H *\<^sub>V a))\ have \H' \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ unfolding H'_def using G_cspan by (smt (verit, del_insts) UN_iff complex_vector.span_clauses(1) complex_vector.span_sum) moreover have \H' = H\ using \cspan basisA' = UNIV\ apply (rule cblinfun_eq_on_UNIV_span) apply (auto simp: H'_def cblinfun.sum_left) apply (subst sum_single) by (auto simp: G) ultimately show \H \ cspan (\a\basisA'. cspan {F a b |b. b \ basisB})\ by simp qed moreover have \cspan basis \ cspan (\a\basisA'. cspan {F a b|b. b\basisB})\ using F_basis by (smt (z3) UN_subset_iff complex_vector.span_alt complex_vector.span_minimal complex_vector.subspace_span mem_Collect_eq subset_iff) ultimately show \cspan basis = UNIV\ by auto qed instance cblinfun :: (\{cfinite_dim,complex_normed_vector}\, \{cfinite_dim,complex_normed_vector}\) cfinite_dim proof intro_classes obtain basisA :: \'a set\ where [simp]: \cspan basisA = UNIV\ \cindependent basisA\ \finite basisA\ using finite_basis by blast obtain basisB :: \'b set\ where [simp]: \cspan basisB = UNIV\ \cindependent basisB\ \finite basisB\ using finite_basis by blast define f where \f a b = cconstruct basisA (\x. if x=a then b else 0)\ for a :: 'a and b :: 'b have f_a: \f a b a = b\ if \a : basisA\ for a b by (simp add: complex_vector.construct_basis f_def that) have f_not_a: \f a b c = 0\ if \a : basisA\ and \c : basisA\ and \a \ c\for a b c using that by (simp add: complex_vector.construct_basis f_def) define F where \F a b = CBlinfun (f a b)\ for a b have \clinear (f a b)\ for a b by (auto intro: complex_vector.linear_construct simp: f_def) then have \bounded_clinear (f a b)\ for a b by auto then have F_apply: \cblinfun_apply (F a b) = f a b\ for a b by (simp add: F_def bounded_clinear_CBlinfun_apply) define basis where \basis = {F a b| a b. a\basisA \ b\basisB}\ have \cspan basis = UNIV\ apply (rule cblinfun_cspan_UNIV[where basisA=basisA and basisB=basisB]) apply (auto simp: basis_def) by (metis F_apply f_a f_not_a) moreover have \finite basis\ unfolding basis_def apply (rule finite_image_set2) by auto ultimately show \\S :: ('a \\<^sub>C\<^sub>L 'b) set. finite S \ cspan S = UNIV\ by auto qed subsection \Relationship to real bounded operators (\<^typ>\_ \\<^sub>L _\)\ instantiation blinfun :: (real_normed_vector, complex_normed_vector) "complex_normed_vector" begin lift_definition scaleC_blinfun :: \complex \ ('a::real_normed_vector, 'b::complex_normed_vector) blinfun \ ('a, 'b) blinfun\ is \\ c::complex. \ f::'a\'b. (\ x. c *\<^sub>C (f x) )\ proof fix c::complex and f :: \'a\'b\ and b1::'a and b2::'a assume \bounded_linear f\ show \c *\<^sub>C f (b1 + b2) = c *\<^sub>C f b1 + c *\<^sub>C f b2\ by (simp add: \bounded_linear f\ linear_simps scaleC_add_right) fix c::complex and f :: \'a\'b\ and b::'a and r::real assume \bounded_linear f\ show \c *\<^sub>C f (r *\<^sub>R b) = r *\<^sub>R (c *\<^sub>C f b)\ by (simp add: \bounded_linear f\ linear_simps(5) scaleR_scaleC) fix c::complex and f :: \'a\'b\ assume \bounded_linear f\ have \\ K. \ x. norm (f x) \ norm x * K\ using \bounded_linear f\ by (simp add: bounded_linear.bounded) then obtain K where \\ x. norm (f x) \ norm x * K\ by blast have \cmod c \ 0\ by simp hence \\ x. (cmod c) * norm (f x) \ (cmod c) * norm x * K\ using \\ x. norm (f x) \ norm x * K\ by (metis ordered_comm_semiring_class.comm_mult_left_mono vector_space_over_itself.scale_scale) moreover have \norm (c *\<^sub>C f x) = (cmod c) * norm (f x)\ for x by simp ultimately show \\K. \x. norm (c *\<^sub>C f x) \ norm x * K\ by (metis ab_semigroup_mult_class.mult_ac(1) mult.commute) qed instance proof have "r *\<^sub>R x = complex_of_real r *\<^sub>C x" for x :: "('a, 'b) blinfun" and r apply transfer by (simp add: scaleR_scaleC) thus "((*\<^sub>R) r::'a \\<^sub>L 'b \ _) = (*\<^sub>C) (complex_of_real r)" for r by auto show "a *\<^sub>C (x + y) = a *\<^sub>C x + a *\<^sub>C y" for a :: complex and x y :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_right) show "(a + b) *\<^sub>C x = a *\<^sub>C x + b *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by (simp add: scaleC_add_left) show "a *\<^sub>C b *\<^sub>C x = (a * b) *\<^sub>C x" for a b :: complex and x :: "'a \\<^sub>L 'b" apply transfer by simp have \1 *\<^sub>C f x = f x\ for f :: \'a\'b\ and x by auto thus "1 *\<^sub>C x = x" for x :: "'a \\<^sub>L 'b" by (simp add: scaleC_blinfun.rep_eq blinfun_eqI) have \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ if \bounded_linear f\ for f :: \'a \ 'b\ and a :: complex proof- have \cmod a \ 0\ by simp have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \\ x. (cmod a) *(\ ereal ((norm (f x)) / (norm x)) \) \ (cmod a) * K\ using \cmod a \ 0\ by (metis abs_ereal.simps(1) abs_ereal_pos abs_pos ereal_mult_left_mono times_ereal.simps(1)) hence \\ x. (\ ereal ((cmod a) * (norm (f x)) / (norm x)) \) \ (cmod a) * K\ by simp hence \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have p1: \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ using \\ x. \ ereal (cmod a * (norm (f x)) / (norm x)) \ \ cmod a * K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) have p2: \\i. i \ UNIV \ 0 \ ereal (cmod a * norm (f i) / norm i)\ by simp hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ (SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\)\ using \bdd_above {ereal (cmod a * (norm (f x)) / (norm x)) | x. True}\ \{ereal (cmod a * (norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I p2 abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal (cmod a * (norm (f x)) / (norm x))\ \ cmod a * K\ using \(SUP x. \ereal (cmod a * (norm (f x)) / (norm x))\) \ cmod a * K\ by simp hence \\ ( SUP i\UNIV::'a set. ereal ((\ x. (cmod a) * (norm (f x)) / norm x) i)) \ \ \\ by auto hence w2: \( SUP i\UNIV::'a set. ereal ((\ x. cmod a * (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. cmod a * (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) have \(UNIV::('a set)) \ {}\ by simp moreover have \\ i. i \ (UNIV::('a set)) \ (\ x. (norm (f x)) / norm x :: ereal) i \ 0\ by simp moreover have \cmod a \ 0\ by simp ultimately have \(SUP i\(UNIV::('a set)). ((cmod a)::ereal) * (\ x. (norm (f x)) / norm x :: ereal) i ) = ((cmod a)::ereal) * ( SUP i\(UNIV::('a set)). (\ x. (norm (f x)) / norm x :: ereal) i )\ by (simp add: Sup_ereal_mult_left') hence \(SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp hence z1: \real_of_ereal ( (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) ) = real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) )\ by simp have z2: \real_of_ereal (SUP x. ((cmod a)::ereal) * ( (norm (f x)) / norm x :: ereal) ) = (SUP x. cmod a * (norm (f x) / norm x))\ using w2 by auto have \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = (cmod a) * real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) )\ by simp moreover have \real_of_ereal ( SUP x. ( (norm (f x)) / norm x :: ereal) ) = ( SUP x. ((norm (f x)) / norm x) )\ proof- have \\ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) \ \ \\ proof- have \\ K::real. \ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ using \bounded_linear f\ le_onorm by fastforce then obtain K::real where \\ x. (\ ereal ((norm (f x)) / (norm x)) \) \ K\ by blast hence \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ by simp moreover have \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by auto ultimately have \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ using \\ x. \ ereal ((norm (f x)) / (norm x)) \ \ K\ Sup_least mem_Collect_eq by (simp add: SUP_le_iff) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ (SUP x. \ereal ((norm (f x)) / (norm x))\)\ using \bdd_above {ereal ((norm (f x)) / (norm x)) | x. True}\ \{ereal ((norm (f x)) / (norm x)) | x. True} \ {}\ by (metis (mono_tags, lifting) SUP_upper2 Sup.SUP_cong UNIV_I \\i. i \ UNIV \ 0 \ ereal (norm (f i) / norm i)\ abs_ereal_ge0 ereal_le_real) hence \\SUP x. ereal ((norm (f x)) / (norm x))\ \ K\ using \(SUP x. \ereal ((norm (f x)) / (norm x))\) \ K\ by simp thus ?thesis by auto qed hence \ ( SUP i\UNIV::'a set. ereal ((\ x. (norm (f x)) / norm x) i)) = ereal ( Sup ((\ x. (norm (f x)) / norm x) ` (UNIV::'a set) ))\ by (simp add: ereal_SUP) thus ?thesis by simp qed have z3: \real_of_ereal ( ((cmod a)::ereal) * ( SUP x. ( (norm (f x)) / norm x :: ereal) ) ) = cmod a * (SUP x. norm (f x) / norm x)\ by (simp add: \real_of_ereal (SUP x. ereal (norm (f x) / norm x)) = (SUP x. norm (f x) / norm x)\) hence w1: \(SUP x. cmod a * (norm (f x) / norm x)) = cmod a * (SUP x. norm (f x) / norm x)\ using z1 z2 by linarith have v1: \onorm (\x. a *\<^sub>C f x) = (SUP x. norm (a *\<^sub>C f x) / norm x)\ by (simp add: onorm_def) have v2: \(SUP x. norm (a *\<^sub>C f x) / norm x) = (SUP x. ((cmod a) * norm (f x)) / norm x)\ by simp have v3: \(SUP x. ((cmod a) * norm (f x)) / norm x) = (SUP x. (cmod a) * ((norm (f x)) / norm x))\ by simp have v4: \(SUP x. (cmod a) * ((norm (f x)) / norm x)) = (cmod a) * (SUP x. ((norm (f x)) / norm x))\ using w1 by blast show \onorm (\x. a *\<^sub>C f x) = cmod a * onorm f\ using v1 v2 v3 v4 by (metis (mono_tags, lifting) onorm_def) qed thus \norm (a *\<^sub>C x) = cmod a * norm x\ for a::complex and x::\('a, 'b) blinfun\ apply transfer by blast qed end (* We do not have clinear_blinfun_compose_right *) lemma clinear_blinfun_compose_left: \clinear (\x. blinfun_compose x y)\ by (auto intro!: clinearI simp: blinfun_eqI scaleC_blinfun.rep_eq bounded_bilinear.add_left bounded_bilinear_blinfun_compose) instantiation blinfun :: (real_normed_vector, cbanach) "cbanach" begin instance.. end lemma blinfun_compose_assoc: "(A o\<^sub>L B) o\<^sub>L C = A o\<^sub>L (B o\<^sub>L C)" by (simp add: blinfun_eqI) lift_definition blinfun_of_cblinfun::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a \\<^sub>L 'b\ is "id" apply transfer by (simp add: bounded_clinear.bounded_linear) lift_definition blinfun_cblinfun_eq :: \'a \\<^sub>L 'b \ 'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ bool\ is "(=)" . lemma blinfun_cblinfun_eq_bi_unique[transfer_rule]: \bi_unique blinfun_cblinfun_eq\ unfolding bi_unique_def apply transfer by auto lemma blinfun_cblinfun_eq_right_total[transfer_rule]: \right_total blinfun_cblinfun_eq\ unfolding right_total_def apply transfer by (simp add: bounded_clinear.bounded_linear) named_theorems cblinfun_blinfun_transfer lemma cblinfun_blinfun_transfer_0[cblinfun_blinfun_transfer]: "blinfun_cblinfun_eq (0::(_,_) blinfun) (0::(_,_) cblinfun)" apply transfer by simp lemma cblinfun_blinfun_transfer_plus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (+) (+)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_minus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (-) (-)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_uminus[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (uminus) (uminus)" unfolding rel_fun_def apply transfer by auto definition "real_complex_eq r c \ complex_of_real r = c" lemma bi_unique_real_complex_eq[transfer_rule]: \bi_unique real_complex_eq\ unfolding real_complex_eq_def bi_unique_def by auto lemma left_total_real_complex_eq[transfer_rule]: \left_total real_complex_eq\ unfolding real_complex_eq_def left_total_def by auto lemma cblinfun_blinfun_transfer_scaleC[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(real_complex_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (scaleR) (scaleC)" unfolding rel_fun_def apply transfer by (simp add: real_complex_eq_def scaleR_scaleC) lemma cblinfun_blinfun_transfer_CBlinfun[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(eq_onp bounded_clinear ===> blinfun_cblinfun_eq) Blinfun CBlinfun" unfolding rel_fun_def blinfun_cblinfun_eq.rep_eq eq_onp_def by (auto simp: CBlinfun_inverse Blinfun_inverse bounded_clinear.bounded_linear) lemma cblinfun_blinfun_transfer_norm[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=)) norm norm" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_dist[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> (=)) dist dist" unfolding rel_fun_def dist_norm apply transfer by auto lemma cblinfun_blinfun_transfer_sgn[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) sgn sgn" unfolding rel_fun_def sgn_blinfun_def sgn_cblinfun_def apply transfer by (auto simp: scaleR_scaleC) lemma cblinfun_blinfun_transfer_Cauchy[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> (=)) Cauchy Cauchy" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding Cauchy_def by transfer_prover qed lemma cblinfun_blinfun_transfer_tendsto[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(((=) ===> blinfun_cblinfun_eq) ===> blinfun_cblinfun_eq ===> (=) ===> (=)) tendsto tendsto" proof - note cblinfun_blinfun_transfer[transfer_rule] show ?thesis unfolding tendsto_iff by transfer_prover qed lemma cblinfun_blinfun_transfer_compose[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> blinfun_cblinfun_eq ===> blinfun_cblinfun_eq) (o\<^sub>L) (o\<^sub>C\<^sub>L)" unfolding rel_fun_def apply transfer by auto lemma cblinfun_blinfun_transfer_apply[cblinfun_blinfun_transfer]: includes lifting_syntax shows "(blinfun_cblinfun_eq ===> (=) ===> (=)) blinfun_apply cblinfun_apply" unfolding rel_fun_def apply transfer by auto lemma blinfun_of_cblinfun_inj: \blinfun_of_cblinfun f = blinfun_of_cblinfun g \ f = g\ by (metis cblinfun_apply_inject blinfun_of_cblinfun.rep_eq) lemma blinfun_of_cblinfun_inv: assumes "\c. \x. f *\<^sub>v (c *\<^sub>C x) = c *\<^sub>C (f *\<^sub>v x)" shows "\g. blinfun_of_cblinfun g = f" using assms proof transfer show "\g\Collect bounded_clinear. id g = f" if "bounded_linear f" and "\c x. f (c *\<^sub>C x) = c *\<^sub>C f x" for f :: "'a \ 'b" using that bounded_linear_bounded_clinear by auto qed lemma blinfun_of_cblinfun_zero: \blinfun_of_cblinfun 0 = 0\ apply transfer by simp lemma blinfun_of_cblinfun_uminus: \blinfun_of_cblinfun (- f) = - (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_minus: \blinfun_of_cblinfun (f - g) = blinfun_of_cblinfun f - blinfun_of_cblinfun g\ apply transfer by auto lemma blinfun_of_cblinfun_scaleC: \blinfun_of_cblinfun (c *\<^sub>C f) = c *\<^sub>C (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_scaleR: \blinfun_of_cblinfun (c *\<^sub>R f) = c *\<^sub>R (blinfun_of_cblinfun f)\ apply transfer by auto lemma blinfun_of_cblinfun_norm: fixes f::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ shows \norm f = norm (blinfun_of_cblinfun f)\ apply transfer by auto subsection \Composition\ lemma blinfun_of_cblinfun_cblinfun_compose: fixes f::\'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector\ and g::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b\ shows \blinfun_of_cblinfun (f o\<^sub>C\<^sub>L g) = (blinfun_of_cblinfun f) o\<^sub>L (blinfun_of_cblinfun g)\ apply transfer by auto lemma cblinfun_compose_assoc: shows "(A o\<^sub>C\<^sub>L B) o\<^sub>C\<^sub>L C = A o\<^sub>C\<^sub>L (B o\<^sub>C\<^sub>L C)" by (metis (no_types, lifting) cblinfun_apply_inject fun.map_comp cblinfun_compose.rep_eq) lemma cblinfun_compose_zero_right[simp]: "U o\<^sub>C\<^sub>L 0 = 0" using bounded_cbilinear.zero_right bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_zero_left[simp]: "0 o\<^sub>C\<^sub>L U = 0" using bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose by blast lemma cblinfun_compose_scaleC_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>C A) o\<^sub>C\<^sub>L B = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ by (simp add: bounded_cbilinear.scaleC_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_scaleR_left[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \(a *\<^sub>R A) o\<^sub>C\<^sub>L B = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_scaleC_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>C B) = a *\<^sub>C (A o\<^sub>C\<^sub>L B)\ apply transfer by (auto intro!: ext bounded_clinear.clinear complex_vector.linear_scale) lemma cblinfun_compose_scaleR_right[simp]: fixes A::"'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::complex_normed_vector" and B::"'a::complex_normed_vector \\<^sub>C\<^sub>L 'b" shows \A o\<^sub>C\<^sub>L (a *\<^sub>R B) = a *\<^sub>R (A o\<^sub>C\<^sub>L B)\ by (simp add: scaleR_scaleC) lemma cblinfun_compose_id_right[simp]: shows "U o\<^sub>C\<^sub>L id_cblinfun = U" apply transfer by auto lemma cblinfun_compose_id_left[simp]: shows "id_cblinfun o\<^sub>C\<^sub>L U = U" apply transfer by auto lemma cblinfun_eq_on: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \t \ closure (cspan G)\ shows "A *\<^sub>V t = B *\<^sub>V t" using assms apply transfer using bounded_clinear_eq_on by blast lemma cblinfun_eq_gen_eqI: fixes A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::complex_normed_vector" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and \ccspan G = \\ shows "A = B" apply (rule cblinfun_eqI) apply (rule cblinfun_eq_on[where G=G]) using assms apply auto by (metis ccspan.rep_eq iso_tuple_UNIV_I top_ccsubspace.rep_eq) lemma cblinfun_compose_add_left: \(a + b) o\<^sub>C\<^sub>L c = (a o\<^sub>C\<^sub>L c) + (b o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma cblinfun_compose_add_right: \a o\<^sub>C\<^sub>L (b + c) = (a o\<^sub>C\<^sub>L b) + (a o\<^sub>C\<^sub>L c)\ by (simp add: bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) lemma cbilinear_cblinfun_compose[simp]: "cbilinear cblinfun_compose" by (auto intro!: clinearI simp add: cbilinear_def bounded_cbilinear.add_left bounded_cbilinear.add_right bounded_cbilinear_cblinfun_compose) subsection \Adjoint\ lift_definition adj :: "'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ 'b \\<^sub>C\<^sub>L 'a" ("_*" [99] 100) is cadjoint by (fact cadjoint_bounded_clinear) lemma id_cblinfun_adjoint[simp]: "id_cblinfun* = id_cblinfun" apply transfer using cadjoint_id by (metis eq_id_iff) lemma double_adj[simp]: "(A*)* = A" apply transfer using double_cadjoint by blast lemma adj_cblinfun_compose[simp]: fixes B::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and A::\'b \\<^sub>C\<^sub>L 'c::complex_inner\ shows "(A o\<^sub>C\<^sub>L B)* = (B*) o\<^sub>C\<^sub>L (A*)" proof transfer fix A :: \'b \ 'c\ and B :: \'a \ 'b\ assume \bounded_clinear A\ and \bounded_clinear B\ hence \bounded_clinear (A \ B)\ by (simp add: comp_bounded_clinear) have \\ (A \ B) u, v \ = \ u, (B\<^sup>\ \ A\<^sup>\) v \\ for u v by (metis (no_types, lifting) cadjoint_univ_prop \bounded_clinear A\ \bounded_clinear B\ cinner_commute' comp_def) thus \(A \ B)\<^sup>\ = B\<^sup>\ \ A\<^sup>\\ using \bounded_clinear (A \ B)\ by (metis cadjoint_eqI cinner_commute') qed lemma scaleC_adj[simp]: "(a *\<^sub>C A)* = (cnj a) *\<^sub>C (A*)" apply transfer by (simp add: Complex_Vector_Spaces0.bounded_clinear.bounded_linear bounded_clinear_def complex_vector.linear_scale scaleC_cadjoint) lemma scaleR_adj[simp]: "(a *\<^sub>R A)* = a *\<^sub>R (A*)" by (simp add: scaleR_scaleC) lemma adj_plus: \(A + B)* = (A*) + (B*)\ proof transfer fix A B::\'b \ 'a\ assume a1: \bounded_clinear A\ and a2: \bounded_clinear B\ define F where \F = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ define G where \G = (\x. A x + B x)\ have \bounded_clinear G\ unfolding G_def by (simp add: a1 a2 bounded_clinear_add) moreover have \\F u, v\ = \u, G v\\ for u v unfolding F_def G_def using cadjoint_univ_prop a1 a2 cinner_add_left by (simp add: cadjoint_univ_prop cinner_add_left cinner_add_right) ultimately have \F = G\<^sup>\ \ using cadjoint_eqI by blast thus \(\x. A x + B x)\<^sup>\ = (\x. (A\<^sup>\) x + (B\<^sup>\) x)\ unfolding F_def G_def by auto qed lemma cinner_sup_norm_cblinfun: fixes A :: \'a::{complex_normed_vector,not_singleton} \\<^sub>C\<^sub>L 'b::complex_inner\ shows \norm A = (SUP (\,\). cmod (cinner \ (A *\<^sub>V \)) / (norm \ * norm \))\ apply transfer apply (rule cinner_sup_onorm) by (simp add: bounded_clinear.bounded_linear) lemma cinner_adj_left: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\G* *\<^sub>V x, y\ = \x, G *\<^sub>V y\\ apply transfer using cadjoint_univ_prop by blast lemma cinner_adj_right: fixes G :: "'b::chilbert_space \\<^sub>C\<^sub>L 'a::complex_inner" shows \\x, G* *\<^sub>V y\ = \G *\<^sub>V x, y\\ apply transfer using cadjoint_univ_prop' by blast lemma adj_0[simp]: \0* = 0\ by (metis add_cancel_right_left adj_plus) lemma norm_adj[simp]: \norm (A*) = norm A\ for A :: \'b::chilbert_space \\<^sub>C\<^sub>L 'c::complex_inner\ proof (cases \(\x y :: 'b. x \ y) \ (\x y :: 'c. x \ y)\) case True then have c1: \class.not_singleton TYPE('b)\ apply intro_classes by simp from True have c2: \class.not_singleton TYPE('c)\ apply intro_classes by simp have normA: \norm A = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c1) have normAadj: \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A* *\<^sub>V \)) / (norm \ * norm \))\ apply (rule cinner_sup_norm_cblinfun[internalize_sort \'a::{complex_normed_vector,not_singleton}\]) apply (rule complex_normed_vector_axioms) by (rule c2) have \norm (A*) = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ unfolding normAadj apply (subst cinner_adj_right) apply (subst cinner_commute) apply (subst complex_mod_cnj) by rule also have \\ = Sup ((\(\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \)) ` prod.swap ` UNIV)\ by auto also have \\ = (SUP (\, \). cmod (\ \\<^sub>C (A *\<^sub>V \)) / (norm \ * norm \))\ apply (subst image_image) by auto also have \\ = norm A\ unfolding normA by (simp add: mult.commute) finally show ?thesis by - next case False then consider (b) \\x::'b. x = 0\ | (c) \\x::'c. x = 0\ by auto then have \A = 0\ apply (cases; transfer) apply (metis (full_types) bounded_clinear_def complex_vector.linear_0) by auto then show \norm (A*) = norm A\ by simp qed lemma antilinear_adj[simp]: \antilinear adj\ apply (rule antilinearI) by (auto simp add: adj_plus) lemma bounded_antilinear_adj[bounded_antilinear, simp]: \bounded_antilinear adj\ by (auto intro!: antilinearI exI[of _ 1] simp: bounded_antilinear_def bounded_antilinear_axioms_def adj_plus) lemma adjoint_eqI: fixes G:: \'b::chilbert_space \\<^sub>C\<^sub>L 'a::chilbert_space\ and F:: \'a \\<^sub>C\<^sub>L 'b\ assumes \\x y. \(cblinfun_apply F) x, y\ = \x, (cblinfun_apply G) y\\ shows \F = G*\ using assms apply transfer using cadjoint_eqI by auto lemma cinner_real_hermiteanI: \ \Prop. II.2.12 in @{cite conway2013course}\ assumes \\\. cinner \ (A *\<^sub>V \) \ \\ shows \A = A*\ proof - { fix g h :: 'a { fix \ :: complex have \cinner h (A h) + cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) + (abs \)\<^sup>2 * cinner g (A g) = cinner (h + \ *\<^sub>C g) (A *\<^sub>V (h + \ *\<^sub>C g))\ (is \?sum4 = _\) apply (auto simp: cinner_add_right cinner_add_left cblinfun.add_right cblinfun.scaleC_right ring_class.ring_distribs) by (metis cnj_x_x mult.commute) also have \\ \ \\ using assms by auto finally have \?sum4 = cnj ?sum4\ using Reals_cnj_iff by fastforce then have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner (A h) g + cnj \ *\<^sub>C cinner (A g) h\ using Reals_cnj_iff abs_complex_real assms by force also have \\ = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by (simp add: cinner_adj_right) finally have \cnj \ *\<^sub>C cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) + cnj \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by - } from this[where \2=1] this[where \2=\] have 1: \cinner g (A h) + cinner h (A g) = cinner h (A* *\<^sub>V g) + cinner g (A* *\<^sub>V h)\ and i: \- \ * cinner g (A h) + \ *\<^sub>C cinner h (A g) = \ *\<^sub>C cinner h (A* *\<^sub>V g) - \ *\<^sub>C cinner g (A* *\<^sub>V h)\ by auto from arg_cong2[OF 1 arg_cong[OF i, where f=\(*) (-\)\], where f=plus] have \cinner h (A g) = cinner h (A* *\<^sub>V g)\ by (auto simp: ring_class.ring_distribs) } then show "A = A*" by (simp add: adjoint_eqI cinner_adj_right) qed lemma norm_AAadj[simp]: \norm (A o\<^sub>C\<^sub>L A*) = (norm A)\<^sup>2\ for A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof (cases \class.not_singleton TYPE('b)\) case True then have [simp]: \class.not_singleton TYPE('b)\ by - have 1: \(norm A)\<^sup>2 * \ \ norm (A o\<^sub>C\<^sub>L A*)\ if \\ < 1\ and \\ \ 0\ for \ proof - obtain \ where \: \norm ((A*) *\<^sub>V \) \ norm (A*) * sqrt \\ and [simp]: \norm \ = 1\ apply atomize_elim apply (rule cblinfun_norm_approx_witness_mult[internalize_sort' 'a]) using \\ < 1\ by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) have \complex_of_real ((norm A)\<^sup>2 * \) = (norm (A*) * sqrt \)\<^sup>2\ by (simp add: ordered_field_class.sign_simps(23) that(2)) also have \\ \ (norm ((A* *\<^sub>V \)))\<^sup>2\ apply (rule complex_of_real_mono) using \ apply (rule power_mono) using \\ \ 0\ by auto also have \\ \ cinner (A* *\<^sub>V \) (A* *\<^sub>V \)\ by (auto simp flip: power2_norm_eq_cinner) also have \\ = cinner \ (A *\<^sub>V A* *\<^sub>V \)\ by (simp add: cinner_adj_left) also have \\ = cinner \ ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ by auto also have \\ \ norm (A o\<^sub>C\<^sub>L A*)\ using \norm \ = 1\ by (smt (verit, best) Im_complex_of_real Re_complex_of_real \(A* *\<^sub>V \) \\<^sub>C (A* *\<^sub>V \) = \ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \)\ \\ \\<^sub>C (A *\<^sub>V A* *\<^sub>V \) = \ \\<^sub>C ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \)\ cdot_square_norm cinner_ge_zero cmod_Re complex_inner_class.Cauchy_Schwarz_ineq2 less_eq_complex_def mult_cancel_left1 mult_cancel_right1 norm_cblinfun) finally show ?thesis - by auto + by (auto simp: less_eq_complex_def) qed then have 1: \(norm A)\<^sup>2 \ norm (A o\<^sub>C\<^sub>L A*)\ by (metis field_le_mult_one_interval less_eq_real_def ordered_field_class.sign_simps(5)) have 2: \norm (A o\<^sub>C\<^sub>L A*) \ (norm A)\<^sup>2\ proof (rule norm_cblinfun_bound) show \0 \ (norm A)\<^sup>2\ by simp fix \ have \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) = norm (A *\<^sub>V A* *\<^sub>V \)\ by auto also have \\ \ norm A * norm (A* *\<^sub>V \)\ by (simp add: norm_cblinfun) also have \\ \ norm A * norm (A*) * norm \\ by (metis mult.assoc norm_cblinfun norm_imp_pos_and_ge ordered_comm_semiring_class.comm_mult_left_mono) also have \\ = (norm A)\<^sup>2 * norm \\ by (simp add: power2_eq_square) finally show \norm ((A o\<^sub>C\<^sub>L A*) *\<^sub>V \) \ (norm A)\<^sup>2 * norm \\ by - qed from 1 2 show ?thesis by simp next case False then have [simp]: \class.CARD_1 TYPE('b)\ by (rule not_singleton_vs_CARD_1) have \A = 0\ apply (rule cblinfun_to_CARD_1_0[internalize_sort' 'b]) by (auto intro: complex_normed_vector_class.complex_normed_vector_axioms) then show ?thesis by auto qed subsection \Unitaries / isometries\ definition isometry::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun\ definition unitary::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner \ bool\ where \unitary U \ (U* o\<^sub>C\<^sub>L U = id_cblinfun) \ (U o\<^sub>C\<^sub>L U* = id_cblinfun)\ lemma unitary_twosided_isometry: "unitary U \ isometry U \ isometry (U*)" unfolding unitary_def isometry_def by simp lemma isometryD[simp]: "isometry U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding isometry_def by simp (* Not [simp] because isometryD[simp] + unitary_isometry[simp] already have the same effect *) lemma unitaryD1: "unitary U \ U* o\<^sub>C\<^sub>L U = id_cblinfun" unfolding unitary_def by simp lemma unitaryD2[simp]: "unitary U \ U o\<^sub>C\<^sub>L U* = id_cblinfun" unfolding unitary_def by simp lemma unitary_isometry[simp]: "unitary U \ isometry U" unfolding unitary_def isometry_def by simp lemma unitary_adj[simp]: "unitary (U*) = unitary U" unfolding unitary_def by auto lemma isometry_cblinfun_compose[simp]: assumes "isometry A" and "isometry B" shows "isometry (A o\<^sub>C\<^sub>L B)" proof- have "B* o\<^sub>C\<^sub>L A* o\<^sub>C\<^sub>L (A o\<^sub>C\<^sub>L B) = id_cblinfun" if "A* o\<^sub>C\<^sub>L A = id_cblinfun" and "B* o\<^sub>C\<^sub>L B = id_cblinfun" using that by (smt (verit, del_insts) adjoint_eqI cblinfun_apply_cblinfun_compose cblinfun_id_cblinfun_apply) thus ?thesis using assms unfolding isometry_def by simp qed lemma unitary_cblinfun_compose[simp]: "unitary (A o\<^sub>C\<^sub>L B)" if "unitary A" and "unitary B" using that by (smt (z3) adj_cblinfun_compose cblinfun_compose_assoc cblinfun_compose_id_right double_adj isometryD isometry_cblinfun_compose unitary_def unitary_isometry) lemma unitary_surj: assumes "unitary U" shows "surj (cblinfun_apply U)" apply (rule surjI[where f=\cblinfun_apply (U*)\]) using assms unfolding unitary_def apply transfer using comp_eq_dest_lhs by force lemma unitary_id[simp]: "unitary id_cblinfun" by (simp add: unitary_def) lemma orthogonal_on_basis_is_isometry: assumes spanB: \ccspan B = \\ assumes orthoU: \\b c. b\B \ c\B \ cinner (U *\<^sub>V b) (U *\<^sub>V c) = cinner b c\ shows \isometry U\ proof - have [simp]: \b \ closure (cspan B)\ for b using spanB apply transfer by simp have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ and \\\B\ for \ \ by (simp add: cinner_adj_left orthoU that(1) that(2)) have *: \cinner (U* *\<^sub>V U *\<^sub>V \) \ = cinner \ \\ if \\\B\ for \ \ apply (rule bounded_clinear_eq_on[where t=\ and G=B]) using bounded_clinear_cinner_right *[OF that] by auto have \U* *\<^sub>V U *\<^sub>V \ = \\ if \\\B\ for \ apply (rule cinner_extensionality) apply (subst cinner_eq_flip) by (simp add: * that) then have \U* o\<^sub>C\<^sub>L U = id_cblinfun\ by (metis cblinfun_apply_cblinfun_compose cblinfun_eq_gen_eqI cblinfun_id_cblinfun_apply spanB) then show \isometry U\ using isometry_def by blast qed subsection \Images\ (* Closure is necessary. See email 47a3bb3d-3cc3-0934-36eb-3ef0f7b70a85@ut.ee *) lift_definition cblinfun_image :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector \ 'a ccsubspace \ 'b ccsubspace\ (infixr "*\<^sub>S" 70) is "\A S. closure (A ` S)" using bounded_clinear_def closed_closure closed_csubspace.intro by (simp add: bounded_clinear_def complex_vector.linear_subspace_image closure_is_closed_csubspace) lemma cblinfun_image_mono: assumes a1: "S \ T" shows "A *\<^sub>S S \ A *\<^sub>S T" using a1 by (simp add: cblinfun_image.rep_eq closure_mono image_mono less_eq_ccsubspace.rep_eq) lemma cblinfun_image_0[simp]: shows "U *\<^sub>S 0 = 0" thm zero_ccsubspace_def apply transfer by (simp add: bounded_clinear_def complex_vector.linear_0) lemma cblinfun_image_bot[simp]: "U *\<^sub>S bot = bot" using cblinfun_image_0 by auto lemma cblinfun_image_sup[simp]: fixes A B :: \'a::chilbert_space ccsubspace\ and U :: "'a \\<^sub>C\<^sub>L'b::chilbert_space" shows \U *\<^sub>S (sup A B) = sup (U *\<^sub>S A) (U *\<^sub>S B)\ apply transfer using bounded_clinear.bounded_linear closure_image_closed_sum by blast lemma scaleC_cblinfun_image[simp]: fixes A :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b :: chilbert_space\ and S :: \'a ccsubspace\ and \ :: complex shows \(\ *\<^sub>C A) *\<^sub>S S = \ *\<^sub>C (A *\<^sub>S S)\ proof- have \closure ( ( ((*\<^sub>C) \) \ (cblinfun_apply A) ) ` space_as_set S) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis closure_scaleC image_comp) hence \(closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = ((*\<^sub>C) \) ` (closure (cblinfun_apply A ` space_as_set S))\ by (metis (mono_tags, lifting) comp_apply image_cong scaleC_cblinfun.rep_eq) hence \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by (metis space_as_set_inverse cblinfun_image.rep_eq scaleC_ccsubspace.rep_eq) have x1: "Abs_clinear_space (closure ((*\<^sub>V) (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure ((*\<^sub>V) A ` space_as_set S))" using \Abs_clinear_space (closure (cblinfun_apply (\ *\<^sub>C A) ` space_as_set S)) = \ *\<^sub>C Abs_clinear_space (closure (cblinfun_apply A ` space_as_set S))\ by blast show ?thesis unfolding cblinfun_image_def using x1 by force qed lemma cblinfun_image_id[simp]: "id_cblinfun *\<^sub>S \ = \" apply transfer by (simp add: closed_csubspace.closed) lemma cblinfun_compose_image: \(A o\<^sub>C\<^sub>L B) *\<^sub>S S = A *\<^sub>S (B *\<^sub>S S)\ apply transfer unfolding image_comp[symmetric] apply (rule closure_bounded_linear_image_subset_eq[symmetric]) by (simp add: bounded_clinear.bounded_linear) lemmas cblinfun_assoc_left = cblinfun_compose_assoc[symmetric] cblinfun_compose_image[symmetric] add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space", symmetric] lemmas cblinfun_assoc_right = cblinfun_compose_assoc cblinfun_compose_image add.assoc[where ?'a="'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space"] lemma cblinfun_image_INF_leq[simp]: fixes U :: "'b::complex_normed_vector \\<^sub>C\<^sub>L 'c::cbanach" and V :: "'a \ 'b ccsubspace" shows \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ apply transfer by (simp add: INT_greatest Inter_lower closure_mono image_mono) lemma isometry_cblinfun_image_inf_distrib': fixes U::\'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::cbanach\ and B C::"'a ccsubspace" shows "U *\<^sub>S (inf B C) \ inf (U *\<^sub>S B) (U *\<^sub>S C)" proof - define V where \V b = (if b then B else C)\ for b have \U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S (V i))\ by auto then show ?thesis unfolding V_def by (metis (mono_tags, lifting) INF_UNIV_bool_expand) qed lemma cblinfun_image_eq: fixes S :: "'a::cbanach ccsubspace" and A B :: "'a::cbanach \\<^sub>C\<^sub>L'b::cbanach" assumes "\x. x \ G \ A *\<^sub>V x = B *\<^sub>V x" and "ccspan G \ S" shows "A *\<^sub>S S = B *\<^sub>S S" proof (use assms in transfer) fix G :: "'a set" and A :: "'a \ 'b" and B :: "'a \ 'b" and S :: "'a set" assume a1: "bounded_clinear A" assume a2: "bounded_clinear B" assume a3: "\x. x \ G \ A x = B x" assume a4: "S \ closure (cspan G)" have "A ` closure S = B ` closure S" by (smt (verit, best) UnCI a1 a2 a3 a4 bounded_clinear_eq_on closure_Un closure_closure image_cong sup.absorb_iff1) then show "closure (A ` S) = closure (B ` S)" by (metis Complex_Vector_Spaces0.bounded_clinear.bounded_linear a1 a2 closure_bounded_linear_image_subset_eq) qed lemma cblinfun_fixes_range: assumes "A o\<^sub>C\<^sub>L B = B" and "\ \ space_as_set (B *\<^sub>S top)" shows "A *\<^sub>V \ = \" proof- define rangeB rangeB' where "rangeB = space_as_set (B *\<^sub>S top)" and "rangeB' = range (cblinfun_apply B)" from assms have "\ \ closure rangeB'" by (simp add: cblinfun_image.rep_eq rangeB'_def top_ccsubspace.rep_eq) then obtain \i where \i_lim: "\i \ \" and \i_B: "\i i \ rangeB'" for i using closure_sequential by blast have A_invariant: "A *\<^sub>V \i i = \i i" for i proof- from \i_B obtain \ where \: "\i i = B *\<^sub>V \" using rangeB'_def by blast hence "A *\<^sub>V \i i = (A o\<^sub>C\<^sub>L B) *\<^sub>V \" by (simp add: cblinfun_compose.rep_eq) also have "\ = B *\<^sub>V \" by (simp add: assms) also have "\ = \i i" by (simp add: \) finally show ?thesis. qed from \i_lim have "(\i. A *\<^sub>V (\i i)) \ A *\<^sub>V \" by (rule isCont_tendsto_compose[rotated], simp) with A_invariant have "(\i. \i i) \ A *\<^sub>V \" by auto with \i_lim show "A *\<^sub>V \ = \" using LIMSEQ_unique by blast qed lemma zero_cblinfun_image[simp]: "0 *\<^sub>S S = (0::_ ccsubspace)" apply transfer by (simp add: complex_vector.subspace_0 image_constant[where x=0]) lemma cblinfun_image_INF_eq_general: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L'c::chilbert_space" and Uinv :: "'c \\<^sub>C\<^sub>L'b" assumes UinvUUinv: "Uinv o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L Uinv = Uinv" and UUinvU: "U o\<^sub>C\<^sub>L Uinv o\<^sub>C\<^sub>L U = U" \ \Meaning: \<^term>\Uinv\ is a Pseudoinverse of \<^term>\U\\ and V: "\i. V i \ Uinv *\<^sub>S top" shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof (rule antisym) show "U *\<^sub>S (INF i. V i) \ (INF i. U *\<^sub>S V i)" by (rule cblinfun_image_INF_leq) next define rangeU rangeUinv where "rangeU = U *\<^sub>S top" and "rangeUinv = Uinv *\<^sub>S top" define INFUV INFV where INFUV_def: "INFUV = (INF i. U *\<^sub>S V i)" and INFV_def: "INFV = (INF i. V i)" from assms have "V i \ rangeUinv" for i unfolding rangeUinv_def by simp moreover have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set rangeUinv" for \ using UinvUUinv cblinfun_fixes_range rangeUinv_def that by fastforce ultimately have "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>V \ = \" if "\ \ space_as_set (V i)" for \ i using less_eq_ccsubspace.rep_eq that by blast hence d1: "(Uinv o\<^sub>C\<^sub>L U) *\<^sub>S (V i) = (V i)" for i proof transfer show "closure ((Uinv \ U) ` V i) = V i" if "pred_fun \ closed_csubspace V" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ (Uinv \ U) \ = \" for V :: "'a \ 'b set" and Uinv :: "'c \ 'b" and U :: "'b \ 'c" and i :: 'a using that proof auto show "x \ V i" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ closure (V i)" for x :: 'b using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure (V i)" if "\x. closed_csubspace (V x)" and "bounded_clinear Uinv" and "bounded_clinear U" and "\\ i. \ \ V i \ Uinv (U \) = \" and "x \ V i" for x :: 'b using that using setdist_eq_0_sing_1 setdist_sing_in_set by blast qed qed have "U *\<^sub>S V i \ rangeU" for i by (simp add: cblinfun_image_mono rangeU_def) hence "INFUV \ rangeU" unfolding INFUV_def by (meson INF_lower UNIV_I order_trans) moreover have "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set rangeU" for \ using UUinvU cblinfun_fixes_range rangeU_def that by fastforce ultimately have x: "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>V \ = \" if "\ \ space_as_set INFUV" for \ by (simp add: in_mono less_eq_ccsubspace.rep_eq that) have "closure ((U \ Uinv) ` INFUV) = INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ (U \ Uinv) \ = \" for INFUV :: "'c set" and U :: "'b \ 'c" and Uinv :: "'c \ 'b" using that proof auto show "x \ INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ closure INFUV" for x :: 'c using that by (metis orthogonal_complement_of_closure closed_csubspace.subspace double_orthogonal_complement_id closure_is_closed_csubspace) show "x \ closure INFUV" if "closed_csubspace INFUV" and "bounded_clinear U" and "bounded_clinear Uinv" and "\\. \ \ INFUV \ U (Uinv \) = \" and "x \ INFUV" for x :: 'c using that using setdist_eq_0_sing_1 setdist_sing_in_set by (simp add: closed_csubspace.closed) qed hence "(U o\<^sub>C\<^sub>L Uinv) *\<^sub>S INFUV = INFUV" by (metis (mono_tags, opaque_lifting) x cblinfun_image.rep_eq cblinfun_image_id id_cblinfun_apply image_cong space_as_set_inject) hence "INFUV = U *\<^sub>S Uinv *\<^sub>S INFUV" by (simp add: cblinfun_compose_image) also have "\ \ U *\<^sub>S (INF i. Uinv *\<^sub>S U *\<^sub>S V i)" unfolding INFUV_def by (metis cblinfun_image_mono cblinfun_image_INF_leq) also have "\ = U *\<^sub>S INFV" using d1 by (metis (no_types, lifting) INFV_def cblinfun_assoc_left(2) image_cong) finally show "INFUV \ U *\<^sub>S INFV". qed lemma unitary_range[simp]: assumes "unitary U" shows "U *\<^sub>S top = top" using assms unfolding unitary_def apply transfer by (metis closure_UNIV comp_apply surj_def) lemma range_adjoint_isometry: assumes "isometry U" shows "U* *\<^sub>S top = top" proof- from assms have "top = U* *\<^sub>S U *\<^sub>S top" by (simp add: cblinfun_assoc_left(2)) also have "\ \ U* *\<^sub>S top" by (simp add: cblinfun_image_mono) finally show ?thesis using top.extremum_unique by blast qed lemma cblinfun_image_INF_eq[simp]: fixes V :: "'a \ 'b::chilbert_space ccsubspace" and U :: "'b \\<^sub>C\<^sub>L 'c::chilbert_space" assumes \isometry U\ shows "U *\<^sub>S (INF i. V i) = (INF i. U *\<^sub>S V i)" proof - from \isometry U\ have "U* o\<^sub>C\<^sub>L U o\<^sub>C\<^sub>L U* = U*" unfolding isometry_def by simp moreover from \isometry U\ have "U o\<^sub>C\<^sub>L U* o\<^sub>C\<^sub>L U = U" unfolding isometry_def by (simp add: cblinfun_compose_assoc) moreover have "V i \ U* *\<^sub>S top" for i by (simp add: range_adjoint_isometry assms) ultimately show ?thesis by (rule cblinfun_image_INF_eq_general) qed lemma isometry_cblinfun_image_inf_distrib[simp]: fixes U::\'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ and X Y::"'a ccsubspace" assumes "isometry U" shows "U *\<^sub>S (inf X Y) = inf (U *\<^sub>S X) (U *\<^sub>S Y)" using cblinfun_image_INF_eq[where V="\b. if b then X else Y" and U=U] unfolding INF_UNIV_bool_expand using assms by auto lemma cblinfun_image_ccspan: shows "A *\<^sub>S ccspan G = ccspan ((*\<^sub>V) A ` G)" apply transfer by (simp add: bounded_clinear.bounded_linear bounded_clinear_def closure_bounded_linear_image_subset_eq complex_vector.linear_span_image) lemma cblinfun_apply_in_image[simp]: "A *\<^sub>V \ \ space_as_set (A *\<^sub>S \)" by (metis cblinfun_image.rep_eq closure_subset in_mono range_eqI top_ccsubspace.rep_eq) lemma cblinfun_plus_image_distr: \(A + B) *\<^sub>S S \ A *\<^sub>S S \ B *\<^sub>S S\ apply transfer by (smt (verit, ccfv_threshold) closed_closure closed_sum_def closure_minimal closure_subset image_subset_iff set_plus_intro subset_eq) lemma cblinfun_sum_image_distr: \(\i\I. A i) *\<^sub>S S \ (SUP i\I. A i *\<^sub>S S)\ proof (cases \finite I\) case True then show ?thesis proof induction case empty then show ?case by auto next case (insert x F) then show ?case apply auto by (smt (z3) cblinfun_plus_image_distr inf_sup_aci(6) le_iff_sup) qed next case False then show ?thesis by auto qed subsection \Sandwiches\ lift_definition sandwich :: \('a::chilbert_space \\<^sub>C\<^sub>L 'b::complex_inner) \ (('a \\<^sub>C\<^sub>L 'a) \\<^sub>C\<^sub>L ('b \\<^sub>C\<^sub>L 'b))\ is \\(A::'a\\<^sub>C\<^sub>L'b) B. A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ proof fix A :: \'a \\<^sub>C\<^sub>L 'b\ and B B1 B2 :: \'a \\<^sub>C\<^sub>L 'a\ and c :: complex show \A o\<^sub>C\<^sub>L (B1 + B2) o\<^sub>C\<^sub>L A* = (A o\<^sub>C\<^sub>L B1 o\<^sub>C\<^sub>L A*) + (A o\<^sub>C\<^sub>L B2 o\<^sub>C\<^sub>L A*)\ by (simp add: cblinfun_compose_add_left cblinfun_compose_add_right) show \A o\<^sub>C\<^sub>L (c *\<^sub>C B) o\<^sub>C\<^sub>L A* = c *\<^sub>C (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*)\ by auto show \\K. \B. norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * K\ proof (rule exI[of _ \norm A * norm (A*)\], rule allI) fix B have \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by blast also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) finally show \norm (A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*) \ norm B * (norm A * norm (A*))\ by (simp add: mult.assoc vector_space_over_itself.scale_left_commute) qed qed lemma sandwich_0[simp]: \sandwich 0 = 0\ by (simp add: cblinfun_eqI sandwich.rep_eq) lemma sandwich_apply: \sandwich A *\<^sub>V B = A o\<^sub>C\<^sub>L B o\<^sub>C\<^sub>L A*\ apply (transfer fixing: A B) by auto lemma norm_sandwich: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'a::{chilbert_space} \\<^sub>C\<^sub>L 'b::{complex_inner}\ proof - have main: \norm (sandwich A) = (norm A)\<^sup>2\ for A :: \'c::{chilbert_space,not_singleton} \\<^sub>C\<^sub>L 'd::{complex_inner}\ proof (rule norm_cblinfun_eqI) show \(norm A)\<^sup>2 \ norm (sandwich A *\<^sub>V id_cblinfun) / norm (id_cblinfun :: 'c \\<^sub>C\<^sub>L _)\ apply (auto simp: sandwich_apply) by - fix B have \norm (sandwich A *\<^sub>V B) \ norm (A o\<^sub>C\<^sub>L B) * norm (A*)\ using norm_cblinfun_compose by (auto simp: sandwich_apply simp del: norm_adj) also have \\ \ (norm A * norm B) * norm (A*)\ by (simp add: mult_right_mono norm_cblinfun_compose) also have \\ \ (norm A)\<^sup>2 * norm B\ by (simp add: power2_eq_square mult.assoc vector_space_over_itself.scale_left_commute) finally show \norm (sandwich A *\<^sub>V B) \ (norm A)\<^sup>2 * norm B\ by - show \0 \ (norm A)\<^sup>2\ by auto qed show ?thesis proof (cases \class.not_singleton TYPE('a)\) case True show ?thesis apply (rule main[internalize_sort' 'c2]) apply standard[1] using True by simp next case False have \A = 0\ apply (rule cblinfun_from_CARD_1_0[internalize_sort' 'a]) apply (rule not_singleton_vs_CARD_1) apply (rule False) by standard then show ?thesis by simp qed qed lemma sandwich_apply_adj: \sandwich A (B*) = (sandwich A B)*\ by (simp add: cblinfun_assoc_left(1) sandwich_apply) lemma sandwich_id[simp]: "sandwich id_cblinfun = id_cblinfun" apply (rule cblinfun_eqI) by (auto simp: sandwich_apply) subsection \Projectors\ lift_definition Proj :: "('a::chilbert_space) ccsubspace \ 'a \\<^sub>C\<^sub>L'a" is \projection\ by (rule projection_bounded_clinear) lemma Proj_range[simp]: "Proj S *\<^sub>S top = S" proof transfer fix S :: \'a set\ assume \closed_csubspace S\ then have "closure (range (projection S)) \ S" by (metis closed_csubspace.closed closed_csubspace.subspace closure_closed complex_vector.subspace_0 csubspace_is_convex dual_order.eq_iff insert_absorb insert_not_empty projection_image) moreover have "S \ closure (range (projection S))" using \closed_csubspace S\ by (metis closed_csubspace_def closure_subset csubspace_is_convex equals0D projection_image subset_iff) ultimately show \closure (range (projection S)) = S\ by auto qed lemma adj_Proj: \(Proj M)* = Proj M\ apply transfer by (simp add: projection_cadjoint) lemma Proj_idempotent[simp]: \Proj M o\<^sub>C\<^sub>L Proj M = Proj M\ proof - have u1: \(cblinfun_apply (Proj M)) = projection (space_as_set M)\ apply transfer by blast have \closed_csubspace (space_as_set M)\ using space_as_set by auto hence u2: \(projection (space_as_set M))\(projection (space_as_set M)) = (projection (space_as_set M))\ using projection_idem by fastforce have \(cblinfun_apply (Proj M)) \ (cblinfun_apply (Proj M)) = cblinfun_apply (Proj M)\ using u1 u2 by simp hence \cblinfun_apply ((Proj M) o\<^sub>C\<^sub>L (Proj M)) = cblinfun_apply (Proj M)\ by (simp add: cblinfun_compose.rep_eq) thus ?thesis using cblinfun_apply_inject by auto qed lift_definition is_Proj::\'a::chilbert_space \\<^sub>C\<^sub>L 'a \ bool\ is \\P. \M. closed_csubspace M \ is_projection_on P M\ . lemma Proj_on_own_range': fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \P o\<^sub>C\<^sub>L P = P\ and \P = P*\ shows \Proj (P *\<^sub>S top) = P\ proof- define M where "M = P *\<^sub>S top" have v3: "x \ (\x. x - P *\<^sub>V x) -` {0}" if "x \ range (cblinfun_apply P)" for x :: 'a proof- have v3_1: \cblinfun_apply P \ cblinfun_apply P = cblinfun_apply P\ by (metis \P o\<^sub>C\<^sub>L P = P\ cblinfun_compose.rep_eq) have \\t. P *\<^sub>V t = x\ using that by blast then obtain t where t_def: \P *\<^sub>V t = x\ by blast hence \x - P *\<^sub>V x = x - P *\<^sub>V (P *\<^sub>V t)\ by simp also have \\ = x - (P *\<^sub>V t)\ using v3_1 by (metis comp_apply) also have \\ = 0\ by (simp add: t_def) finally have \x - P *\<^sub>V x = 0\ by blast thus ?thesis by simp qed have v1: "range (cblinfun_apply P) \ (\x. x - cblinfun_apply P x) -` {0}" using v3 by blast have "x \ range (cblinfun_apply P)" if "x \ (\x. x - P *\<^sub>V x) -` {0}" for x :: 'a proof- have x1:\x - P *\<^sub>V x = 0\ using that by blast have \x = P *\<^sub>V x\ by (simp add: x1 eq_iff_diff_eq_0) thus ?thesis by blast qed hence v2: "(\x. x - cblinfun_apply P x) -` {0} \ range (cblinfun_apply P)" by blast have i1: \range (cblinfun_apply P) = (\ x. x - cblinfun_apply P x) -` {0}\ using v1 v2 by (simp add: v1 dual_order.antisym) have p1: \closed {(0::'a)}\ by simp have p2: \continuous (at x) (\ x. x - P *\<^sub>V x)\ for x proof- have \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by (simp add: id_cblinfun.rep_eq minus_cblinfun.rep_eq) hence \bounded_clinear (cblinfun_apply (id_cblinfun - P))\ using cblinfun_apply by blast hence \continuous (at x) (cblinfun_apply (id_cblinfun - P))\ by (simp add: clinear_continuous_at) thus ?thesis using \cblinfun_apply (id_cblinfun - P) = (\ x. x - P *\<^sub>V x)\ by simp qed have i2: \closed ( (\ x. x - P *\<^sub>V x) -` {0} )\ using p1 p2 by (rule Abstract_Topology.continuous_closed_vimage) have \closed (range (cblinfun_apply P))\ using i1 i2 by simp have u2: \cblinfun_apply P x \ space_as_set M\ for x by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq top_ccsubspace.rep_eq) have xy: \\ x - P *\<^sub>V x, y \ = 0\ if y1: \y \ space_as_set M\ for x y proof- have \\t. y = P *\<^sub>V t\ using y1 by (simp add: M_def \closed (range ((*\<^sub>V) P))\ cblinfun_image.rep_eq image_iff top_ccsubspace.rep_eq) then obtain t where t_def: \y = P *\<^sub>V t\ by blast have \\ x - P *\<^sub>V x, y \ = \ x - P *\<^sub>V x, P *\<^sub>V t \\ by (simp add: t_def) also have \\ = \ P *\<^sub>V (x - P *\<^sub>V x), t \\ by (metis \P = P*\ cinner_adj_left) also have \\ = \ P *\<^sub>V x - P *\<^sub>V (P *\<^sub>V x), t \\ by (simp add: cblinfun.diff_right) also have \\ = \ P *\<^sub>V x - P *\<^sub>V x, t \\ by (metis assms(1) comp_apply cblinfun_compose.rep_eq) also have \\ = \ 0, t \\ by simp also have \\ = 0\ by simp finally show ?thesis by blast qed hence u1: \x - P *\<^sub>V x \ orthogonal_complement (space_as_set M)\ for x by (simp add: orthogonal_complementI) have "closed_csubspace (space_as_set M)" using space_as_set by auto hence f1: "(Proj M) *\<^sub>V a = P *\<^sub>V a" for a by (simp add: Proj.rep_eq projection_eqI u1 u2) have "(+) ((P - Proj M) *\<^sub>V a) = id" for a using f1 by (auto intro!: ext simp add: minus_cblinfun.rep_eq) hence "b - b = cblinfun_apply (P - Proj M) a" for a b by (metis (no_types) add_diff_cancel_right' id_apply) hence "cblinfun_apply (id_cblinfun - (P - Proj M)) a = a" for a by (simp add: minus_cblinfun.rep_eq) thus ?thesis using u1 u2 cblinfun_apply_inject diff_diff_eq2 diff_eq_diff_eq eq_id_iff id_cblinfun.rep_eq by (metis (no_types, opaque_lifting) M_def) qed lemma Proj_range_closed: assumes "is_Proj P" shows "closed (range (cblinfun_apply P))" using assms apply transfer using closed_csubspace.closed is_projection_on_image by blast lemma Proj_is_Proj[simp]: fixes M::\'a::chilbert_space ccsubspace\ shows \is_Proj (Proj M)\ proof- have u1: "closed_csubspace (space_as_set M)" using space_as_set by blast have v1: "h - Proj M *\<^sub>V h \ orthogonal_complement (space_as_set M)" for h by (simp add: Proj.rep_eq orthogonal_complementI projection_orthogonal u1) have v2: "Proj M *\<^sub>V h \ space_as_set M" for h by (metis Proj.rep_eq mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) have u2: "is_projection_on ((*\<^sub>V) (Proj M)) (space_as_set M)" unfolding is_projection_on_def by (simp add: smallest_dist_is_ortho u1 v1 v2) show ?thesis using u1 u2 is_Proj.rep_eq by blast qed lemma is_Proj_algebraic: fixes P::\'a::chilbert_space \\<^sub>C\<^sub>L 'a\ shows \is_Proj P \ P o\<^sub>C\<^sub>L P = P \ P = P*\ proof have "P o\<^sub>C\<^sub>L P = P" if "is_Proj P" using that apply transfer using is_projection_on_idem by fastforce moreover have "P = P*" if "is_Proj P" using that apply transfer by (metis is_projection_on_cadjoint) ultimately show "P o\<^sub>C\<^sub>L P = P \ P = P*" if "is_Proj P" using that by blast show "is_Proj P" if "P o\<^sub>C\<^sub>L P = P \ P = P*" using that Proj_on_own_range' Proj_is_Proj by metis qed lemma Proj_on_own_range: fixes P :: \'a::chilbert_space \\<^sub>C\<^sub>L'a\ assumes \is_Proj P\ shows \Proj (P *\<^sub>S top) = P\ using Proj_on_own_range' assms is_Proj_algebraic by blast lemma Proj_image_leq: "(Proj S) *\<^sub>S A \ S" by (metis Proj_range inf_top_left le_inf_iff isometry_cblinfun_image_inf_distrib') lemma Proj_sandwich: fixes A::"'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space" assumes "isometry A" shows "sandwich A *\<^sub>V Proj S = Proj (A *\<^sub>S S)" proof- define P where \P = A o\<^sub>C\<^sub>L Proj S o\<^sub>C\<^sub>L (A*)\ have \P o\<^sub>C\<^sub>L P = P\ using assms unfolding P_def isometry_def by (metis (no_types, lifting) Proj_idempotent cblinfun_assoc_left(1) cblinfun_compose_id_left) moreover have \P = P*\ unfolding P_def by (metis adj_Proj adj_cblinfun_compose cblinfun_assoc_left(1) double_adj) ultimately have \\M. P = Proj M \ space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ using P_def Proj_on_own_range' by (metis Proj_is_Proj Proj_range_closed cblinfun_image.rep_eq closure_closed top_ccsubspace.rep_eq) then obtain M where \P = Proj M\ and \space_as_set M = range (cblinfun_apply (A o\<^sub>C\<^sub>L (Proj S) o\<^sub>C\<^sub>L (A*)))\ by blast have f1: "A o\<^sub>C\<^sub>L Proj S = P o\<^sub>C\<^sub>L A" by (simp add: P_def assms cblinfun_compose_assoc) hence "P o\<^sub>C\<^sub>L A o\<^sub>C\<^sub>L A* = P" using P_def by presburger hence "(P o\<^sub>C\<^sub>L A) *\<^sub>S (c \ A* *\<^sub>S d) = P *\<^sub>S (A *\<^sub>S c \ d)" for c d by (simp add: cblinfun_assoc_left(2)) hence "P *\<^sub>S (A *\<^sub>S \ \ c) = (P o\<^sub>C\<^sub>L A) *\<^sub>S \" for c by (metis sup_top_left) hence \M = A *\<^sub>S S\ using f1 by (metis \P = Proj M\ cblinfun_assoc_left(2) Proj_range sup_top_right) thus ?thesis using \P = Proj M\ unfolding P_def sandwich_apply by blast qed lemma Proj_orthog_ccspan_union: assumes "\x y. x \ X \ y \ Y \ is_orthogonal x y" shows \Proj (ccspan (X \ Y)) = Proj (ccspan X) + Proj (ccspan Y)\ proof - have \x \ cspan X \ y \ cspan Y \ is_orthogonal x y\ for x y apply (rule is_orthogonal_closure_cspan[where X=X and Y=Y]) using closure_subset assms by auto then have \x \ closure (cspan X) \ y \ closure (cspan Y) \ is_orthogonal x y\ for x y by (metis orthogonal_complementI orthogonal_complement_of_closure orthogonal_complement_orthoI') then show ?thesis apply (transfer fixing: X Y) apply (subst projection_plus[symmetric]) by auto qed abbreviation proj :: "'a::chilbert_space \ 'a \\<^sub>C\<^sub>L 'a" where "proj \ \ Proj (ccspan {\})" lemma proj_0[simp]: \proj 0 = 0\ apply transfer by auto lemma surj_isometry_is_unitary: fixes U :: \'a::chilbert_space \\<^sub>C\<^sub>L 'b::chilbert_space\ assumes \isometry U\ assumes \U *\<^sub>S \ = \\ shows \unitary U\ by (metis Proj_sandwich sandwich_apply Proj_on_own_range' assms(1) assms(2) cblinfun_compose_id_right isometry_def unitary_def unitary_id unitary_range) lemma ccsubspace_supI_via_Proj: fixes A B C::"'a::chilbert_space ccsubspace" assumes a1: \Proj (- C) *\<^sub>S A \ B\ shows "A \ sup B C" proof- have x2: \x \ space_as_set B\ if "x \ closure ( (projection (orthogonal_complement (space_as_set C))) ` space_as_set A)" for x using that by (metis Proj.rep_eq cblinfun_image.rep_eq assms less_eq_ccsubspace.rep_eq subsetD uminus_ccsubspace.rep_eq) have q1: \x \ closure {\ + \ |\ \. \ \ space_as_set B \ \ \ space_as_set C}\ if \x \ space_as_set A\ for x proof- have p1: \closed_csubspace (space_as_set C)\ using space_as_set by auto hence \x = (projection (space_as_set C)) x + (projection (orthogonal_complement (space_as_set C))) x\ by simp hence \x = (projection (orthogonal_complement (space_as_set C))) x + (projection (space_as_set C)) x\ by (metis ordered_field_class.sign_simps(2)) moreover have \(projection (orthogonal_complement (space_as_set C))) x \ space_as_set B\ using x2 by (meson closure_subset image_subset_iff that) moreover have \(projection (space_as_set C)) x \ space_as_set C\ by (metis mem_Collect_eq orthog_proj_exists projection_eqI space_as_set) ultimately show ?thesis using closure_subset by fastforce qed have x1: \x \ (space_as_set B +\<^sub>M space_as_set C)\ if "x \ space_as_set A" for x proof - have f1: "x \ closure {a + b |a b. a \ space_as_set B \ b \ space_as_set C}" by (simp add: q1 that) have "{a + b |a b. a \ space_as_set B \ b \ space_as_set C} = {a. \p. p \ space_as_set B \ (\q. q \ space_as_set C \ a = p + q)}" by blast hence "x \ closure {a. \b\space_as_set B. \c\space_as_set C. a = b + c}" using f1 by (simp add: Bex_def_raw) thus ?thesis using that unfolding closed_sum_def set_plus_def by blast qed hence \x \ space_as_set (Abs_clinear_space (space_as_set B +\<^sub>M space_as_set C))\ if "x \ space_as_set A" for x using that by (metis space_as_set_inverse sup_ccsubspace.rep_eq) thus ?thesis by (simp add: x1 less_eq_ccsubspace.rep_eq subset_eq sup_ccsubspace.rep_eq) qed lemma is_Proj_idempotent: assumes "is_Proj P" shows "P o\<^sub>C\<^sub>L P = P" using assms unfolding is_Proj_def using assms is_Proj_algebraic by auto lemma is_proj_selfadj: assumes "is_Proj P" shows "P* = P" using assms unfolding is_Proj_def by (metis is_Proj_algebraic is_Proj_def) lemma is_Proj_I: assumes "P o\<^sub>C\<^sub>L P = P" and "P* = P" shows "is_Proj P" using assms is_Proj_algebraic by metis lemma is_Proj_0[simp]: "is_Proj 0" by (metis add_left_cancel adj_plus bounded_cbilinear.zero_left bounded_cbilinear_cblinfun_compose group_cancel.rule0 is_Proj_I) lemma is_Proj_complement[simp]: assumes a1: "is_Proj P" shows "is_Proj (id_cblinfun-P)" by (smt (z3) add_diff_cancel_left add_diff_cancel_left' adj_cblinfun_compose adj_plus assms bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel id_cblinfun_adjoint is_Proj_algebraic cblinfun_compose_id_left) lemma Proj_bot[simp]: "Proj bot = 0" by (metis zero_cblinfun_image Proj_on_own_range' is_Proj_0 is_Proj_algebraic zero_ccsubspace_def) lemma Proj_ortho_compl: "Proj (- X) = id_cblinfun - Proj X" by (transfer , auto) lemma Proj_inj: assumes "Proj X = Proj Y" shows "X = Y" by (metis assms Proj_range) subsection \Kernel\ lift_definition kernel :: "'a::complex_normed_vector \\<^sub>C\<^sub>L'b::complex_normed_vector \ 'a ccsubspace" is "\ f. f -` {0}" by (metis kernel_is_closed_csubspace) definition eigenspace :: "complex \ 'a::complex_normed_vector \\<^sub>C\<^sub>L'a \ 'a ccsubspace" where "eigenspace a A = kernel (A - a *\<^sub>C id_cblinfun)" lemma kernel_scaleC[simp]: "a\0 \ kernel (a *\<^sub>C A) = kernel A" for a :: complex and A :: "(_,_) cblinfun" apply transfer using complex_vector.scale_eq_0_iff by blast lemma kernel_0[simp]: "kernel 0 = top" apply transfer by auto lemma kernel_id[simp]: "kernel id_cblinfun = 0" apply transfer by simp lemma eigenspace_scaleC[simp]: assumes a1: "a \ 0" shows "eigenspace b (a *\<^sub>C A) = eigenspace (b/a) A" proof - have "b *\<^sub>C (id_cblinfun::('a, _) cblinfun) = a *\<^sub>C (b / a) *\<^sub>C id_cblinfun" using a1 by (metis ceq_vector_fraction_iff) hence "kernel (a *\<^sub>C A - b *\<^sub>C id_cblinfun) = kernel (A - (b / a) *\<^sub>C id_cblinfun)" using a1 by (metis (no_types) complex_vector.scale_right_diff_distrib kernel_scaleC) thus ?thesis unfolding eigenspace_def by blast qed lemma eigenspace_memberD: assumes "x \ space_as_set (eigenspace e A)" shows "A *\<^sub>V x = e *\<^sub>C x" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberD: assumes "x \ space_as_set (kernel A)" shows "A *\<^sub>V x = 0" using assms apply transfer by auto lemma eigenspace_memberI: assumes "A *\<^sub>V x = e *\<^sub>C x" shows "x \ space_as_set (eigenspace e A)" using assms unfolding eigenspace_def apply transfer by auto lemma kernel_memberI: assumes "A *\<^sub>V x = 0" shows "x \ space_as_set (kernel A)" using assms apply transfer by auto subsection \Isomorphisms and inverses\ definition iso_cblinfun :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ bool\ where \iso_cblinfun A = (\ B. A o\<^sub>C\<^sub>L B = id_cblinfun \ B o\<^sub>C\<^sub>L A = id_cblinfun)\ definition cblinfun_inv :: \('a::complex_normed_vector, 'b::complex_normed_vector) cblinfun \ ('b,'a) cblinfun\ where \cblinfun_inv A = (SOME B. B o\<^sub>C\<^sub>L A = id_cblinfun)\ lemma assumes \iso_cblinfun A\ shows cblinfun_inv_left: \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and cblinfun_inv_right: \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ proof - from assms obtain B where AB: \A o\<^sub>C\<^sub>L B = id_cblinfun\ and BA: \B o\<^sub>C\<^sub>L A = id_cblinfun\ using iso_cblinfun_def by blast from BA have \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ by (metis (mono_tags, lifting) cblinfun_inv_def someI_ex) with AB BA have \cblinfun_inv A = B\ by (metis cblinfun_assoc_left(1) cblinfun_compose_id_right) with AB BA show \cblinfun_inv A o\<^sub>C\<^sub>L A = id_cblinfun\ and \A o\<^sub>C\<^sub>L cblinfun_inv A = id_cblinfun\ by auto qed lemma cblinfun_inv_uniq: assumes "A o\<^sub>C\<^sub>L B = id_cblinfun" and "B o\<^sub>C\<^sub>L A = id_cblinfun" shows "cblinfun_inv A = B" using assms by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_inv_left iso_cblinfun_def) subsection \One-dimensional spaces\ instantiation cblinfun :: (one_dim, one_dim) complex_inner begin text \Once we have a theory for the trace, we could instead define the Hilbert-Schmidt inner product and relax the \<^class>\one_dim\-sort constraint to (\<^class>\cfinite_dim\,\<^class>\complex_normed_vector\) or similar\ definition "cinner_cblinfun (A::'a \\<^sub>C\<^sub>L 'b) (B::'a \\<^sub>C\<^sub>L 'b) = cnj (one_dim_iso (A *\<^sub>V 1)) * one_dim_iso (B *\<^sub>V 1)" instance proof intro_classes fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "\A, B\ = cnj \B, A\" unfolding cinner_cblinfun_def by auto show "\A + B, C\ = \A, C\ + \B, C\" by (simp add: cinner_cblinfun_def algebra_simps plus_cblinfun.rep_eq) show "\c *\<^sub>C A, B\ = cnj c * \A, B\" by (simp add: cblinfun.scaleC_left cinner_cblinfun_def) show "0 \ \A, A\" unfolding cinner_cblinfun_def by auto have "bounded_clinear A \ A 1 = 0 \ A = (\_. 0)" for A::"'a \ 'b" proof (rule one_dim_clinear_eqI [where x = 1] , auto) show "clinear A" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: bounded_clinear.clinear) show "clinear ((\_. 0)::'a \ 'b)" if "bounded_clinear A" and "A 1 = 0" for A :: "'a \ 'b" using that by (simp add: complex_vector.module_hom_zero) qed hence "A *\<^sub>V 1 = 0 \ A = 0" by transfer hence "one_dim_iso (A *\<^sub>V 1) = 0 \ A = 0" by (metis one_dim_iso_of_zero one_dim_iso_inj) thus "(\A, A\ = 0) = (A = 0)" by (auto simp: cinner_cblinfun_def) show "norm A = sqrt (cmod \A, A\)" unfolding cinner_cblinfun_def apply transfer by (simp add: norm_mult abs_complex_def one_dim_onorm' cnj_x_x power2_eq_square bounded_clinear.clinear) qed end instantiation cblinfun :: (one_dim, one_dim) one_dim begin lift_definition one_cblinfun :: "'a \\<^sub>C\<^sub>L 'b" is "one_dim_iso" by (rule bounded_clinear_one_dim_iso) lift_definition times_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f g. f o one_dim_iso o g" by (simp add: comp_bounded_clinear) lift_definition inverse_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" is "\f. ((*) (one_dim_iso (inverse (f 1)))) o one_dim_iso" by (auto intro!: comp_bounded_clinear bounded_clinear_mult_right) definition divide_cblinfun :: "'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b \ 'a \\<^sub>C\<^sub>L 'b" where "divide_cblinfun A B = A * inverse B" definition "canonical_basis_cblinfun = [1 :: 'a \\<^sub>C\<^sub>L 'b]" instance proof intro_classes let ?basis = "canonical_basis :: ('a \\<^sub>C\<^sub>L 'b) list" fix A B C :: "'a \\<^sub>C\<^sub>L 'b" and c c' :: complex show "distinct ?basis" unfolding canonical_basis_cblinfun_def by simp have "(1::'a \\<^sub>C\<^sub>L 'b) \ (0::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_one zero_neq_one) thus "cindependent (set ?basis)" unfolding canonical_basis_cblinfun_def by simp have "A \ cspan (set ?basis)" for A proof - define c :: complex where "c = one_dim_iso (A *\<^sub>V 1)" have "A x = one_dim_iso (A 1) *\<^sub>C one_dim_iso x" for x by (smt (z3) cblinfun.scaleC_right complex_vector.scale_left_commute one_dim_iso_idem one_dim_scaleC_1) hence "A = one_dim_iso (A *\<^sub>V 1) *\<^sub>C 1" apply transfer by metis thus "A \ cspan (set ?basis)" unfolding canonical_basis_cblinfun_def by (smt complex_vector.span_base complex_vector.span_scale list.set_intros(1)) qed thus "cspan (set ?basis) = UNIV" by auto have "A = (1::'a \\<^sub>C\<^sub>L 'b) \ norm (1::'a \\<^sub>C\<^sub>L 'b) = (1::real)" apply transfer by simp thus "A \ set ?basis \ norm A = 1" unfolding canonical_basis_cblinfun_def by simp show "?basis = [1]" unfolding canonical_basis_cblinfun_def by simp show "c *\<^sub>C 1 * c' *\<^sub>C 1 = (c * c') *\<^sub>C (1::'a\\<^sub>C\<^sub>L'b)" apply transfer by auto have "(1::'a \\<^sub>C\<^sub>L 'b) = (0::'a \\<^sub>C\<^sub>L 'b) \ False" by (metis cblinfun.zero_left one_cblinfun.rep_eq one_dim_iso_of_zero' zero_neq_neg_one) thus "is_ortho_set (set ?basis)" unfolding is_ortho_set_def canonical_basis_cblinfun_def by auto show "A div B = A * inverse B" by (simp add: divide_cblinfun_def) show "inverse (c *\<^sub>C 1) = (1::'a\\<^sub>C\<^sub>L'b) /\<^sub>C c" apply transfer by (simp add: o_def one_dim_inverse) qed end lemma id_cblinfun_eq_1[simp]: \id_cblinfun = 1\ apply transfer by auto lemma one_dim_apply_is_times[simp]: fixes A :: "'a::one_dim \\<^sub>C\<^sub>L 'a" and B :: "'a \\<^sub>C\<^sub>L 'a" shows "A o\<^sub>C\<^sub>L B = A * B" apply transfer by simp lemma one_comp_one_cblinfun[simp]: "1 o\<^sub>C\<^sub>L 1 = 1" apply transfer unfolding o_def by simp lemma one_cblinfun_adj[simp]: "1* = 1" apply transfer by simp lemma scaleC_1_right[simp]: \scaleC x (1::'a::one_dim) = of_complex x\ unfolding of_complex_def by simp lemma scaleC_of_complex[simp]: \scaleC x (of_complex y) = of_complex (x * y)\ unfolding of_complex_def using scaleC_scaleC by blast lemma scaleC_1_apply[simp]: \(x *\<^sub>C 1) *\<^sub>V y = x *\<^sub>C y\ by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_apply_1_left[simp]: \1 *\<^sub>V y = y\ by (metis cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma of_complex_cblinfun_apply[simp]: \of_complex x *\<^sub>V y = x *\<^sub>C y\ unfolding of_complex_def by (metis cblinfun.scaleC_left cblinfun_id_cblinfun_apply id_cblinfun_eq_1) lemma cblinfun_compose_1_left[simp]: \1 o\<^sub>C\<^sub>L x = x\ apply transfer by auto lemma cblinfun_compose_1_right[simp]: \x o\<^sub>C\<^sub>L 1 = x\ apply transfer by auto lemma one_dim_iso_id_cblinfun: \one_dim_iso id_cblinfun = id_cblinfun\ by simp lemma one_dim_iso_id_cblinfun_eq_1: \one_dim_iso id_cblinfun = 1\ by simp lemma one_dim_iso_comp_distr[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a o\<^sub>C\<^sub>L one_dim_iso b\ by (smt (z3) cblinfun_compose_scaleC_left cblinfun_compose_scaleC_right one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC) lemma one_dim_iso_comp_distr_times[simp]: \one_dim_iso (a o\<^sub>C\<^sub>L b) = one_dim_iso a * one_dim_iso b\ by (smt (verit, del_insts) mult.left_neutral mult_scaleC_left one_cinner_a_scaleC_one one_comp_one_cblinfun one_dim_iso_of_one one_dim_iso_scaleC cblinfun_compose_scaleC_right cblinfun_compose_scaleC_left) lemma one_dim_iso_adjoint[simp]: \one_dim_iso (A*) = (one_dim_iso A)*\ by (smt (z3) one_cblinfun_adj one_cinner_a_scaleC_one one_dim_iso_of_one one_dim_iso_scaleC scaleC_adj) lemma one_dim_iso_adjoint_complex[simp]: \one_dim_iso (A*) = cnj (one_dim_iso A)\ by (metis (mono_tags, lifting) one_cblinfun_adj one_dim_iso_idem one_dim_scaleC_1 scaleC_adj) lemma one_dim_cblinfun_compose_commute: \a o\<^sub>C\<^sub>L b = b o\<^sub>C\<^sub>L a\ for a b :: \('a::one_dim,'a) cblinfun\ by (simp add: one_dim_iso_inj) lemma one_cblinfun_apply_one[simp]: \1 *\<^sub>V 1 = 1\ by (simp add: one_cblinfun.rep_eq) subsection \Loewner order\ lift_definition heterogenous_cblinfun_id :: \'a::complex_normed_vector \\<^sub>C\<^sub>L 'b::complex_normed_vector\ is \if bounded_clinear (heterogenous_identity :: 'a::complex_normed_vector \ 'b::complex_normed_vector) then heterogenous_identity else (\_. 0)\ by auto lemma heterogenous_cblinfun_id_def'[simp]: "heterogenous_cblinfun_id = id_cblinfun" apply transfer by auto definition "heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'b::chilbert_space itself) \ unitary (heterogenous_cblinfun_id :: 'a \\<^sub>C\<^sub>L 'b) \ unitary (heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a)" lemma heterogenous_same_type_cblinfun[simp]: \heterogenous_same_type_cblinfun (x::'a::chilbert_space itself) (y::'a::chilbert_space itself)\ unfolding heterogenous_same_type_cblinfun_def by auto instantiation cblinfun :: (chilbert_space, chilbert_space) ord begin definition less_eq_cblinfun :: \('a \\<^sub>C\<^sub>L 'b) \ ('a \\<^sub>C\<^sub>L 'b) \ bool\ where less_eq_cblinfun_def_heterogenous: \less_eq_cblinfun A B = (if heterogenous_same_type_cblinfun TYPE('a) TYPE('b) then \\::'b. cinner \ ((B-A) *\<^sub>V heterogenous_cblinfun_id *\<^sub>V \) \ 0 else (A=B))\ definition \less_cblinfun (A :: 'a \\<^sub>C\<^sub>L 'b) B \ A \ B \ \ B \ A\ instance.. end lemma less_eq_cblinfun_def: \A \ B \ (\\. cinner \ (A *\<^sub>V \) \ cinner \ (B *\<^sub>V \))\ unfolding less_eq_cblinfun_def_heterogenous by (auto simp del: less_eq_complex_def simp: cblinfun.diff_left cinner_diff_right) instantiation cblinfun :: (chilbert_space, chilbert_space) ordered_complex_vector begin instance proof intro_classes note less_eq_complex_def[simp del] fix x y z :: \'a \\<^sub>C\<^sub>L 'b\ fix a b :: complex define pos where \pos X \ (\\. cinner \ (X *\<^sub>V \) \ 0)\ for X :: \'b \\<^sub>C\<^sub>L 'b\ consider (unitary) \heterogenous_same_type_cblinfun TYPE('a) TYPE('b)\ \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B = pos ((B-A) o\<^sub>C\<^sub>L (heterogenous_cblinfun_id :: 'b\\<^sub>C\<^sub>L'a))\ | (trivial) \\A B :: 'a \\<^sub>C\<^sub>L 'b. A \ B \ A = B\ apply atomize_elim by (auto simp: pos_def less_eq_cblinfun_def_heterogenous) note cases = this have [simp]: \pos 0\ unfolding pos_def by auto have pos_nondeg: \X = 0\ if \pos X\ and \pos (-X)\ for X apply (rule cblinfun_cinner_eqI, simp) using that by (metis (no_types, lifting) cblinfun.minus_left cinner_minus_right dual_order.antisym equation_minus_iff neg_le_0_iff_le pos_def) have pos_add: \pos (X+Y)\ if \pos X\ and \pos Y\ for X Y by (smt (z3) pos_def cblinfun.diff_left cinner_minus_right cinner_simps(3) diff_ge_0_iff_ge diff_minus_eq_add neg_le_0_iff_le order_trans that(1) that(2) uminus_cblinfun.rep_eq) have pos_scaleC: \pos (a *\<^sub>C X)\ if \a\0\ and \pos X\ for X a using that unfolding pos_def by (auto simp: cblinfun.scaleC_left) let ?id = \heterogenous_cblinfun_id :: 'b \\<^sub>C\<^sub>L 'a\ show \x \ x\ apply (cases rule:cases) by auto show \(x < y) \ (x \ y \ \ y \ x)\ unfolding less_cblinfun_def by simp show \x \ z\ if \x \ y\ and \y \ z\ proof (cases rule:cases) case unitary define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ and \b = (z-y) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ with unitary that have \pos a\ and \pos b\ by auto then have \pos (a + b)\ by (rule pos_add) moreover have \a + b = (z - x) o\<^sub>C\<^sub>L heterogenous_cblinfun_id\ unfolding a_def b_def by (metis (no_types, lifting) bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose diff_add_cancel ordered_field_class.sign_simps(2) ordered_field_class.sign_simps(8)) ultimately show ?thesis using unitary by auto next case trivial with that show ?thesis by auto qed show \x = y\ if \x \ y\ and \y \ x\ proof (cases rule:cases) case unitary then have \unitary ?id\ by (auto simp: heterogenous_same_type_cblinfun_def) define a b :: \'b \\<^sub>C\<^sub>L 'b\ where \a = (y-x) o\<^sub>C\<^sub>L ?id\ and \b = (x-y) o\<^sub>C\<^sub>L ?id\ with unitary that have \pos a\ and \pos b\ by auto then have \a = 0\ apply (rule_tac pos_nondeg) apply (auto simp: a_def b_def) by (smt (verit, best) add.commute bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose cblinfun_compose_zero_left diff_0 diff_add_cancel group_cancel.rule0 group_cancel.sub1) then show ?thesis unfolding a_def using \unitary ?id\ by (metis cblinfun_compose_assoc cblinfun_compose_id_right cblinfun_compose_zero_left eq_iff_diff_eq_0 unitaryD2) next case trivial with that show ?thesis by simp qed show \x + y \ x + z\ if \y \ z\ proof (cases rule:cases) case unitary with that show ?thesis by auto next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ a *\<^sub>C y\ if \x \ y\ and \0 \ a\ proof (cases rule:cases) case unitary with that pos_scaleC show ?thesis by (metis cblinfun_compose_scaleC_left complex_vector.scale_right_diff_distrib) next case trivial with that show ?thesis by auto qed show \a *\<^sub>C x \ b *\<^sub>C x\ if \a \ b\ and \0 \ x\ proof (cases rule:cases) case unitary with that show ?thesis by (auto intro!: pos_scaleC simp flip: scaleC_diff_left) next case trivial with that show ?thesis by auto qed qed end lemma positive_id_cblinfun[simp]: "id_cblinfun \ 0" unfolding less_eq_cblinfun_def using cinner_ge_zero by auto lemma positive_hermitianI: \A = A*\ if \A \ 0\ apply (rule cinner_real_hermiteanI) using that by (auto simp del: less_eq_complex_def simp: reals_zero_comparable_iff less_eq_cblinfun_def) lemma positive_cblinfunI: \A \ 0\ if \\x. cinner x (A *\<^sub>V x) \ 0\ unfolding less_eq_cblinfun_def using that by auto (* Note: this does not require B to be a square operator *) lemma positive_cblinfun_squareI: \A = B* o\<^sub>C\<^sub>L B \ A \ 0\ apply (rule positive_cblinfunI) by (metis cblinfun_apply_cblinfun_compose cinner_adj_right cinner_ge_zero) lemma one_dim_loewner_order: \A \ B \ one_dim_iso A \ (one_dim_iso B :: complex)\ for A B :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ proof - note less_eq_complex_def[simp del] have A: \A = one_dim_iso A *\<^sub>C id_cblinfun\ by simp have B: \B = one_dim_iso B *\<^sub>C id_cblinfun\ by simp have \A \ B \ (\\. cinner \ (A \) \ cinner \ (B \))\ by (simp add: less_eq_cblinfun_def) also have \\ \ (\\::'a. one_dim_iso B * (\ \\<^sub>C \) \ one_dim_iso A * (\ \\<^sub>C \))\ apply (subst A, subst B) by (metis (no_types, opaque_lifting) cinner_scaleC_right id_cblinfun_apply scaleC_cblinfun.rep_eq) also have \\ \ one_dim_iso A \ (one_dim_iso B :: complex)\ by (auto intro!: mult_right_mono elim!: allE[where x=1]) finally show ?thesis by - qed lemma one_dim_positive: \A \ 0 \ one_dim_iso A \ (0::complex)\ for A :: \'a \\<^sub>C\<^sub>L 'a::{chilbert_space, one_dim}\ using one_dim_loewner_order[where B=0] by auto subsection \Embedding vectors to operators\ lift_definition vector_to_cblinfun :: \'a::complex_normed_vector \ 'b::one_dim \\<^sub>C\<^sub>L 'a\ is \\\ \. one_dim_iso \ *\<^sub>C \\ by (simp add: bounded_clinear_scaleC_const) lemma vector_to_cblinfun_cblinfun_apply: "vector_to_cblinfun (A *\<^sub>V \) = A o\<^sub>C\<^sub>L (vector_to_cblinfun \)" apply transfer unfolding comp_def bounded_clinear_def clinear_def Vector_Spaces.linear_def module_hom_def module_hom_axioms_def by simp lemma vector_to_cblinfun_add: \vector_to_cblinfun (x + y) = vector_to_cblinfun x + vector_to_cblinfun y\ apply transfer by (simp add: scaleC_add_right) lemma norm_vector_to_cblinfun[simp]: "norm (vector_to_cblinfun x) = norm x" proof transfer have "bounded_clinear (one_dim_iso::'a \ complex)" by simp moreover have "onorm (one_dim_iso::'a \ complex) * norm x = norm x" for x :: 'b by simp ultimately show "onorm (\\. one_dim_iso (\::'a) *\<^sub>C x) = norm x" for x :: 'b by (subst onorm_scaleC_left) qed lemma bounded_clinear_vector_to_cblinfun[bounded_clinear]: "bounded_clinear vector_to_cblinfun" apply (rule bounded_clinearI[where K=1]) apply (transfer, simp add: scaleC_add_right) apply (transfer, simp add: mult.commute) by simp lemma vector_to_cblinfun_scaleC[simp]: "vector_to_cblinfun (a *\<^sub>C \) = a *\<^sub>C vector_to_cblinfun \" for a::complex proof (subst asm_rl [of "a *\<^sub>C \ = (a *\<^sub>C id_cblinfun) *\<^sub>V \"]) show "a *\<^sub>C \ = a *\<^sub>C id_cblinfun *\<^sub>V \" by (simp add: scaleC_cblinfun.rep_eq) show "vector_to_cblinfun (a *\<^sub>C id_cblinfun *\<^sub>V \) = a *\<^sub>C (vector_to_cblinfun \::'a \\<^sub>C\<^sub>L 'b)" by (metis cblinfun_id_cblinfun_apply cblinfun_compose_scaleC_left vector_to_cblinfun_cblinfun_apply) qed lemma vector_to_cblinfun_apply_one_dim[simp]: shows "vector_to_cblinfun \ *\<^sub>V \ = one_dim_iso \ *\<^sub>C \" apply transfer by (rule refl) lemma vector_to_cblinfun_adj_apply[simp]: shows "vector_to_cblinfun \* *\<^sub>V \ = of_complex (cinner \ \)" by (simp add: cinner_adj_right one_dim_iso_def one_dim_iso_inj) lemma vector_to_cblinfun_comp_one[simp]: "(vector_to_cblinfun s :: 'a::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L 1 = (vector_to_cblinfun s :: 'b::one_dim \\<^sub>C\<^sub>L _)" apply (transfer fixing: s) by fastforce lemma vector_to_cblinfun_0[simp]: "vector_to_cblinfun 0 = 0" by (metis cblinfun.zero_left cblinfun_compose_zero_left vector_to_cblinfun_cblinfun_apply) lemma image_vector_to_cblinfun[simp]: "vector_to_cblinfun x *\<^sub>S top = ccspan {x}" proof transfer show "closure (range (\\::'b. one_dim_iso \ *\<^sub>C x)) = closure (cspan {x})" for x :: 'a proof (rule arg_cong [where f = closure]) have "k *\<^sub>C x \ range (\\. one_dim_iso \ *\<^sub>C x)" for k by (smt (z3) id_apply one_dim_iso_id one_dim_iso_idem range_eqI) thus "range (\\. one_dim_iso (\::'b) *\<^sub>C x) = cspan {x}" unfolding complex_vector.span_singleton by auto qed qed lemma vector_to_cblinfun_adj_comp_vector_to_cblinfun[simp]: shows "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \ = cinner \ \ *\<^sub>C id_cblinfun" proof - have "one_dim_iso \ *\<^sub>C one_dim_iso (of_complex \\, \\) = \\, \\ *\<^sub>C one_dim_iso \" for \ :: "'c::one_dim" by (metis complex_vector.scale_left_commute of_complex_def one_dim_iso_of_one one_dim_iso_scaleC one_dim_scaleC_1) hence "one_dim_iso ((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = one_dim_iso ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by simp hence "((vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \) *\<^sub>V \) = ((cinner \ \ *\<^sub>C id_cblinfun) *\<^sub>V \)" for \ :: "'c::one_dim" by (rule one_dim_iso_inj) thus ?thesis using cblinfun_eqI[where x = "vector_to_cblinfun \* o\<^sub>C\<^sub>L vector_to_cblinfun \" and y = "\\, \\ *\<^sub>C id_cblinfun"] by auto qed lemma isometry_vector_to_cblinfun[simp]: assumes "norm x = 1" shows "isometry (vector_to_cblinfun x)" using assms cnorm_eq_1 isometry_def by force subsection \Butterflies (rank-1 projectors)\ definition butterfly_def: "butterfly (s::'a::complex_normed_vector) (t::'b::chilbert_space) = vector_to_cblinfun s o\<^sub>C\<^sub>L (vector_to_cblinfun t :: complex \\<^sub>C\<^sub>L _)*" abbreviation "selfbutter s \ butterfly s s" lemma butterfly_add_left: \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (simp add: butterfly_def vector_to_cblinfun_add cbilinear_add_left bounded_cbilinear.add_left bounded_cbilinear_cblinfun_compose) lemma butterfly_add_right: \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (simp add: butterfly_def adj_plus vector_to_cblinfun_add cblinfun_compose_add_right) lemma butterfly_def_one_dim: "butterfly s t = (vector_to_cblinfun s :: 'c::one_dim \\<^sub>C\<^sub>L _) o\<^sub>C\<^sub>L (vector_to_cblinfun t :: 'c \\<^sub>C\<^sub>L _)*" (is "_ = ?rhs") for s :: "'a::complex_normed_vector" and t :: "'b::chilbert_space" proof - let ?isoAC = "1 :: 'c \\<^sub>C\<^sub>L complex" let ?isoCA = "1 :: complex \\<^sub>C\<^sub>L 'c" let ?vector = "vector_to_cblinfun :: _ \ ('c \\<^sub>C\<^sub>L _)" have "butterfly s t = (?vector s o\<^sub>C\<^sub>L ?isoCA) o\<^sub>C\<^sub>L (?vector t o\<^sub>C\<^sub>L ?isoCA)*" unfolding butterfly_def vector_to_cblinfun_comp_one by simp also have "\ = ?vector s o\<^sub>C\<^sub>L (?isoCA o\<^sub>C\<^sub>L ?isoCA*) o\<^sub>C\<^sub>L (?vector t)*" by (metis (no_types, lifting) cblinfun_compose_assoc adj_cblinfun_compose) also have "\ = ?rhs" by simp finally show ?thesis by simp qed lemma butterfly_comp_cblinfun: "butterfly \ \ o\<^sub>C\<^sub>L a = butterfly \ (a* *\<^sub>V \)" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma cblinfun_comp_butterfly: "a o\<^sub>C\<^sub>L butterfly \ \ = butterfly (a *\<^sub>V \) \" unfolding butterfly_def by (simp add: cblinfun_compose_assoc vector_to_cblinfun_cblinfun_apply) lemma butterfly_apply[simp]: "butterfly \ \' *\<^sub>V \ = \\', \\ *\<^sub>C \" by (simp add: butterfly_def scaleC_cblinfun.rep_eq) lemma butterfly_scaleC_left[simp]: "butterfly (c *\<^sub>C \) \ = c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleC_right[simp]: "butterfly \ (c *\<^sub>C \) = cnj c *\<^sub>C butterfly \ \" unfolding butterfly_def vector_to_cblinfun_scaleC scaleC_adj by (simp add: cnj_x_x) lemma butterfly_scaleR_left[simp]: "butterfly (r *\<^sub>R \) \ = r *\<^sub>C butterfly \ \" by (simp add: scaleR_scaleC) lemma butterfly_scaleR_right[simp]: "butterfly \ (r *\<^sub>R \) = r *\<^sub>C butterfly \ \" by (simp add: butterfly_scaleC_right scaleR_scaleC) lemma butterfly_adjoint[simp]: "(butterfly \ \)* = butterfly \ \" unfolding butterfly_def by auto lemma butterfly_comp_butterfly[simp]: "butterfly \1 \2 o\<^sub>C\<^sub>L butterfly \3 \4 = \\2, \3\ *\<^sub>C butterfly \1 \4" by (simp add: butterfly_comp_cblinfun) lemma butterfly_0_left[simp]: "butterfly 0 a = 0" by (simp add: butterfly_def) lemma butterfly_0_right[simp]: "butterfly a 0 = 0" by (simp add: butterfly_def) lemma norm_butterfly: "norm (butterfly \ \) = norm \ * norm \" proof (cases "\=0") case True then show ?thesis by simp next case False show ?thesis unfolding norm_cblinfun.rep_eq thm onormI[OF _ False] proof (rule onormI[OF _ False]) fix x have "cmod \\, x\ * norm \ \ norm \ * norm \ * norm x" by (metis ab_semigroup_mult_class.mult_ac(1) complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult_left_mono norm_ge_zero) thus "norm (butterfly \ \ *\<^sub>V x) \ norm \ * norm \ * norm x" by (simp add: power2_eq_square) show "norm (butterfly \ \ *\<^sub>V \) = norm \ * norm \ * norm \" by (smt (z3) ab_semigroup_mult_class.mult_ac(1) butterfly_apply mult.commute norm_eq_sqrt_cinner norm_ge_zero norm_scaleC power2_eq_square real_sqrt_abs real_sqrt_eq_iff) qed qed lemma bounded_sesquilinear_butterfly[bounded_sesquilinear]: \bounded_sesquilinear (\(b::'b::chilbert_space) (a::'a::chilbert_space). butterfly a b)\ proof standard fix a a' :: 'a and b b' :: 'b and r :: complex show \butterfly (a + a') b = butterfly a b + butterfly a' b\ by (rule butterfly_add_left) show \butterfly a (b + b') = butterfly a b + butterfly a b'\ by (rule butterfly_add_right) show \butterfly (r *\<^sub>C a) b = r *\<^sub>C butterfly a b\ by simp show \butterfly a (r *\<^sub>C b) = cnj r *\<^sub>C butterfly a b\ by simp show \\K. \b a. norm (butterfly a b) \ norm b * norm a * K \ apply (rule exI[of _ 1]) by (simp add: norm_butterfly) qed lemma inj_selfbutter_upto_phase: assumes "selfbutter x = selfbutter y" shows "\c. cmod c = 1 \ x = c *\<^sub>C y" proof (cases "x = 0") case True from assms have "y = 0" using norm_butterfly by (metis True butterfly_0_left divisors_zero norm_eq_zero) with True show ?thesis using norm_one by fastforce next case False define c where "c = \y, x\ / \x, x\" have "\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x" by (simp add: butterfly_apply) also have "\ = selfbutter y *\<^sub>V x" using assms by simp also have "\ = \y, x\ *\<^sub>C y" by (simp add: butterfly_apply) finally have xcy: "x = c *\<^sub>C y" by (simp add: c_def ceq_vector_fraction_iff) have "cmod c * norm x = cmod c * norm y" using assms norm_butterfly by (smt (verit, ccfv_SIG) \\x, x\ *\<^sub>C x = selfbutter x *\<^sub>V x\ \selfbutter y *\<^sub>V x = \y, x\ *\<^sub>C y\ cinner_scaleC_right complex_vector.scale_left_commute complex_vector.scale_right_imp_eq mult_cancel_left norm_eq_sqrt_cinner norm_eq_zero scaleC_scaleC xcy) also have "cmod c * norm y = norm (c *\<^sub>C y)" by simp also have "\ = norm x" unfolding xcy[symmetric] by simp finally have c: "cmod c = 1" by (simp add: False) from c xcy show ?thesis by auto qed lemma butterfly_eq_proj: assumes "norm x = 1" shows "selfbutter x = proj x" proof - define B and \ :: "complex \\<^sub>C\<^sub>L 'a" where "B = selfbutter x" and "\ = vector_to_cblinfun x" then have B: "B = \ o\<^sub>C\<^sub>L \*" unfolding butterfly_def by simp have \adj\: "\* o\<^sub>C\<^sub>L \ = id_cblinfun" using \_def assms isometry_def isometry_vector_to_cblinfun by blast have "B o\<^sub>C\<^sub>L B = \ o\<^sub>C\<^sub>L (\* o\<^sub>C\<^sub>L \) o\<^sub>C\<^sub>L \*" by (simp add: B cblinfun_assoc_left(1)) also have "\ = B" unfolding \adj\ by (simp add: B) finally have idem: "B o\<^sub>C\<^sub>L B = B". have herm: "B = B*" unfolding B by simp from idem herm have BProj: "B = Proj (B *\<^sub>S top)" by (rule Proj_on_own_range'[symmetric]) have "B *\<^sub>S top = ccspan {x}" by (simp add: B \_def assms cblinfun_compose_image range_adjoint_isometry) with BProj show "B = proj x" by simp qed lemma butterfly_is_Proj: \norm x = 1 \ is_Proj (selfbutter x)\ by (subst butterfly_eq_proj, simp_all) lemma cspan_butterfly_UNIV: assumes \cspan basisA = UNIV\ assumes \cspan basisB = UNIV\ assumes \is_ortho_set basisB\ assumes \\b. b \ basisB \ norm b = 1\ shows \cspan {butterfly a b| (a::'a::{complex_normed_vector}) (b::'b::{chilbert_space,cfinite_dim}). a \ basisA \ b \ basisB} = UNIV\ proof - have F: \\F\{butterfly a b |a b. a \ basisA \ b \ basisB}. \b'\basisB. F *\<^sub>V b' = (if b' = b then a else 0)\ if \a \ basisA\ and \b \ basisB\ for a b apply (rule bexI[where x=\butterfly a b\]) using assms that by (auto simp: is_ortho_set_def cnorm_eq_1) show ?thesis apply (rule cblinfun_cspan_UNIV[where basisA=basisB and basisB=basisA]) using assms apply auto[2] using F by (smt (verit, ccfv_SIG) image_iff) qed lemma cindependent_butterfly: fixes basisA :: \'a::chilbert_space set\ and basisB :: \'b::chilbert_space set\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes normA: \\a. a\basisA \ norm a = 1\ and normB: \\b. b\basisB \ norm b = 1\ shows \cindependent {butterfly a b| a b. a\basisA \ b\basisB}\ proof (unfold complex_vector.independent_explicit_module, intro allI impI, rename_tac T f g) fix T :: \('b \\<^sub>C\<^sub>L 'a) set\ and f :: \'b \\<^sub>C\<^sub>L 'a \ complex\ and g :: \'b \\<^sub>C\<^sub>L 'a\ assume \finite T\ assume T_subset: \T \ {butterfly a b |a b. a \ basisA \ b \ basisB}\ define lin where \lin = (\g\T. f g *\<^sub>C g)\ assume \lin = 0\ assume \g \ T\ (* To show: f g = 0 *) then obtain a b where g: \g = butterfly a b\ and [simp]: \a \ basisA\ \b \ basisB\ using T_subset by auto have *: "(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0" if \g \ T - {butterfly a b}\ for g proof - from that obtain a' b' where g: \g = butterfly a' b'\ and [simp]: \a' \ basisA\ \b' \ basisB\ using T_subset by auto from that have \g \ butterfly a b\ by auto with g consider (a) \a\a'\ | (b) \b\b'\ by auto then show \(vector_to_cblinfun a)* *\<^sub>V f g *\<^sub>C g *\<^sub>V b = 0\ proof cases case a then show ?thesis using \is_ortho_set basisA\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) next case b then show ?thesis using \is_ortho_set basisB\ unfolding g by (auto simp: is_ortho_set_def butterfly_def scaleC_cblinfun.rep_eq) qed qed have \0 = (vector_to_cblinfun a)* *\<^sub>V lin *\<^sub>V b\ using \lin = 0\ by auto also have \\ = (\g\T. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ unfolding lin_def apply (rule complex_vector.linear_sum) by (smt (z3) cblinfun.scaleC_left cblinfun.scaleC_right cblinfun.add_right clinearI plus_cblinfun.rep_eq) also have \\ = (\g\{butterfly a b}. (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b)\ apply (rule sum.mono_neutral_right) using \finite T\ * \g \ T\ g by auto also have \\ = (vector_to_cblinfun a)* *\<^sub>V (f g *\<^sub>C g) *\<^sub>V b\ by (simp add: g) also have \\ = f g\ unfolding g using normA normB by (auto simp: butterfly_def scaleC_cblinfun.rep_eq cnorm_eq_1) finally show \f g = 0\ by simp qed lemma clinear_eq_butterflyI: fixes F G :: \('a::{chilbert_space,cfinite_dim} \\<^sub>C\<^sub>L 'b::complex_inner) \ 'c::complex_vector\ assumes "clinear F" and "clinear G" assumes \cspan basisA = UNIV\ \cspan basisB = UNIV\ assumes \is_ortho_set basisA\ \is_ortho_set basisB\ assumes "\a b. a\basisA \ b\basisB \ F (butterfly a b) = G (butterfly a b)" assumes \\b. b\basisB \ norm b = 1\ shows "F = G" apply (rule complex_vector.linear_eq_on_span[where f=F, THEN ext, rotated 3]) apply (subst cspan_butterfly_UNIV) using assms by auto subsection \Bifunctionals\ lift_definition bifunctional :: \'a::complex_normed_vector \\<^sub>C\<^sub>L (('a \\<^sub>C\<^sub>L complex) \\<^sub>C\<^sub>L complex)\ is \\x f. f *\<^sub>V x\ by (simp add: cblinfun.flip) lemma bifunctional_apply[simp]: \(bifunctional *\<^sub>V x) *\<^sub>V f = f *\<^sub>V x\ by (transfer fixing: x f, simp) lemma bifunctional_isometric[simp]: \norm (bifunctional *\<^sub>V x) = norm x\ for x :: \'a::complex_inner\ proof - define f :: \'a \\<^sub>C\<^sub>L complex\ where \f = CBlinfun (\y. cinner x y)\ then have [simp]: \f *\<^sub>V y = cinner x y\ for y by (simp add: bounded_clinear_CBlinfun_apply bounded_clinear_cinner_right) then have [simp]: \norm f = norm x\ apply (auto intro!: norm_cblinfun_eqI[where x=x] simp: power2_norm_eq_cinner[symmetric]) apply (smt (verit, best) norm_eq_sqrt_cinner norm_ge_zero power2_norm_eq_cinner real_div_sqrt) using Cauchy_Schwarz_ineq2 by blast show ?thesis apply (auto intro!: norm_cblinfun_eqI[where x=f]) apply (metis norm_eq_sqrt_cinner norm_imp_pos_and_ge real_div_sqrt) by (metis norm_cblinfun ordered_field_class.sign_simps(33)) qed lemma norm_bifunctional[simp]: \norm (bifunctional :: 'a::{complex_inner, not_singleton} \\<^sub>C\<^sub>L _) = 1\ proof - obtain x :: 'a where [simp]: \norm x = 1\ by (meson UNIV_not_singleton ex_norm1) show ?thesis by (auto intro!: norm_cblinfun_eqI[where x=x]) qed subsection \Banach-Steinhaus\ theorem cbanach_steinhaus: fixes F :: \'c \ 'a::cbanach \\<^sub>C\<^sub>L 'b::complex_normed_vector\ assumes \\x. \M. \n. norm ((F n) *\<^sub>V x) \ M\ shows \\M. \ n. norm (F n) \ M\ using cblinfun_blinfun_transfer[transfer_rule] apply (rule TrueI)? (* Deletes current facts *) proof (use assms in transfer) fix F :: \'c \ 'a \\<^sub>L 'b\ assume \(\x. \M. \n. norm (F n *\<^sub>v x) \ M)\ hence \\x. bounded (range (\n. blinfun_apply (F n) x))\ by (metis (no_types, lifting) boundedI rangeE) hence \bounded (range F)\ by (simp add: banach_steinhaus) thus \\M. \n. norm (F n) \ M\ by (simp add: bounded_iff) qed subsection \Riesz-representation theorem\ theorem riesz_frechet_representation_cblinfun_existence: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ shows \\t. \x. f *\<^sub>V x = \t, x\\ apply transfer by (rule riesz_frechet_representation_existence) lemma riesz_frechet_representation_cblinfun_unique: \ \Theorem 3.4 in @{cite conway2013course}\ fixes f::\'a::complex_inner \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ assumes \\x. f *\<^sub>V x = \u, x\\ shows \t = u\ using assms by (rule riesz_frechet_representation_unique) theorem riesz_frechet_representation_cblinfun_norm: includes notation_norm fixes f::\'a::chilbert_space \\<^sub>C\<^sub>L complex\ assumes \\x. f *\<^sub>V x = \t, x\\ shows \\f\ = \t\\ using assms proof transfer fix f::\'a \ complex\ and t assume \bounded_clinear f\ and \\x. f x = \t, x\\ from \\x. f x = \t, x\\ have \(norm (f x)) / (norm x) \ norm t\ for x proof(cases \norm x = 0\) case True thus ?thesis by simp next case False have \norm (f x) = norm (\t, x\)\ using \\x. f x = \t, x\\ by simp also have \norm \t, x\ \ norm t * norm x\ by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) finally have \norm (f x) \ norm t * norm x\ by blast thus ?thesis by (metis False linordered_field_class.divide_right_mono nonzero_mult_div_cancel_right norm_ge_zero) qed moreover have \(norm (f t)) / (norm t) = norm t\ proof(cases \norm t = 0\) case True thus ?thesis by simp next case False have \f t = \t, t\\ using \\x. f x = \t, x\\ by blast also have \\ = (norm t)^2\ by (meson cnorm_eq_square) also have \\ = (norm t)*(norm t)\ by (simp add: power2_eq_square) finally have \f t = (norm t)*(norm t)\ by blast thus ?thesis by (metis False Re_complex_of_real \\x. f x = cinner t x\ cinner_ge_zero complex_of_real_cmod nonzero_divide_eq_eq) qed ultimately have \Sup {(norm (f x)) / (norm x)| x. True} = norm t\ by (smt cSup_eq_maximum mem_Collect_eq) moreover have \Sup {(norm (f x)) / (norm x)| x. True} = (SUP x. (norm (f x)) / (norm x))\ by (simp add: full_SetCompr_eq) ultimately show \onorm f = norm t\ by (simp add: onorm_def) qed subsection \Extension of complex bounded operators\ definition cblinfun_extension where "cblinfun_extension S \ = (SOME B. \x\S. B *\<^sub>V x = \ x)" definition cblinfun_extension_exists where "cblinfun_extension_exists S \ = (\B. \x\S. B *\<^sub>V x = \ x)" lemma cblinfun_extension_existsI: assumes "\x. x\S \ B *\<^sub>V x = \ x" shows "cblinfun_extension_exists S \" using assms cblinfun_extension_exists_def by blast lemma cblinfun_extension_exists_finite_dim: fixes \::"'a::{complex_normed_vector,cfinite_dim} \ 'b::complex_normed_vector" assumes "cindependent S" and "cspan S = UNIV" shows "cblinfun_extension_exists S \" proof- define f::"'a \ 'b" where "f = complex_vector.construct S \" have "clinear f" by (simp add: complex_vector.linear_construct assms linear_construct f_def) have "bounded_clinear f" using \clinear f\ assms by auto then obtain B::"'a \\<^sub>C\<^sub>L 'b" where "B *\<^sub>V x = f x" for x using cblinfun_apply_cases by blast have "B *\<^sub>V x = \ x" if c1: "x\S" for x proof- have "B *\<^sub>V x = f x" by (simp add: \\x. B *\<^sub>V x = f x\) also have "\ = \ x" using assms complex_vector.construct_basis f_def that by (simp add: complex_vector.construct_basis) finally show?thesis by blast qed thus ?thesis unfolding cblinfun_extension_exists_def by blast qed lemma cblinfun_extension_exists_bounded_dense: fixes f :: \'a::complex_normed_vector \ 'b::cbanach\ assumes \csubspace S\ assumes \closure S = UNIV\ assumes f_add: \\x y. x \ S \ y \ S \ f (x + y) = f x + f y\ assumes f_scale: \\c x y. x \ S \ f (c *\<^sub>C x) = c *\<^sub>C f x\ assumes bounded: \\x. x \ S \ norm (f x) \ B * norm x\ shows \cblinfun_extension_exists S f\ proof - obtain B where bounded: \\x. x \ S \ norm (f x) \ B * norm x\ and \B > 0\ using bounded by (smt (z3) mult_mono norm_ge_zero) have \\xi. (xi \ x) \ (\i. xi i \ S)\ for x using assms(2) closure_sequential by blast then obtain seq :: \'a \ nat \ 'a\ where seq_lim: \seq x \ x\ and seq_S: \seq x i \ S\ for x i apply (atomize_elim, subst all_conj_distrib[symmetric]) apply (rule choice) by auto define g where \g x = lim (\i. f (seq x i))\ for x have \Cauchy (\i. f (seq x i))\ for x proof (rule CauchyI) fix e :: real assume \e > 0\ have \Cauchy (seq x)\ using LIMSEQ_imp_Cauchy seq_lim by blast then obtain M where less_eB: \norm (seq x m - seq x n) < e/B\ if \n \ M\ and \m \ M\ for n m apply atomize_elim by (meson CauchyD \0 < B\ \0 < e\ linordered_field_class.divide_pos_pos) have \norm (f (seq x m) - f (seq x n)) < e\ if \n \ M\ and \m \ M\ for n m proof - have \norm (f (seq x m) - f (seq x n)) = norm (f (seq x m - seq x n))\ using f_add f_scale seq_S by (metis add_diff_cancel assms(1) complex_vector.subspace_diff diff_add_cancel) also have \\ \ B * norm (seq x m - seq x n)\ apply (rule bounded) by (simp add: assms(1) complex_vector.subspace_diff seq_S) also from less_eB have \\ < B * (e/B)\ by (meson \0 < B\ linordered_semiring_strict_class.mult_strict_left_mono that) also have \\ \ e\ using \0 < B\ by auto finally show ?thesis by - qed then show \\M. \m\M. \n\M. norm (f (seq x m) - f (seq x n)) < e\ by auto qed then have f_seq_lim: \(\i. f (seq x i)) \ g x\ for x by (simp add: Cauchy_convergent_iff convergent_LIMSEQ_iff g_def) have f_xi_lim: \(\i. f (xi i)) \ g x\ if \xi \ x\ and \\i. xi i \ S\ for xi x proof - from seq_lim that have \(\i. B * norm (xi i - seq x i)) \ 0\ by (metis (no_types) \0 < B\ cancel_comm_monoid_add_class.diff_cancel norm_not_less_zero norm_zero tendsto_diff tendsto_norm_zero_iff tendsto_zero_mult_left_iff) then have \(\i. f (xi i + (-1) *\<^sub>C seq x i)) \ 0\ apply (rule Lim_null_comparison[rotated]) using bounded by (simp add: assms(1) complex_vector.subspace_diff seq_S that(2)) then have \(\i. f (xi i) - f (seq x i)) \ 0\ apply (subst (asm) f_add) apply (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S)[2] apply (subst (asm) f_scale) by (auto simp: that \csubspace S\ complex_vector.subspace_neg seq_S) then show \(\i. f (xi i)) \ g x\ using Lim_transform f_seq_lim by fastforce qed have g_add: \g (x + y) = g x + g y\ for x y proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto obtain yi :: \nat \ 'a\ where \yi \ y\ and \yi i \ S\ for i using seq_S seq_lim by auto have \(\i. xi i + yi i) \ x + y\ using \xi \ x\ \yi \ y\ tendsto_add by blast then have lim1: \(\i. f (xi i + yi i)) \ g (x + y)\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ assms(1) complex_vector.subspace_add f_xi_lim) have \(\i. f (xi i + yi i)) = (\i. f (xi i) + f (yi i))\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ f_add) also have \\ \ g x + g y\ by (simp add: \\i. xi i \ S\ \\i. yi i \ S\ \xi \ x\ \yi \ y\ f_xi_lim tendsto_add) finally show ?thesis using lim1 LIMSEQ_unique by blast qed have g_scale: \g (c *\<^sub>C x) = c *\<^sub>C g x\ for c x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto have \(\i. c *\<^sub>C xi i) \ c *\<^sub>C x\ using \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at isCont_tendsto_compose by blast then have lim1: \(\i. f (c *\<^sub>C xi i)) \ g (c *\<^sub>C x)\ by (simp add: \\i. xi i \ S\ assms(1) complex_vector.subspace_scale f_xi_lim) have \(\i. f (c *\<^sub>C xi i)) = (\i. c *\<^sub>C f (xi i))\ by (simp add: \\i. xi i \ S\ f_scale) also have \\ \ c *\<^sub>C g x\ using \\i. xi i \ S\ \xi \ x\ bounded_clinear_scaleC_right clinear_continuous_at f_xi_lim isCont_tendsto_compose by blast finally show ?thesis using lim1 LIMSEQ_unique by blast qed have [simp]: \f x = g x\ if \x \ S\ for x proof - have \(\_. x) \ x\ by auto then have \(\_. f x) \ g x\ using that by (rule f_xi_lim) then show \f x = g x\ by (simp add: LIMSEQ_const_iff) qed have g_bounded: \norm (g x) \ B * norm x\ for x proof - obtain xi :: \nat \ 'a\ where \xi \ x\ and \xi i \ S\ for i using seq_S seq_lim by auto then have \(\i. f (xi i)) \ g x\ using f_xi_lim by presburger then have \(\i. norm (f (xi i))) \ norm (g x)\ by (metis tendsto_norm) moreover have \(\i. B * norm (xi i)) \ B * norm x\ by (simp add: \xi \ x\ tendsto_mult_left tendsto_norm) ultimately show \norm (g x) \ B * norm x\ apply (rule lim_mono[rotated]) using bounded using \xi _ \ S\ by blast qed have \bounded_clinear g\ using g_add g_scale apply (rule bounded_clinearI[where K=B]) using g_bounded by (simp add: ordered_field_class.sign_simps(5)) then have [simp]: \CBlinfun g *\<^sub>V x = g x\ for x by (subst CBlinfun_inverse, auto) show \cblinfun_extension_exists S f\ apply (rule cblinfun_extension_existsI[where B=\CBlinfun g\]) by auto qed lemma cblinfun_extension_apply: assumes "cblinfun_extension_exists S f" and "v \ S" shows "(cblinfun_extension S f) *\<^sub>V v = f v" by (smt assms cblinfun_extension_def cblinfun_extension_exists_def tfl_some) subsection \Notation\ bundle cblinfun_notation begin notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) notation cblinfun_apply (infixr "*\<^sub>V" 70) notation cblinfun_image (infixr "*\<^sub>S" 70) notation adj ("_*" [99] 100) end bundle no_cblinfun_notation begin no_notation cblinfun_compose (infixl "o\<^sub>C\<^sub>L" 55) no_notation cblinfun_apply (infixr "*\<^sub>V" 70) no_notation cblinfun_image (infixr "*\<^sub>S" 70) no_notation adj ("_*" [99] 100) end bundle blinfun_notation begin notation blinfun_apply (infixr "*\<^sub>V" 70) end bundle no_blinfun_notation begin no_notation blinfun_apply (infixr "*\<^sub>V" 70) end unbundle no_cblinfun_notation end diff --git a/thys/Complex_Bounded_Operators/Complex_Euclidean_Space0.thy b/thys/Complex_Bounded_Operators/Complex_Euclidean_Space0.thy --- a/thys/Complex_Bounded_Operators/Complex_Euclidean_Space0.thy +++ b/thys/Complex_Bounded_Operators/Complex_Euclidean_Space0.thy @@ -1,393 +1,393 @@ (* Title: HOL/Analysis/Euclidean_Space.thy Author: Johannes Hölzl, TU München Author: Brian Huffman, Portland State University *) section \\Complex_Euclidean_Space0\ -- Finite-Dimensional Inner Product Spaces\ theory Complex_Euclidean_Space0 imports "HOL-Analysis.L2_Norm" "Complex_Inner_Product" "HOL-Analysis.Product_Vector" "HOL-Library.Rewrite" begin (* subsection\<^marker>\tag unimportant\ \Interlude: Some properties of real sets\ *) (* Complex analogue not needed *) (* lemma seq_mono_lemma: assumes "\(n::nat) \ m. (d n :: real) < e n" and "\n \ m. e n \ e m" shows "\n \ m. d n < e m" *) subsection \Type class of Euclidean spaces\ class ceuclidean_space = complex_inner + fixes CBasis :: "'a set" assumes nonempty_CBasis [simp]: "CBasis \ {}" assumes finite_CBasis [simp]: "finite CBasis" assumes cinner_CBasis: "\u \ CBasis; v \ CBasis\ \ cinner u v = (if u = v then 1 else 0)" assumes ceuclidean_all_zero_iff: "(\u\CBasis. cinner x u = 0) \ (x = 0)" syntax "_type_cdimension" :: "type \ nat" ("(1CDIM/(1'(_')))") translations "CDIM('a)" \ "CONST card (CONST CBasis :: 'a set)" typed_print_translation \ [(\<^const_syntax>\card\, fn ctxt => fn _ => fn [Const (\<^const_syntax>\CBasis\, Type (\<^type_name>\set\, [T]))] => Syntax.const \<^syntax_const>\_type_cdimension\ $ Syntax_Phases.term_of_typ ctxt T)] \ lemma (in ceuclidean_space) norm_CBasis[simp]: "u \ CBasis \ norm u = 1" unfolding norm_eq_sqrt_cinner by (simp add: cinner_CBasis) lemma (in ceuclidean_space) cinner_same_CBasis[simp]: "u \ CBasis \ cinner u u = 1" by (simp add: cinner_CBasis) lemma (in ceuclidean_space) cinner_not_same_CBasis: "u \ CBasis \ v \ CBasis \ u \ v \ cinner u v = 0" by (simp add: cinner_CBasis) lemma (in ceuclidean_space) sgn_CBasis: "u \ CBasis \ sgn u = u" unfolding sgn_div_norm by (simp add: scaleR_one) lemma (in ceuclidean_space) CBasis_zero [simp]: "0 \ CBasis" proof assume "0 \ CBasis" thus "False" using cinner_CBasis [of 0 0] by simp qed lemma (in ceuclidean_space) nonzero_CBasis: "u \ CBasis \ u \ 0" by clarsimp lemma (in ceuclidean_space) SOME_CBasis: "(SOME i. i \ CBasis) \ CBasis" by (metis ex_in_conv nonempty_CBasis someI_ex) lemma norm_some_CBasis [simp]: "norm (SOME i. i \ CBasis) = 1" by (simp add: SOME_CBasis) lemma (in ceuclidean_space) cinner_sum_left_CBasis[simp]: "b \ CBasis \ cinner (\i\CBasis. f i *\<^sub>C i) b = cnj (f b)" by (simp add: cinner_sum_left cinner_CBasis if_distrib comm_monoid_add_class.sum.If_cases) (* Not present in Euclidean_Space *) (* lemma (in ceuclidean_space) cinner_sum_right_CBasis[simp]: "b \ CBasis \ cinner b (\i\CBasis. f i *\<^sub>C i) = f b" by (metis (mono_tags, lifting) cinner_commute cinner_sum_left_CBasis comm_monoid_add_class.sum.cong complex_cnj_cnj) *) lemma (in ceuclidean_space) ceuclidean_eqI: assumes b: "\b. b \ CBasis \ cinner x b = cinner y b" shows "x = y" proof - from b have "\b\CBasis. cinner (x - y) b = 0" by (simp add: cinner_diff_left) then show "x = y" by (simp add: ceuclidean_all_zero_iff) qed lemma (in ceuclidean_space) ceuclidean_eq_iff: "x = y \ (\b\CBasis. cinner x b = cinner y b)" by (auto intro: ceuclidean_eqI) lemma (in ceuclidean_space) ceuclidean_representation_sum: "(\i\CBasis. f i *\<^sub>C i) = b \ (\i\CBasis. f i = cnj (cinner b i))" apply (subst ceuclidean_eq_iff) apply simp by (metis complex_cnj_cnj cinner_commute) lemma (in ceuclidean_space) ceuclidean_representation_sum': "b = (\i\CBasis. f i *\<^sub>C i) \ (\i\CBasis. f i = cinner i b)" apply (auto simp add: ceuclidean_representation_sum[symmetric]) apply (metis ceuclidean_representation_sum cinner_commute) by (metis local.ceuclidean_representation_sum local.cinner_commute) lemma (in ceuclidean_space) ceuclidean_representation: "(\b\CBasis. cinner b x *\<^sub>C b) = x" unfolding ceuclidean_representation_sum using local.cinner_commute by blast lemma (in ceuclidean_space) ceuclidean_cinner: "cinner x y = (\b\CBasis. cinner x b * cnj (cinner y b))" apply (subst (1 2) ceuclidean_representation [symmetric]) apply (simp add: cinner_sum_right cinner_CBasis ac_simps) by (metis local.cinner_commute mult.commute) lemma (in ceuclidean_space) choice_CBasis_iff: fixes P :: "'a \ complex \ bool" shows "(\i\CBasis. \x. P i x) \ (\x. \i\CBasis. P i (cinner x i))" unfolding bchoice_iff proof safe fix f assume "\i\CBasis. P i (f i)" then show "\x. \i\CBasis. P i (cinner x i)" by (auto intro!: exI[of _ "\i\CBasis. cnj (f i) *\<^sub>C i"]) qed auto lemma (in ceuclidean_space) bchoice_CBasis_iff: fixes P :: "'a \ complex \ bool" shows "(\i\CBasis. \x\A. P i x) \ (\x. \i\CBasis. cinner x i \ A \ P i (cinner x i))" by (simp add: choice_CBasis_iff Bex_def) lemma (in ceuclidean_space) ceuclidean_representation_sum_fun: "(\x. \b\CBasis. cinner b (f x) *\<^sub>C b) = f" apply (rule ext) apply (simp add: ceuclidean_representation_sum) by (meson local.cinner_commute) lemma euclidean_isCont: assumes "\b. b \ CBasis \ isCont (\x. (cinner b (f x)) *\<^sub>C b) x" shows "isCont f x" apply (subst ceuclidean_representation_sum_fun [symmetric]) apply (rule isCont_sum) by (blast intro: assms) lemma CDIM_positive [simp]: "0 < CDIM('a::ceuclidean_space)" by (simp add: card_gt_0_iff) lemma CDIM_ge_Suc0 [simp]: "Suc 0 \ card CBasis" by (meson CDIM_positive Suc_leI) lemma sum_cinner_CBasis_scaleC [simp]: fixes f :: "'a::ceuclidean_space \ 'b::complex_vector" assumes "b \ CBasis" shows "(\i\CBasis. (cinner i b) *\<^sub>C f i) = f b" by (simp add: comm_monoid_add_class.sum.remove [OF finite_CBasis assms] assms cinner_not_same_CBasis comm_monoid_add_class.sum.neutral) lemma sum_cinner_CBasis_eq [simp]: assumes "b \ CBasis" shows "(\i\CBasis. (cinner i b) * f i) = f b" by (simp add: comm_monoid_add_class.sum.remove [OF finite_CBasis assms] assms cinner_not_same_CBasis comm_monoid_add_class.sum.neutral) lemma sum_if_cinner [simp]: assumes "i \ CBasis" "j \ CBasis" shows "cinner (\k\CBasis. if k = i then f i *\<^sub>C i else g k *\<^sub>C k) j = (if j=i then cnj (f j) else cnj (g j))" proof (cases "i=j") case True with assms show ?thesis by (auto simp: cinner_sum_left if_distrib [of "\x. cinner x j"] cinner_CBasis cong: if_cong) next case False have "(\k\CBasis. cinner (if k = i then f i *\<^sub>C i else g k *\<^sub>C k) j) = (\k\CBasis. if k = j then cnj (g k) else 0)" apply (rule sum.cong) using False assms by (auto simp: cinner_CBasis) also have "... = cnj (g j)" using assms by auto finally show ?thesis using False by (auto simp: cinner_sum_left) qed lemma norm_le_componentwise: "(\b. b \ CBasis \ cmod(cinner x b) \ cmod(cinner y b)) \ norm x \ norm y" apply (auto simp: cnorm_le ceuclidean_cinner [of x x] ceuclidean_cinner [of y y] power2_eq_square intro!: sum_mono) apply (smt (verit, best) mult.commute sum.cong) - by (simp add: ordered_field_class.sign_simps(33)) + done lemma CBasis_le_norm: "b \ CBasis \ cmod (cinner x b) \ norm x" by (rule order_trans [OF Cauchy_Schwarz_ineq2]) simp lemma norm_bound_CBasis_le: "b \ CBasis \ norm x \ e \ cmod (inner x b) \ e" by (metis inner_commute mult.left_neutral norm_CBasis norm_of_real order_trans real_inner_class.Cauchy_Schwarz_ineq2) lemma norm_bound_CBasis_lt: "b \ CBasis \ norm x < e \ cmod (inner x b) < e" by (metis inner_commute le_less_trans mult.left_neutral norm_CBasis norm_of_real real_inner_class.Cauchy_Schwarz_ineq2) lemma cnorm_le_l1: "norm x \ (\b\CBasis. cmod (cinner x b))" apply (subst ceuclidean_representation[of x, symmetric]) apply (rule order_trans[OF norm_sum]) apply (auto intro!: sum_mono) by (metis cinner_commute complex_inner_1_left complex_inner_class.Cauchy_Schwarz_ineq2 mult.commute mult.left_neutral norm_one) (* Maybe it holds in the complex case but the proof does not adapt trivially *) (* lemma csum_norm_allsubsets_bound: fixes f :: "'a \ 'n::ceuclidean_space" assumes fP: "finite P" and fPs: "\Q. Q \ P \ norm (sum f Q) \ e" shows "(\x\P. norm (f x)) \ 2 * real CDIM('n) * e" *) (* subsection\<^marker>\tag unimportant\ \Subclass relationships\ *) (* Everything is commented out, so we comment out the heading, too. *) (* If we include this, instantiation prod :: (ceuclidean_space, ceuclidean_space) ceuclidean_space below fails *) (* instance ceuclidean_space \ perfect_space proof fix x :: 'a show "\ open {x}" proof assume "open {x}" then obtain e where "0 < e" and e: "\y. dist y x < e \ y = x" unfolding open_dist by fast define y where "y = x + scaleR (e/2) (SOME b. b \ CBasis)" have [simp]: "(SOME b. b \ CBasis) \ CBasis" by (rule someI_ex) (auto simp: ex_in_conv) from \0 < e\ have "y \ x" unfolding y_def by (auto intro!: nonzero_CBasis) from \0 < e\ have "dist y x < e" unfolding y_def by (simp add: dist_norm) from \y \ x\ and \dist y x < e\ show "False" using e by simp qed qed *) subsection \Class instances\ subsubsection\<^marker>\tag unimportant\ \Type \<^typ>\complex\\ (* No analogue *) (* instantiation real :: ceuclidean_space *) instantiation complex :: ceuclidean_space begin definition [simp]: "CBasis = {1::complex}" instance by standard auto end lemma CDIM_complex[simp]: "CDIM(complex) = 1" by simp (* lemma CDIM_complex[simp]: "DIM(complex) = 2" lemma complex_CBasis_1 [iff]: "(1::complex) \ CBasis" lemma complex_CBasis_i [iff]: "\ \ CBasis" *) subsubsection\<^marker>\tag unimportant\ \Type \<^typ>\'a \ 'b\\ 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)" lemma cinner_Pair [simp]: "cinner (a, b) (c, d) = cinner a c + cinner b d" unfolding cinner_prod_def by simp instance proof fix r :: complex fix x y z :: "'a::complex_inner \ 'b::complex_inner" show "cinner x y = cnj (cinner y x)" unfolding cinner_prod_def by simp show "cinner (x + y) z = cinner x z + cinner y z" unfolding cinner_prod_def by (simp add: cinner_add_left) show "cinner (scaleC r x) y = cnj r * cinner x y" unfolding cinner_prod_def by (simp add: distrib_left) show "0 \ cinner x x" unfolding cinner_prod_def by (intro add_nonneg_nonneg cinner_ge_zero) show "cinner x x = 0 \ x = 0" unfolding cinner_prod_def prod_eq_iff by (metis antisym cinner_eq_zero_iff cinner_ge_zero fst_zero le_add_same_cancel2 snd_zero verit_sum_simplify) show "norm x = sqrt (cmod (cinner x x))" unfolding norm_prod_def cinner_prod_def by (metis (no_types, lifting) Re_complex_of_real add_nonneg_nonneg cinner_ge_zero complex_of_real_cmod plus_complex.simps(1) power2_norm_eq_cinner') qed end lemma cinner_Pair_0: "cinner x (0, b) = cinner (snd x) b" "cinner x (a, 0) = cinner (fst x) a" by (cases x, simp)+ instantiation prod :: (ceuclidean_space, ceuclidean_space) ceuclidean_space begin definition "CBasis = (\u. (u, 0)) ` CBasis \ (\v. (0, v)) ` CBasis" lemma sum_CBasis_prod_eq: fixes f::"('a*'b)\('a*'b)" shows "sum f CBasis = sum (\i. f (i, 0)) CBasis + sum (\i. f (0, i)) CBasis" proof - have "inj_on (\u. (u::'a, 0::'b)) CBasis" "inj_on (\u. (0::'a, u::'b)) CBasis" by (auto intro!: inj_onI Pair_inject) thus ?thesis unfolding CBasis_prod_def by (subst sum.union_disjoint) (auto simp: CBasis_prod_def sum.reindex) qed instance proof show "(CBasis :: ('a \ 'b) set) \ {}" unfolding CBasis_prod_def by simp next show "finite (CBasis :: ('a \ 'b) set)" unfolding CBasis_prod_def by simp next fix u v :: "'a \ 'b" assume "u \ CBasis" and "v \ CBasis" thus "cinner u v = (if u = v then 1 else 0)" unfolding CBasis_prod_def cinner_prod_def by (auto simp add: cinner_CBasis split: if_split_asm) next fix x :: "'a \ 'b" show "(\u\CBasis. cinner x u = 0) \ x = 0" unfolding CBasis_prod_def ball_Un ball_simps by (simp add: cinner_prod_def prod_eq_iff ceuclidean_all_zero_iff) qed lemma CDIM_prod[simp]: "CDIM('a \ 'b) = CDIM('a) + CDIM('b)" unfolding CBasis_prod_def by (subst card_Un_disjoint) (auto intro!: card_image arg_cong2[where f="(+)"] inj_onI) end subsection \Locale instances\ lemma finite_dimensional_vector_space_euclidean: "finite_dimensional_vector_space (*\<^sub>C) CBasis" proof unfold_locales show "finite (CBasis::'a set)" by (metis finite_CBasis) show "complex_vector.independent (CBasis::'a set)" unfolding complex_vector.dependent_def cdependent_raw_def[symmetric] apply (subst complex_vector.span_finite) apply simp apply clarify apply (drule_tac f="cinner a" in arg_cong) by (simp add: cinner_CBasis cinner_sum_right eq_commute) show "module.span (*\<^sub>C) CBasis = UNIV" unfolding complex_vector.span_finite [OF finite_CBasis] cspan_raw_def[symmetric] by (auto intro!: ceuclidean_representation[symmetric]) qed interpretation ceucl: finite_dimensional_vector_space "scaleC :: complex => 'a => 'a::ceuclidean_space" "CBasis" rewrites "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" and "Vector_Spaces.linear (*) (*\<^sub>C) = clinear" and "finite_dimensional_vector_space.dimension CBasis = CDIM('a)" (* and "dimension = CDIM('a)" *) (* This line leads to a type error. Not sure why *) by (auto simp add: cdependent_raw_def crepresentation_raw_def csubspace_raw_def cspan_raw_def cextend_basis_raw_def cdim_raw_def clinear_def complex_scaleC_def[abs_def] finite_dimensional_vector_space.dimension_def intro!: finite_dimensional_vector_space.dimension_def finite_dimensional_vector_space_euclidean) interpretation ceucl: finite_dimensional_vector_space_pair_1 "scaleC::complex\'a::ceuclidean_space\'a" CBasis "scaleC::complex\'b::complex_vector \ 'b" by unfold_locales interpretation ceucl?: finite_dimensional_vector_space_prod scaleC scaleC CBasis CBasis rewrites "Basis_pair = CBasis" and "module_prod.scale (*\<^sub>C) (*\<^sub>C) = (scaleC::_\_\('a \ 'b))" proof - show "finite_dimensional_vector_space_prod (*\<^sub>C) (*\<^sub>C) CBasis CBasis" by unfold_locales interpret finite_dimensional_vector_space_prod "(*\<^sub>C)" "(*\<^sub>C)" "CBasis::'a set" "CBasis::'b set" by fact show "Basis_pair = CBasis" unfolding Basis_pair_def CBasis_prod_def by auto show "module_prod.scale (*\<^sub>C) (*\<^sub>C) = scaleC" by (fact module_prod_scale_eq_scaleC) qed end diff --git a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy --- a/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy +++ b/thys/Complex_Bounded_Operators/Complex_Inner_Product.thy @@ -1,2230 +1,2230 @@ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) section \\Complex_Inner_Product\ -- Complex Inner Product Spaces\ theory Complex_Inner_Product imports Complex_Vector_Spaces "HOL-Analysis.Infinite_Set_Sum" Complex_Inner_Product0 begin subsection \Complex inner product spaces\ bundle cinner_bracket_notation begin notation cinner ("\_, _\") end unbundle cinner_bracket_notation bundle no_cinner_bracket_notation begin no_notation cinner ("\_, _\") end lemma cinner_real: "cinner x x \ \" by (meson cinner_ge_zero reals_zero_comparable_iff) lemmas cinner_commute' [simp] = cinner_commute[symmetric] lemma (in complex_inner) cinner_eq_flip: \(cinner x y = cinner z w) \ (cinner y x = cinner w z)\ by (metis cinner_commute) lemma Im_cinner_x_x[simp]: "Im \x , x\ = 0" using comp_Im_same[OF cinner_ge_zero] by simp lemma of_complex_inner_1' [simp]: "cinner (1 :: 'a :: {complex_inner, complex_normed_algebra_1}) (of_complex x) = x" by (metis cinner_commute complex_cnj_cnj of_complex_inner_1) class chilbert_space = complex_inner + complete_space begin subclass cbanach by standard end instantiation complex :: "chilbert_space" begin instance .. end subsection \Misc facts\ text \This is a useful rule for establishing the equality of vectors\ lemma cinner_extensionality: assumes \\\. \\, \\ = \\, \\\ shows \\ = \\ by (metis assms cinner_eq_zero_iff cinner_simps(3) right_minus_eq) lemma polar_identity: includes notation_norm shows \\x + y\^2 = \x\^2 + \y\^2 + 2*Re \x, y\\ \ \Shown in the proof of Corollary 1.5 in @{cite conway2013course}\ proof - have \\x , y\ + \y , x\ = \x , y\ + cnj \x , y\\ by simp hence \\x , y\ + \y , x\ = 2 * Re \x , y\ \ using complex_add_cnj by presburger have \\x + y\^2 = \x+y, x+y\\ by (simp add: cdot_square_norm) 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 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 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 text \Orthogonal set\ definition is_ortho_set :: "'a::complex_inner set \ bool" where \is_ortho_set S = ((\x\S. \y\S. x \ y \ \x, y\ = 0) \ 0 \ S)\ lemma is_ortho_set_empty[simp]: "is_ortho_set {}" unfolding is_ortho_set_def by auto lemma is_ortho_set_antimono: \A \ B \ is_ortho_set B \ is_ortho_set A\ unfolding is_ortho_set_def by auto lemma orthogonal_complement_of_closure: fixes A ::"('a::complex_inner) set" shows "orthogonal_complement A = orthogonal_complement (closure A)" proof- have s1: \\ y, x \ = 0\ 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 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 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) + 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 instance conjugate_space :: (chilbert_space) chilbert_space.. end diff --git a/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy b/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy --- a/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy +++ b/thys/Complex_Bounded_Operators/Complex_Inner_Product0.thy @@ -1,548 +1,548 @@ (* Based on HOL/Real_Vector_Spaces.thy by Brian Huffman Adapted to the complex case by Dominique Unruh *) section \\Complex_Inner_Product0\ -- Inner Product Spaces and Gradient Derivative\ theory Complex_Inner_Product0 imports Complex_Main Complex_Vector_Spaces "HOL-Analysis.Inner_Product" "Complex_Bounded_Operators.Extra_Ordered_Fields" begin subsection \Complex inner product spaces\ text \ Temporarily relax type constraints for \<^term>\open\, \<^term>\uniformity\, \<^term>\dist\, and \<^term>\norm\. \ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::open set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::dist \ 'a \ real\)\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniformity \ 'a) filter\)\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::norm \ real\)\ class complex_inner = complex_vector + sgn_div_norm + dist_norm + uniformity_dist + open_uniformity + fixes cinner :: "'a \ 'a \ complex" assumes cinner_commute: "cinner x y = cnj (cinner y x)" and cinner_add_left: "cinner (x + y) z = cinner x z + cinner y z" and cinner_scaleC_left [simp]: "cinner (scaleC r x) y = (cnj r) * (cinner x y)" and cinner_ge_zero [simp]: "0 \ cinner x x" and cinner_eq_zero_iff [simp]: "cinner x x = 0 \ x = 0" and norm_eq_sqrt_cinner: "norm x = sqrt (cmod (cinner x x))" begin lemma cinner_zero_left [simp]: "cinner 0 x = 0" using cinner_add_left [of 0 0 x] by simp lemma cinner_minus_left [simp]: "cinner (- x) y = - cinner x y" using cinner_add_left [of x "- x" y] by (simp add: group_add_class.add_eq_0_iff) lemma cinner_diff_left: "cinner (x - y) z = cinner x z - cinner y z" using cinner_add_left [of x "- y" z] by simp lemma cinner_sum_left: "cinner (\x\A. f x) y = (\x\A. cinner (f x) y)" by (cases "finite A", induct set: finite, simp_all add: cinner_add_left) lemma call_zero_iff [simp]: "(\u. cinner x u = 0) \ (x = 0)" by auto (use cinner_eq_zero_iff in blast) text \Transfer distributivity rules to right argument.\ lemma cinner_add_right: "cinner x (y + z) = cinner x y + cinner x z" using cinner_add_left [of y z x] by (metis complex_cnj_add local.cinner_commute) lemma cinner_scaleC_right [simp]: "cinner x (scaleC r y) = r * (cinner x y)" using cinner_scaleC_left [of r y x] by (metis complex_cnj_cnj complex_cnj_mult local.cinner_commute) lemma cinner_zero_right [simp]: "cinner x 0 = 0" using cinner_zero_left [of x] by (metis (mono_tags, opaque_lifting) complex_cnj_zero local.cinner_commute) lemma cinner_minus_right [simp]: "cinner x (- y) = - cinner x y" using cinner_minus_left [of y x] by (metis complex_cnj_minus local.cinner_commute) lemma cinner_diff_right: "cinner x (y - z) = cinner x y - cinner x z" using cinner_diff_left [of y z x] by (metis complex_cnj_diff local.cinner_commute) lemma cinner_sum_right: "cinner x (\y\A. f y) = (\y\A. cinner x (f y))" proof (subst cinner_commute) have "(\y\A. cinner (f y) x) = (\y\A. cinner (f y) x)" by blast hence "cnj (\y\A. cinner (f y) x) = cnj (\y\A. (cinner (f y) x))" by simp hence "cnj (cinner (sum f A) x) = (\y\A. cnj (cinner (f y) x))" by (simp add: cinner_sum_left) thus "cnj (cinner (sum f A) x) = (\y\A. (cinner x (f y)))" by (subst (2) cinner_commute) qed lemmas cinner_add [algebra_simps] = cinner_add_left cinner_add_right lemmas cinner_diff [algebra_simps] = cinner_diff_left cinner_diff_right lemmas cinner_scaleC = cinner_scaleC_left cinner_scaleC_right (* text \Legacy theorem names\ lemmas cinner_left_distrib = cinner_add_left lemmas cinner_right_distrib = cinner_add_right lemmas cinner_distrib = cinner_left_distrib cinner_right_distrib *) lemma cinner_gt_zero_iff [simp]: "0 < cinner x x \ x \ 0" by (smt (verit) less_irrefl local.cinner_eq_zero_iff local.cinner_ge_zero order.not_eq_order_implies_strict) (* In Inner_Product, we have lemma power2_norm_eq_cinner: "(norm x)\<^sup>2 = cinner x x" The following are two ways of inserting the conversions between real and complex into this: *) lemma power2_norm_eq_cinner: shows "(complex_of_real (norm x))\<^sup>2 = (cinner x x)" by (smt (verit, del_insts) Im_complex_of_real Re_complex_of_real cinner_gt_zero_iff cinner_zero_right cmod_def complex_eq_0 complex_eq_iff less_complex_def local.norm_eq_sqrt_cinner of_real_power real_sqrt_abs real_sqrt_pow2_iff zero_complex.sel(1)) lemma power2_norm_eq_cinner': shows "(norm x)\<^sup>2 = Re (cinner x x)" by (metis Re_complex_of_real of_real_power power2_norm_eq_cinner) text \Identities involving real multiplication and division.\ lemma cinner_mult_left: "cinner (of_complex m * a) b = cnj m * (cinner a b)" by (simp add: of_complex_def) lemma cinner_mult_right: "cinner a (of_complex m * b) = m * (cinner a b)" by (metis complex_inner_class.cinner_scaleC_right scaleC_conv_of_complex) lemma cinner_mult_left': "cinner (a * of_complex m) b = cnj m * (cinner a b)" by (metis cinner_mult_left mult.right_neutral mult_scaleC_right scaleC_conv_of_complex) lemma cinner_mult_right': "cinner a (b * of_complex m) = (cinner a b) * m" by (simp add: complex_inner_class.cinner_scaleC_right of_complex_def) (* In Inner_Product, we have lemma Cauchy_Schwarz_ineq: "(cinner x y)\<^sup>2 \ cinner x x * cinner y y" The following are two ways of inserting the conversions between real and complex into this: *) lemma Cauchy_Schwarz_ineq: "(cinner x y) * (cinner y x) \ cinner x x * cinner y y" proof (cases) assume "y = 0" thus ?thesis by simp next assume y: "y \ 0" have [simp]: "cnj (cinner y y) = cinner y y" for y by (metis cinner_commute) define r where "r = cnj (cinner x y) / cinner y y" have "0 \ cinner (x - scaleC r y) (x - scaleC r y)" by (rule cinner_ge_zero) also have "\ = cinner x x - r * cinner x y - cnj r * cinner y x + r * cnj r * cinner y y" unfolding cinner_diff_left cinner_diff_right cinner_scaleC_left cinner_scaleC_right by (smt (z3) cancel_comm_monoid_add_class.diff_cancel cancel_comm_monoid_add_class.diff_zero complex_cnj_divide group_add_class.diff_add_cancel local.cinner_commute local.cinner_eq_zero_iff local.cinner_scaleC_left mult.assoc mult.commute mult_eq_0_iff nonzero_eq_divide_eq r_def y) also have "\ = cinner x x - cinner y x * cnj r" unfolding r_def by auto also have "\ = cinner x x - cinner x y * cnj (cinner x y) / cinner y y" unfolding r_def by (metis complex_cnj_divide local.cinner_commute mult.commute times_divide_eq_left) finally have "0 \ cinner x x - cinner x y * cnj (cinner x y) / cinner y y" . hence "cinner x y * cnj (cinner x y) / cinner y y \ cinner x x" by (simp add: le_diff_eq) thus "cinner x y * cinner y x \ cinner x x * cinner y y" by (metis cinner_gt_zero_iff local.cinner_commute nice_ordered_field_class.pos_divide_le_eq y) qed lemma Cauchy_Schwarz_ineq2: shows "norm (cinner x y) \ norm x * norm y" proof (rule power2_le_imp_le) have "(norm (cinner x y))^2 = Re (cinner x y * cinner y x)" by (metis (full_types) Re_complex_of_real complex_norm_square local.cinner_commute) also have "\ \ Re (cinner x x * cinner y y)" using Cauchy_Schwarz_ineq by (rule Re_mono) also have "\ = Re (complex_of_real ((norm x)^2) * complex_of_real ((norm y)^2))" by (simp add: power2_norm_eq_cinner) also have "\ = (norm x * norm y)\<^sup>2" by (simp add: power_mult_distrib) finally show "(cmod (cinner x y))^2 \ (norm x * norm y)\<^sup>2" . show "0 \ norm x * norm y" by (simp add: local.norm_eq_sqrt_cinner) qed (* The following variant does not hold in the complex case: *) (* lemma norm_cauchy_schwarz: "cinner x y \ norm x * norm y" using Cauchy_Schwarz_ineq2 [of x y] by auto *) subclass complex_normed_vector proof fix a :: complex and r :: real and x y :: 'a show "norm x = 0 \ x = 0" unfolding norm_eq_sqrt_cinner by simp show "norm (x + y) \ norm x + norm y" proof (rule power2_le_imp_le) have "Re (cinner x y) \ cmod (cinner x y)" if "\x. Re x \ cmod x" and "\x y. x \ y \ complex_of_real x \ complex_of_real y" using that by simp hence a1: "2 * Re (cinner x y) \ 2 * cmod (cinner x y)" if "\x. Re x \ cmod x" and "\x y. x \ y \ complex_of_real x \ complex_of_real y" using that by simp have "cinner x y + cinner y x = complex_of_real (2 * Re (cinner x y))" by (metis complex_add_cnj local.cinner_commute) also have "\ \ complex_of_real (2 * cmod (cinner x y))" using complex_Re_le_cmod complex_of_real_mono a1 by blast also have "\ = 2 * abs (cinner x y)" unfolding abs_complex_def by simp also have "\ \ 2 * complex_of_real (norm x) * complex_of_real (norm y)" - using Cauchy_Schwarz_ineq2 unfolding abs_complex_def by auto + using Cauchy_Schwarz_ineq2 unfolding abs_complex_def less_eq_complex_def by auto finally have xyyx: "cinner x y + cinner y x \ complex_of_real (2 * norm x * norm y)" by auto have "complex_of_real ((norm (x + y))\<^sup>2) = cinner (x+y) (x+y)" by (simp add: power2_norm_eq_cinner) also have "\ = cinner x x + cinner x y + cinner y x + cinner y y" by (simp add: cinner_add) also have "\ = complex_of_real ((norm x)\<^sup>2) + complex_of_real ((norm y)\<^sup>2) + cinner x y + cinner y x" by (simp add: power2_norm_eq_cinner) also have "\ \ complex_of_real ((norm x)\<^sup>2) + complex_of_real ((norm y)\<^sup>2) + complex_of_real (2 * norm x * norm y)" using xyyx by auto also have "\ = complex_of_real ((norm x + norm y)\<^sup>2)" unfolding power2_sum by auto finally show "(norm (x + y))\<^sup>2 \ (norm x + norm y)\<^sup>2" using complex_of_real_mono_iff by blast show "0 \ norm x + norm y" unfolding norm_eq_sqrt_cinner by simp qed show norm_scaleC: "norm (a *\<^sub>C x) = cmod a * norm x" for a proof (rule power2_eq_imp_eq) show "(norm (a *\<^sub>C x))\<^sup>2 = (cmod a * norm x)\<^sup>2" by (simp_all add: norm_eq_sqrt_cinner norm_mult power2_eq_square) show "0 \ norm (a *\<^sub>C x)" by (simp_all add: norm_eq_sqrt_cinner) show "0 \ cmod a * norm x" by (simp_all add: norm_eq_sqrt_cinner) qed show "norm (r *\<^sub>R x) = \r\ * norm x" unfolding scaleR_scaleC norm_scaleC by auto qed end (* Does not hold in the complex case *) (* lemma csquare_bound_lemma: fixes x :: complex shows "x < (1 + x) * (1 + x)" *) lemma csquare_continuous: fixes e :: real shows "e > 0 \ \d. 0 < d \ (\y. cmod (y - x) < d \ cmod (y * y - x * x) < e)" using isCont_power[OF continuous_ident, of x, unfolded isCont_def LIM_eq, rule_format, of e 2] by (force simp add: power2_eq_square) lemma cnorm_le: "norm x \ norm y \ cinner x x \ cinner y y" by (smt (verit) complex_of_real_mono_iff norm_eq_sqrt_cinner norm_ge_zero of_real_power power2_norm_eq_cinner real_sqrt_le_mono real_sqrt_pow2) lemma cnorm_lt: "norm x < norm y \ cinner x x < cinner y y" by (meson cnorm_le less_le_not_le) lemma cnorm_eq: "norm x = norm y \ cinner x x = cinner y y" by (metis norm_eq_sqrt_cinner power2_norm_eq_cinner) lemma cnorm_eq_1: "norm x = 1 \ cinner x x = 1" by (metis cinner_ge_zero complex_of_real_cmod norm_eq_sqrt_cinner norm_one of_real_1 real_sqrt_eq_iff real_sqrt_one) lemma cinner_divide_left: fixes a :: "'a :: {complex_inner,complex_div_algebra}" shows "cinner (a / of_complex m) b = (cinner a b) / cnj m" by (metis cinner_mult_left' complex_cnj_inverse divide_inverse of_complex_inverse ordered_field_class.sign_simps(33)) lemma cinner_divide_right: fixes a :: "'a :: {complex_inner,complex_div_algebra}" shows "cinner a (b / of_complex m) = (cinner a b) / m" by (metis cinner_mult_right' divide_inverse of_complex_inverse) text \ Re-enable constraints for \<^term>\open\, \<^term>\uniformity\, \<^term>\dist\, and \<^term>\norm\. \ setup \Sign.add_const_constraint (\<^const_name>\open\, SOME \<^typ>\'a::topological_space set \ bool\)\ setup \Sign.add_const_constraint (\<^const_name>\uniformity\, SOME \<^typ>\('a::uniform_space \ 'a) filter\)\ setup \Sign.add_const_constraint (\<^const_name>\dist\, SOME \<^typ>\'a::metric_space \ 'a \ real\)\ setup \Sign.add_const_constraint (\<^const_name>\norm\, SOME \<^typ>\'a::real_normed_vector \ real\)\ lemma bounded_sesquilinear_cinner: "bounded_sesquilinear (cinner::'a::complex_inner \ 'a \ complex)" proof fix x y z :: 'a and r :: complex show "cinner (x + y) z = cinner x z + cinner y z" by (rule cinner_add_left) show "cinner x (y + z) = cinner x y + cinner x z" by (rule cinner_add_right) show "cinner (scaleC r x) y = scaleC (cnj r) (cinner x y)" unfolding complex_scaleC_def by (rule cinner_scaleC_left) show "cinner x (scaleC r y) = scaleC r (cinner x y)" unfolding complex_scaleC_def by (rule cinner_scaleC_right) have "\x y::'a. norm (cinner x y) \ norm x * norm y * 1" by (simp add: complex_inner_class.Cauchy_Schwarz_ineq2) thus "\K. \x y::'a. norm (cinner x y) \ norm x * norm y * K" by metis qed lemmas tendsto_cinner [tendsto_intros] = bounded_bilinear.tendsto [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas isCont_cinner [simp] = bounded_bilinear.isCont [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas has_derivative_cinner [derivative_intros] = bounded_bilinear.FDERIV [OF bounded_sesquilinear_cinner[THEN bounded_sesquilinear.bounded_bilinear]] lemmas bounded_antilinear_cinner_left = bounded_sesquilinear.bounded_antilinear_left [OF bounded_sesquilinear_cinner] lemmas bounded_clinear_cinner_right = bounded_sesquilinear.bounded_clinear_right [OF bounded_sesquilinear_cinner] lemmas bounded_antilinear_cinner_left_comp = bounded_antilinear_cinner_left[THEN bounded_antilinear_o_bounded_clinear] lemmas bounded_clinear_cinner_right_comp = bounded_clinear_cinner_right[THEN bounded_clinear_compose] lemmas has_derivative_cinner_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_clinear_cinner_right[THEN bounded_clinear.bounded_linear]] lemmas has_derivative_cinner_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_antilinear_cinner_left[THEN bounded_antilinear.bounded_linear]] lemma differentiable_cinner [simp]: "f differentiable (at x within s) \ g differentiable at x within s \ (\x. cinner (f x) (g x)) differentiable at x within s" unfolding differentiable_def by (blast intro: has_derivative_cinner) subsection \Class instances\ instantiation complex :: complex_inner begin definition cinner_complex_def [simp]: "cinner x y = cnj x * y" instance proof fix x y z r :: complex show "cinner x y = cnj (cinner y x)" unfolding cinner_complex_def by auto show "cinner (x + y) z = cinner x z + cinner y z" unfolding cinner_complex_def by (simp add: ring_class.ring_distribs(2)) show "cinner (scaleC r x) y = cnj r * cinner x y" unfolding cinner_complex_def complex_scaleC_def by simp show "0 \ cinner x x" by simp show "cinner x x = 0 \ x = 0" unfolding cinner_complex_def by simp have "cmod (Complex x1 x2) = sqrt (cmod (cinner (Complex x1 x2) (Complex x1 x2)))" for x1 x2 unfolding cinner_complex_def complex_cnj complex_mult complex_norm by (simp add: power2_eq_square) thus "norm x = sqrt (cmod (cinner x x))" by (cases x, hypsubst_thin) qed end lemma shows complex_inner_1_left[simp]: "cinner 1 x = x" and complex_inner_1_right[simp]: "cinner x 1 = cnj x" by simp_all (* No analogous to \instantiation complex :: real_inner\ or to lemma complex_inner_1 [simp]: "inner 1 x = Re x" lemma complex_inner_1_right [simp]: "inner x 1 = Re x" lemma complex_inner_i_left [simp]: "inner \ x = Im x" lemma complex_inner_i_right [simp]: "inner x \ = Im x" *) lemma cdot_square_norm: "cinner x x = complex_of_real ((norm x)\<^sup>2)" by (metis Im_complex_of_real Re_complex_of_real cinner_ge_zero complex_eq_iff less_eq_complex_def power2_norm_eq_cinner' zero_complex.simps(2)) lemma cnorm_eq_square: "norm x = a \ 0 \ a \ cinner x x = complex_of_real (a\<^sup>2)" by (metis cdot_square_norm norm_ge_zero of_real_eq_iff power2_eq_iff_nonneg) lemma cnorm_le_square: "norm x \ a \ 0 \ a \ cinner x x \ complex_of_real (a\<^sup>2)" by (smt (verit) cdot_square_norm complex_of_real_mono_iff norm_ge_zero power2_le_imp_le) lemma cnorm_ge_square: "norm x \ a \ a \ 0 \ cinner x x \ complex_of_real (a\<^sup>2)" by (smt (verit, best) antisym_conv cnorm_eq_square cnorm_le_square complex_of_real_nn_iff nn_comparable zero_le_power2) lemma norm_lt_square: "norm x < a \ 0 < a \ cinner x x < complex_of_real (a\<^sup>2)" by (meson cnorm_ge_square cnorm_le_square less_le_not_le) lemma norm_gt_square: "norm x > a \ a < 0 \ cinner x x > complex_of_real (a\<^sup>2)" by (smt (verit, ccfv_SIG) cdot_square_norm complex_of_real_strict_mono_iff norm_ge_zero power2_eq_imp_eq power_mono) text\Dot product in terms of the norm rather than conversely.\ lemmas cinner_simps = cinner_add_left cinner_add_right cinner_diff_right cinner_diff_left cinner_scaleC_left cinner_scaleC_right (* Analogue to both dot_norm and dot_norm_neg *) lemma cdot_norm: "cinner x y = ((norm (x+y))\<^sup>2 - (norm (x-y))\<^sup>2 - \ * (norm (x + \ *\<^sub>C y))\<^sup>2 + \ * (norm (x - \ *\<^sub>C y))\<^sup>2) / 4" unfolding power2_norm_eq_cinner by (simp add: power2_norm_eq_cinner cinner_add_left cinner_add_right cinner_diff_left cinner_diff_right ring_distribs) lemma of_complex_inner_1 [simp]: "cinner (of_complex x) (1 :: 'a :: {complex_inner, complex_normed_algebra_1}) = cnj x" by (metis Complex_Inner_Product0.complex_inner_1_right cinner_complex_def cinner_mult_left complex_cnj_one norm_one of_complex_def power2_norm_eq_cinner scaleC_conv_of_complex) lemma summable_of_complex_iff: "summable (\x. of_complex (f x) :: 'a :: {complex_normed_algebra_1,complex_inner}) \ summable f" proof assume *: "summable (\x. of_complex (f x) :: 'a)" have "bounded_clinear (cinner (1::'a))" by (rule bounded_clinear_cinner_right) then interpret bounded_linear "\x::'a. cinner 1 x" by (rule bounded_clinear.bounded_linear) from summable [OF *] show "summable f" apply (subst (asm) cinner_commute) by simp next assume sum: "summable f" thus "summable (\x. of_complex (f x) :: 'a)" by (rule summable_of_complex) qed subsection \Gradient derivative\ definition\<^marker>\tag important\ cgderiv :: "['a::complex_inner \ complex, 'a, 'a] \ bool" ("(cGDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where (* Must be "cinner D" not "\h. cinner h D", otherwise not even "cGDERIV id x :> 1" holds *) "cGDERIV f x :> D \ FDERIV f x :> cinner D" lemma cgderiv_deriv [simp]: "cGDERIV f x :> D \ DERIV f x :> cnj D" by (simp only: cgderiv_def has_field_derivative_def cinner_complex_def[THEN ext]) lemma cGDERIV_DERIV_compose: assumes "cGDERIV f x :> df" and "DERIV g (f x) :> cnj dg" shows "cGDERIV (\x. g (f x)) x :> scaleC dg df" proof (insert assms) show "cGDERIV (\x. g (f x)) x :> dg *\<^sub>C df" if "cGDERIV f x :> df" and "(g has_field_derivative cnj dg) (at (f x))" unfolding cgderiv_def has_field_derivative_def cinner_scaleC_left complex_cnj_cnj using that by (simp add: cgderiv_def has_derivative_compose has_field_derivative_imp_has_derivative) qed (* Not specific to complex/real *) (* lemma has_derivative_subst: "\FDERIV f x :> df; df = d\ \ FDERIV f x :> d" *) lemma cGDERIV_subst: "\cGDERIV f x :> df; df = d\ \ cGDERIV f x :> d" by simp lemma cGDERIV_const: "cGDERIV (\x. k) x :> 0" unfolding cgderiv_def cinner_zero_left[THEN ext] by (rule has_derivative_const) lemma cGDERIV_add: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x + g x) x :> df + dg" unfolding cgderiv_def cinner_add_left[THEN ext] by (rule has_derivative_add) lemma cGDERIV_minus: "cGDERIV f x :> df \ cGDERIV (\x. - f x) x :> - df" unfolding cgderiv_def cinner_minus_left[THEN ext] by (rule has_derivative_minus) lemma cGDERIV_diff: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x - g x) x :> df - dg" unfolding cgderiv_def cinner_diff_left by (rule has_derivative_diff) lemma cGDERIV_scaleC: "\DERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. scaleC (f x) (g x)) x :> (scaleC (cnj (f x)) dg + scaleC (cnj df) (cnj (g x)))" unfolding cgderiv_def has_field_derivative_def cinner_add_left cinner_scaleC_left apply (rule has_derivative_subst) apply (erule (1) has_derivative_scaleC) by (simp add: ac_simps) lemma GDERIV_mult: "\cGDERIV f x :> df; cGDERIV g x :> dg\ \ cGDERIV (\x. f x * g x) x :> cnj (f x) *\<^sub>C dg + cnj (g x) *\<^sub>C df" unfolding cgderiv_def apply (rule has_derivative_subst) apply (erule (1) has_derivative_mult) apply (rule ext) by (simp add: cinner_add ac_simps) lemma cGDERIV_inverse: "\cGDERIV f x :> df; f x \ 0\ \ cGDERIV (\x. inverse (f x)) x :> - cnj ((inverse (f x))\<^sup>2) *\<^sub>C df" by (metis DERIV_inverse cGDERIV_DERIV_compose complex_cnj_cnj complex_cnj_minus numerals(2)) (* Don't know if this holds: *) (* lemma cGDERIV_norm: assumes "x \ 0" shows "cGDERIV (\x. norm x) x :> sgn x" *) lemma has_derivative_norm[derivative_intros]: fixes x :: "'a::complex_inner" assumes "x \ 0" shows "(norm has_derivative (\h. Re (cinner (sgn x) h))) (at x)" thm has_derivative_norm proof - have Re_pos: "0 < Re (cinner x x)" using assms by (metis Re_strict_mono cinner_gt_zero_iff zero_complex.simps(1)) have Re_plus_Re: "Re (cinner x y) + Re (cinner y x) = 2 * Re (cinner x y)" for x y :: 'a by (metis cinner_commute cnj.simps(1) mult_2_right semiring_normalization_rules(7)) have norm: "norm x = sqrt (Re (cinner x x))" for x :: 'a apply (subst norm_eq_sqrt_cinner, subst cmod_Re) using cinner_ge_zero by auto have v2:"((\x. sqrt (Re (cinner x x))) has_derivative (\xa. (Re (cinner x xa) + Re (cinner xa x)) * (inverse (sqrt (Re (cinner x x))) / 2))) (at x)" by (rule derivative_eq_intros | simp add: Re_pos)+ have v1: "((\x. sqrt (Re (cinner x x))) has_derivative (\y. Re (cinner x y) / sqrt (Re (cinner x x)))) (at x)" if "((\x. sqrt (Re (cinner x x))) has_derivative (\xa. Re (cinner x xa) * inverse (sqrt (Re (cinner x x))))) (at x)" using that apply (subst divide_real_def) by simp have \(norm has_derivative (\y. Re (cinner x y) / norm x)) (at x)\ using v2 apply (auto simp: Re_plus_Re norm [abs_def]) using v1 by blast then show ?thesis by (auto simp: power2_eq_square sgn_div_norm scaleR_scaleC) qed bundle cinner_syntax begin notation cinner (infix "\\<^sub>C" 70) end bundle no_cinner_syntax begin no_notation cinner (infix "\\<^sub>C" 70) end end diff --git a/thys/Complex_Bounded_Operators/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,518 +1,518 @@ section \\Extra_General\ -- General missing things\ theory Extra_General imports "HOL-Library.Cardinality" "HOL-Analysis.Elementary_Topology" Jordan_Normal_Form.Conjugate "HOL-Analysis.Uniform_Limit" "HOL-Library.Set_Algebras" "HOL-Types_To_Sets.Types_To_Sets" begin subsection \Misc\ lemma reals_zero_comparable_iff: "(x::complex)\\ \ x \ 0 \ x \ 0" unfolding complex_is_Real_iff less_eq_complex_def by auto lemma reals_zero_comparable: fixes x::complex assumes "x\\" shows "x \ 0 \ x \ 0" using assms unfolding reals_zero_comparable_iff by assumption lemma unique_choice: "\x. \!y. Q x y \ \!f. \x. Q x (f x)" apply (auto intro!: choice ext) by metis 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 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 bdd_above_image_mono: assumes \\x. x\S \ f x \ g x\ assumes \bdd_above (g ` S)\ shows \bdd_above (f ` S)\ by (smt (verit, ccfv_threshold) assms(1) assms(2) bdd_aboveI2 bdd_above_def order_trans rev_image_eqI) 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" 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 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 (rule filter_leI) show "eventually P (filtermap (\F. F \ A) (finite_subsets_at_top B))" if "eventually P (finite_subsets_at_top A)" for P :: "'a set \ bool" using that unfolding eventually_filtermap unfolding eventually_finite_subsets_at_top by (metis Int_subset_iff assms finite_Int inf_le2 subset_trans) 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 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" proof (cases x) show "0 \ cnj x * x" if "x = Complex x1 x2" for x1 :: real and x2 :: real - using that by (auto simp: complex_cnj complex_mult complex_of_real_def) + using that by (auto simp: complex_cnj complex_mult complex_of_real_def less_eq_complex_def) qed 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 end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy b/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Jordan_Normal_Form.thy @@ -1,432 +1,432 @@ section \\Extra_Jordan_Normal_Form\ -- Additional results for \<^session>\Jordan_Normal_Form\\ (* Authors: Dominique Unruh, University of Tartu, unruh@ut.ee Jose Manuel Rodriguez Caballero, University of Tartu, jose.manuel.rodriguez.caballero@ut.ee *) theory Extra_Jordan_Normal_Form imports Jordan_Normal_Form.Matrix Jordan_Normal_Form.Schur_Decomposition begin text \We define bundles to activate/deactivate the notation from \<^session>\Jordan_Normal_Form\. Reactivate the notation locally via "@{theory_text \includes jnf_notation\}" in a lemma statement. (Or sandwich a declaration using that notation between "@{theory_text \unbundle jnf_notation ... unbundle no_jnf_notation\}.) \ bundle jnf_notation begin notation transpose_mat ("(_\<^sup>T)" [1000]) notation cscalar_prod (infix "\c" 70) notation vec_index (infixl "$" 100) notation smult_vec (infixl "\\<^sub>v" 70) notation scalar_prod (infix "\" 70) notation index_mat (infixl "$$" 100) notation smult_mat (infixl "\\<^sub>m" 70) notation mult_mat_vec (infixl "*\<^sub>v" 70) notation pow_mat (infixr "^\<^sub>m" 75) notation append_vec (infixr "@\<^sub>v" 65) notation append_rows (infixr "@\<^sub>r" 65) end bundle no_jnf_notation begin no_notation transpose_mat ("(_\<^sup>T)" [1000]) no_notation cscalar_prod (infix "\c" 70) no_notation vec_index (infixl "$" 100) no_notation smult_vec (infixl "\\<^sub>v" 70) no_notation scalar_prod (infix "\" 70) no_notation index_mat (infixl "$$" 100) no_notation smult_mat (infixl "\\<^sub>m" 70) no_notation mult_mat_vec (infixl "*\<^sub>v" 70) no_notation pow_mat (infixr "^\<^sub>m" 75) no_notation append_vec (infixr "@\<^sub>v" 65) no_notation append_rows (infixr "@\<^sub>r" 65) end unbundle jnf_notation lemma mat_entry_explicit: fixes M :: "'a::field mat" assumes "M \ carrier_mat m n" and "i < m" and "j < n" shows "vec_index (M *\<^sub>v unit_vec n j) i = M $$ (i,j)" using assms by auto lemma mat_adjoint_def': "mat_adjoint M = transpose_mat (map_mat conjugate M)" apply (rule mat_eq_iff[THEN iffD2]) apply (auto simp: mat_adjoint_def transpose_mat_def) apply (subst mat_of_rows_index) by auto lemma mat_adjoint_swap: fixes M ::"complex mat" assumes "M \ carrier_mat nB nA" and "iA < dim_row M" and "iB < dim_col M" shows "(mat_adjoint M)$$(iB,iA) = cnj (M$$(iA,iB))" unfolding transpose_mat_def map_mat_def by (simp add: assms(2) assms(3) mat_adjoint_def') lemma cscalar_prod_adjoint: fixes M:: "complex mat" assumes "M \ carrier_mat nB nA" and "dim_vec v = nA" and "dim_vec u = nB" shows "v \c ((mat_adjoint M) *\<^sub>v u) = (M *\<^sub>v v) \c u" unfolding mat_adjoint_def using assms(1) assms(2,3)[symmetric] apply (simp add: scalar_prod_def sum_distrib_left field_simps) by (intro sum.swap) lemma scaleC_minus1_left_vec: "-1 \\<^sub>v v = - v" for v :: "_::ring_1 vec" unfolding smult_vec_def uminus_vec_def by auto lemma square_nneg_complex: fixes x :: complex assumes "x \ \" shows "x^2 \ 0" - apply (cases x) using assms unfolding Reals_def by auto + apply (cases x) using assms unfolding Reals_def less_eq_complex_def by auto definition "vec_is_zero n v = (\i vec_is_zero n v \ v = 0\<^sub>v n" unfolding vec_is_zero_def apply auto by (metis index_zero_vec(1)) fun gram_schmidt_sub0 where "gram_schmidt_sub0 n us [] = us" | "gram_schmidt_sub0 n us (w # ws) = (let w' = adjuster n w us + w in if vec_is_zero n w' then gram_schmidt_sub0 n us ws else gram_schmidt_sub0 n (w' # us) ws)" lemma (in cof_vec_space) adjuster_already_in_span: assumes "w \ carrier_vec n" assumes us_carrier: "set us \ carrier_vec n" assumes "corthogonal us" assumes "w \ span (set us)" shows "adjuster n w us + w = 0\<^sub>v n" proof - define v U where "v = adjuster n w us + w" and "U = set us" have span: "v \ span U" unfolding v_def U_def apply (rule adjust_preserves_span[THEN iffD1]) using assms corthogonal_distinct by simp_all have v_carrier: "v \ carrier_vec n" by (simp add: v_def assms corthogonal_distinct) have "v \c us!i = 0" if "i < length us" for i unfolding v_def apply (rule adjust_zero) using that assms by simp_all hence "v \c u = 0" if "u \ U" for u by (metis assms(3) U_def corthogonal_distinct distinct_Ex1 that) hence ortho: "u \c v = 0" if "u \ U" for u apply (subst conjugate_zero_iff[symmetric]) apply (subst conjugate_vec_sprod_comm) using that us_carrier v_carrier apply (auto simp: U_def)[2] apply (subst conjugate_conjugate_sprod) using that us_carrier v_carrier by (auto simp: U_def) from span obtain a where v: "lincomb a U = v" apply atomize_elim apply (rule finite_in_span[simplified]) unfolding U_def using us_carrier by auto have "v \c v = (\u\U. (a u \\<^sub>v u) \c v)" apply (subst v[symmetric]) unfolding lincomb_def apply (subst finsum_scalar_prod_sum) using U_def span us_carrier by auto also have "\ = (\u\U. a u * (u \c v))" using U_def assms(1) in_mono us_carrier v_def by fastforce also have "\ = (\u\U. a u * conjugate 0)" apply (rule sum.cong, simp) using span span_closed U_def us_carrier ortho by auto also have "\ = 0" by auto finally have "v \c v = 0" by - thus "v = 0\<^sub>v n" using U_def conjugate_square_eq_0_vec span span_closed us_carrier by blast qed lemma (in cof_vec_space) gram_schmidt_sub0_result: assumes "gram_schmidt_sub0 n us ws = us'" and "set ws \ carrier_vec n" and "set us \ carrier_vec n" and "distinct us" and "~ lin_dep (set us)" and "corthogonal us" shows "set us' \ carrier_vec n \ distinct us' \ corthogonal us' \ span (set (us @ ws)) = span (set us')" using assms proof (induct ws arbitrary: us us') case (Cons w ws) show ?case proof (cases "w \ span (set us)") case False let ?v = "adjuster n w us" have wW[simp]: "set (w#ws) \ carrier_vec n" using Cons by simp hence W[simp]: "set ws \ carrier_vec n" and w[simp]: "w : carrier_vec n" by auto have U[simp]: "set us \ carrier_vec n" using Cons by simp have UW: "set (us@ws) \ carrier_vec n" by simp have wU: "set (w#us) \ carrier_vec n" by simp have dist_U: "distinct us" using Cons by simp have w_U: "w \ set us" using False using span_mem by auto have ind_U: "~ lin_dep (set us)" using Cons by simp have ind_wU: "~ lin_dep (insert w (set us))" apply (subst lin_dep_iff_in_span[simplified, symmetric]) using w_U ind_U False by auto thm lin_dep_iff_in_span[simplified, symmetric] have corth: "corthogonal us" using Cons by simp have "?v + w \ 0\<^sub>v n" by (simp add: False adjust_nonzero dist_U) hence "\ vec_is_zero n (?v + w)" by (simp add: vec_is_zero) hence U'def: "gram_schmidt_sub0 n ((?v + w)#us) ws = us'" using Cons by simp have v: "?v : carrier_vec n" using dist_U by auto hence vw: "?v + w : carrier_vec n" by auto hence vwU: "set ((?v + w) # us) \ carrier_vec n" by auto have vsU: "?v : span (set us)" apply (rule adjuster_in_span[OF w]) using Cons by simp_all hence vsUW: "?v : span (set (us @ ws))" using span_is_monotone[of "set us" "set (us@ws)"] by auto have wsU: "w \ span (set us)" using lin_dep_iff_in_span[OF U ind_U w w_U] ind_wU by auto hence vwU: "?v + w \ span (set us)" using adjust_not_in_span[OF w U dist_U] by auto have span: "?v + w \ span (set us)" apply (subst span_add[symmetric]) by (simp_all add: False vsU) hence vwUS: "?v + w \ set us" using span_mem by auto have vwU: "set ((?v + w) # us) \ carrier_vec n" using U w vw by simp have dist2: "distinct (((?v + w) # us))" using vwUS by (simp add: dist_U) have orth2: "corthogonal ((adjuster n w us + w) # us)" using adjust_orthogonal[OF U corth w wsU]. have ind_vwU: "~ lin_dep (set ((adjuster n w us + w) # us))" apply simp apply (subst lin_dep_iff_in_span[simplified, symmetric]) by (simp_all add: ind_U vw vwUS span) have span_UwW_U': "span (set (us @ w # ws)) = span (set us')" using Cons(1)[OF U'def W vwU dist2 ind_vwU orth2] using span_Un[OF vwU wU gram_schmidt_sub_span[OF w U dist_U] W W refl] by simp show ?thesis apply (intro conjI) using Cons(1)[OF U'def W vwU dist2 ind_vwU orth2] span_UwW_U' by simp_all next case True let ?v = "adjuster n w us" have "?v + w = 0\<^sub>v n" apply (rule adjuster_already_in_span) using True Cons by auto hence "vec_is_zero n (?v + w)" by (simp add: vec_is_zero) hence U'_def: "us' = gram_schmidt_sub0 n us ws" using Cons by simp have span: "span (set (us @ w # ws)) = span (set us')" proof - have wU_U: "span (set (w # us)) = span (set us)" apply (subst already_in_span[OF _ True, simplified]) using Cons by auto have "span (set (us @ w # ws)) = span (set (w # us) \ set ws)" by simp also have "\ = span (set us \ set ws)" apply (rule span_Un) using wU_U Cons by auto also have "\ = local.span (set us')" using Cons U'_def by auto finally show ?thesis by - qed moreover have "set us' \ carrier_vec n \ distinct us' \ corthogonal us'" unfolding U'_def using Cons by simp ultimately show ?thesis by auto qed qed simp text \This is a variant of \<^term>\Gram_Schmidt.gram_schmidt\ that does not require the input vectors \<^term>\ws\ to be distinct or orthogonal. (In comparison to \<^term>\Gram_Schmidt.gram_schmidt\, our version also returns the result in reversed order.)\ definition "gram_schmidt0 n ws = gram_schmidt_sub0 n [] ws" lemma (in cof_vec_space) gram_schmidt0_result: fixes ws defines "us' \ gram_schmidt0 n ws" assumes ws: "set ws \ carrier_vec n" shows "set us' \ carrier_vec n" (is ?thesis1) and "distinct us'" (is ?thesis2) and "corthogonal us'" (is ?thesis3) and "span (set ws) = span (set us')" (is ?thesis4) proof - have carrier_empty: "set [] \ carrier_vec n" by auto have distinct_empty: "distinct []" by simp have indep_empty: "lin_indpt (set [])" using basis_def subset_li_is_li unit_vecs_basis by auto have ortho_empty: "corthogonal []" by auto note gram_schmidt_sub0_result' = gram_schmidt_sub0_result [OF us'_def[symmetric, THEN meta_eq_to_obj_eq, unfolded gram_schmidt0_def] ws carrier_empty distinct_empty indep_empty ortho_empty] thus ?thesis1 ?thesis2 ?thesis3 ?thesis4 by auto qed locale complex_vec_space = cof_vec_space n "TYPE(complex)" for n :: nat lemma gram_schmidt0_corthogonal: assumes a1: "corthogonal R" and a2: "\x. x \ set R \ dim_vec x = d" shows "gram_schmidt0 d R = rev R" proof - have "gram_schmidt_sub0 d U R = rev R @ U" if "corthogonal ((rev U) @ R)" and "\x. x \ set U \ set R \ dim_vec x = d" for U proof (insert that, induction R arbitrary: U) case Nil show ?case by auto next case (Cons a R) have "a \ set (rev U @ a # R)" by simp moreover have uar: "corthogonal (rev U @ a # R)" by (simp add: Cons.prems(1)) ultimately have \a \ 0\<^sub>v d\ unfolding corthogonal_def by (metis conjugate_zero_vec in_set_conv_nth scalar_prod_right_zero zero_carrier_vec) then have nonzero_a: "\ vec_is_zero d a" by (simp add: Cons.prems(2) vec_is_zero) define T where "T = rev U @ a # R" have "T ! length (rev U) = a" unfolding T_def by (meson nth_append_length) moreover have "(T ! i \c T ! j = 0) = (i \ j)" if "ic T ! j = 0) = (length (rev U) \ j)" if "jc T ! j = 0" if "j length (rev U)" for j using that(1) that(2) by blast hence "a \c T ! j = 0" if "j < length (rev U)" for j using \T ! length (rev U) = a\ that(1) \length (rev U) < length T\ dual_order.strict_trans by blast moreover have "T ! j = (rev U) ! j" if "j < length (rev U)" for j by (smt T_def \length (rev U) < length T\ dual_order.strict_trans list_update_append1 list_update_id nth_list_update_eq that) ultimately have "a \c u = 0" if "u \ set (rev U)" for u by (metis in_set_conv_nth that) hence "a \c u = 0" if "u \ set U" for u by (simp add: that) moreover have "\x. x \ set U \ dim_vec x = d" by (simp add: Cons.prems(2)) ultimately have "adjuster d a U = 0\<^sub>v d" proof(induction U) case Nil then show ?case by simp next case (Cons u U) moreover have "0 \\<^sub>v u + 0\<^sub>v d = 0\<^sub>v d" proof- have "dim_vec u = d" by (simp add: calculation(3)) thus ?thesis by auto qed ultimately show ?case by auto qed hence adjuster_a: "adjuster d a U + a = a" by (simp add: Cons.prems(2) carrier_vecI) have "gram_schmidt_sub0 d U (a # R) = gram_schmidt_sub0 d (a # U) R" by (simp add: adjuster_a nonzero_a) also have "\ = rev (a # R) @ U" apply (subst Cons.IH) using Cons.prems by simp_all finally show ?case by - qed from this[where U="[]"] show ?thesis unfolding gram_schmidt0_def using assms by auto qed lemma adjuster_carrier': (* Like adjuster_carrier but with one assm less *) assumes w: "(w :: 'a::conjugatable_field vec) : carrier_vec n" and us: "set (us :: 'a vec list) \ carrier_vec n" shows "adjuster n w us \ carrier_vec n" by (insert us, induction us, auto) lemma eq_mat_on_vecI: fixes M N :: \'a::field mat\ assumes eq: \\v. v\carrier_vec nA \ M *\<^sub>v v = N *\<^sub>v v\ assumes [simp]: \M \ carrier_mat nB nA\ \N \ carrier_mat nB nA\ shows \M = N\ proof (rule eq_matI) show [simp]: \dim_row M = dim_row N\ \dim_col M = dim_col N\ using assms(2) assms(3) by blast+ fix i j assume [simp]: \i < dim_row N\ \j < dim_col N\ show \M $$ (i, j) = N $$ (i, j)\ thm mat_entry_explicit[where M=M] apply (subst mat_entry_explicit[symmetric]) using assms apply auto[3] apply (subst mat_entry_explicit[symmetric]) using assms apply auto[3] apply (subst eq) apply auto using assms(3) unit_vec_carrier by blast qed lemma list_of_vec_plus: fixes v1 v2 :: \complex vec\ assumes \dim_vec v1 = dim_vec v2\ shows \list_of_vec (v1 + v2) = map2 (+) (list_of_vec v1) (list_of_vec v2)\ proof- have \i < dim_vec v1 \ (list_of_vec (v1 + v2)) ! i = (map2 (+) (list_of_vec v1) (list_of_vec v2)) ! i\ for i by (simp add: assms) thus ?thesis by (metis assms index_add_vec(2) length_list_of_vec length_map map_fst_zip nth_equalityI) qed lemma list_of_vec_mult: fixes v :: \complex vec\ shows \list_of_vec (c \\<^sub>v v) = map ((*) c) (list_of_vec v)\ by (metis (mono_tags, lifting) index_smult_vec(1) index_smult_vec(2) length_list_of_vec length_map nth_equalityI nth_list_of_vec nth_map) unbundle no_jnf_notation end diff --git a/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy b/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy --- a/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy +++ b/thys/Complex_Bounded_Operators/extra/Extra_Ordered_Fields.thy @@ -1,958 +1,953 @@ section \\Extra_Ordered_Fields\ -- Additional facts about ordered fields\ theory Extra_Ordered_Fields imports Complex_Main Jordan_Normal_Form.Conjugate (* Defines ordering for complex. We have to use theirs, otherwise there will be conflicts *) begin subsection\Ordered Fields\ text \In this section we introduce some type classes for ordered rings/fields/etc. that are weakenings of existing classes. Most theorems in this section are copies of the eponymous theorems from Isabelle/HOL, except that they are now proven requiring weaker type classes (usually the need for a total order is removed). Since the lemmas are identical to the originals except for weaker type constraints, we use the same names as for the original lemmas. (In fact, the new lemmas could replace the original ones in Isabelle/HOL with at most minor incompatibilities.\ subsection \Missing from Orderings.thy\ text \This class is analogous to \<^class>\unbounded_dense_linorder\, except that it does not require a total order\ class unbounded_dense_order = dense_order + no_top + no_bot instance unbounded_dense_linorder \ unbounded_dense_order .. subsection \Missing from Rings.thy\ text \The existing class \<^class>\abs_if\ requires \<^term>\\a\ = (if a < 0 then - a else a)\. However, if \<^term>\(<)\ is not a total order, this condition is too strong when \<^term>\a\ is incomparable with \<^term>\0\. (Namely, it requires the absolute value to be the identity on such elements. E.g., the absolute value for complex numbers does not satisfy this.) The following class \partial_abs_if\ is analogous to \<^class>\abs_if\ but does not require anything if \<^term>\a\ is incomparable with \<^term>\0\.\ class partial_abs_if = minus + uminus + ord + zero + abs + assumes abs_neg: "a \ 0 \ abs a = -a" assumes abs_pos: "a \ 0 \ abs a = a" class ordered_semiring_1 = ordered_semiring + semiring_1 \ \missing class analogous to \<^class>\linordered_semiring_1\ without requiring a total order\ begin lemma convex_bound_le: assumes "x \ a" and "y \ a" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u * x + v * y \ a" proof- from assms have "u * x + v * y \ u * a + v * a" by (simp add: add_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end subclass (in linordered_semiring_1) ordered_semiring_1 .. class ordered_semiring_strict = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add + \ \missing class analogous to \<^class>\linordered_semiring_strict\ without requiring a total order\ assumes mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" assumes mult_strict_right_mono: "a < b \ 0 < c \ a * c < b * c" begin subclass semiring_0_cancel .. subclass ordered_semiring proof fix a b c :: 'a assume t1: "a \ b" and t2: "0 \ c" thus "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto from t2 show "a * c \ b * c" unfolding le_less by (metis local.antisym_conv2 local.mult_not_zero local.mult_strict_right_mono t1) qed lemma mult_pos_pos[simp]: "0 < a \ 0 < b \ 0 < a * b" using mult_strict_left_mono [of 0 b a] by simp lemma mult_pos_neg: "0 < a \ b < 0 \ a * b < 0" using mult_strict_left_mono [of b 0 a] by simp lemma mult_neg_pos: "a < 0 \ 0 < b \ a * b < 0" using mult_strict_right_mono [of a 0 b] by simp text \Strict monotonicity in both arguments\ lemma mult_strict_mono: assumes t1: "a < b" and t2: "c < d" and t3: "0 < b" and t4: "0 \ c" shows "a * c < b * d" proof- have "a * c < b * d" by (metis local.dual_order.order_iff_strict local.dual_order.strict_trans2 local.mult_strict_left_mono local.mult_strict_right_mono local.mult_zero_right t1 t2 t3 t4) thus ?thesis using assms by blast qed text \This weaker variant has more natural premises\ lemma mult_strict_mono': assumes "a < b" and "c < d" and "0 \ a" and "0 \ c" shows "a * c < b * d" by (rule mult_strict_mono) (insert assms, auto) lemma mult_less_le_imp_less: assumes t1: "a < b" and t2: "c \ d" and t3: "0 \ a" and t4: "0 < c" shows "a * c < b * d" using local.mult_strict_mono' local.mult_strict_right_mono local.order.order_iff_strict t1 t2 t3 t4 by auto lemma mult_le_less_imp_less: assumes "a \ b" and "c < d" and "0 < a" and "0 \ c" shows "a * c < b * d" by (metis assms(1) assms(2) assms(3) assms(4) local.antisym_conv2 local.dual_order.strict_trans1 local.mult_strict_left_mono local.mult_strict_mono) end subclass (in linordered_semiring_strict) ordered_semiring_strict proof show "c * a < c * b" if "a < b" and "0 < c" for a :: 'a and b and c using that by (simp add: local.mult_strict_left_mono) show "a * c < b * c" if "a < b" and "0 < c" for a :: 'a and b and c using that by (simp add: local.mult_strict_right_mono) qed class ordered_semiring_1_strict = ordered_semiring_strict + semiring_1 \ \missing class analogous to \<^class>\linordered_semiring_1_strict\ without requiring a total order\ begin subclass ordered_semiring_1 .. lemma convex_bound_lt: assumes "x < a" and "y < a" and "0 \ u" and "0 \ v" and "u + v = 1" shows "u * x + v * y < a" proof - from assms have "u * x + v * y < u * a + v * a" by (cases "u = 0") (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono) with assms show ?thesis unfolding distrib_right[symmetric] by simp qed end subclass (in linordered_semiring_1_strict) ordered_semiring_1_strict .. class ordered_comm_semiring_strict = comm_semiring_0 + ordered_cancel_ab_semigroup_add + \ \missing class analogous to \<^class>\linordered_comm_semiring_strict\ without requiring a total order\ assumes comm_mult_strict_left_mono: "a < b \ 0 < c \ c * a < c * b" begin subclass ordered_semiring_strict proof fix a b c :: 'a assume "a < b" and "0 < c" thus "c * a < c * b" by (rule comm_mult_strict_left_mono) thus "a * c < b * c" by (simp only: mult.commute) qed subclass ordered_cancel_comm_semiring proof fix a b c :: 'a assume "a \ b" and "0 \ c" thus "c * a \ c * b" unfolding le_less using mult_strict_left_mono by (cases "c = 0") auto qed end subclass (in linordered_comm_semiring_strict) ordered_comm_semiring_strict apply standard by (simp add: local.mult_strict_left_mono) class ordered_ring_strict = ring + ordered_semiring_strict + ordered_ab_group_add + partial_abs_if \ \missing class analogous to \<^class>\linordered_ring_strict\ without requiring a total order\ begin subclass ordered_ring .. lemma mult_strict_left_mono_neg: "b < a \ c < 0 \ c * a < c * b" using mult_strict_left_mono [of b a "- c"] by simp lemma mult_strict_right_mono_neg: "b < a \ c < 0 \ a * c < b * c" using mult_strict_right_mono [of b a "- c"] by simp lemma mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" using mult_strict_right_mono_neg [of a 0 b] by simp end lemmas mult_sign_intros = mult_nonneg_nonneg mult_nonneg_nonpos mult_nonpos_nonneg mult_nonpos_nonpos mult_pos_pos mult_pos_neg mult_neg_pos mult_neg_neg subsection \Ordered fields\ class ordered_field = field + order + ordered_comm_semiring_strict + ordered_ab_group_add + partial_abs_if \ \missing class analogous to \<^class>\linordered_field\ without requiring a total order\ begin lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) lemmas sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff lemmas (in -) sign_simps = algebra_simps zero_less_mult_iff mult_less_0_iff text\Simplify expressions equated with 1\ lemma zero_eq_1_divide_iff [simp]: "0 = 1 / a \ a = 0" by (cases "a = 0") (auto simp: field_simps) lemma one_divide_eq_0_iff [simp]: "1 / a = 0 \ a = 0" using zero_eq_1_divide_iff[of a] by simp text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ text\Simplify quotients that are compared with the value 1.\ text \Conditional Simplification Rules: No Case Splits\ lemma eq_divide_eq_1 [simp]: "(1 = b/a) = ((a \ 0 & a = b))" by (auto simp add: eq_divide_eq) lemma divide_eq_eq_1 [simp]: "(b/a = 1) = ((a \ 0 & a = b))" by (auto simp add: divide_eq_eq) end (* class ordered_field *) text \The following type class intends to capture some important properties that are common both to the real and the complex numbers. The purpose is to be able to state and prove lemmas that apply both to the real and the complex numbers without needing to state the lemma twice. \ class nice_ordered_field = ordered_field + zero_less_one + idom_abs_sgn + assumes positive_imp_inverse_positive: "0 < a \ 0 < inverse a" and inverse_le_imp_le: "inverse a \ inverse b \ 0 < a \ b \ a" and dense_le: "(\x. x < y \ x \ z) \ y \ z" and nn_comparable: "0 \ a \ 0 \ b \ a \ b \ b \ a" and abs_nn: "\x\ \ 0" begin subclass (in linordered_field) nice_ordered_field proof show "\a\ = - a" if "a \ 0" for a :: 'a using that by simp show "\a\ = a" if "0 \ a" for a :: 'a using that by simp show "0 < inverse a" if "0 < a" for a :: 'a using that by simp show "b \ a" if "inverse a \ inverse b" and "0 < a" for a :: 'a and b using that using local.inverse_le_imp_le by blast show "y \ z" if "\x::'a. x < y \ x \ z" for y and z using that using local.dense_le by blast show "a \ b \ b \ a" if "0 \ a" and "0 \ b" for a :: 'a and b using that by auto show "0 \ \x\" for x :: 'a by simp qed lemma comparable: assumes h1: "a \ c \ a \ c" and h2: "b \ c \ b \ c" shows "a \ b \ b \ a" proof- have "a \ b" if t1: "\ b \ a" and t2: "a \ c" and t3: "b \ c" proof- have "0 \ c-a" by (simp add: t2) moreover have "0 \ c-b" by (simp add: t3) ultimately have "c-a \ c-b \ c-a \ c-b" by (rule nn_comparable) hence "-a \ -b \ -a \ -b" using local.add_le_imp_le_right local.uminus_add_conv_diff by presburger thus ?thesis by (simp add: t1) qed moreover have "a \ b" if t1: "\ b \ a" and t2: "c \ a" and t3: "b \ c" proof- have "b \ a" using local.dual_order.trans t2 t3 by blast thus ?thesis using t1 by auto qed moreover have "a \ b" if t1: "\ b \ a" and t2: "c \ a" and t3: "c \ b" proof- have "0 \ a-c" by (simp add: t2) moreover have "0 \ b-c" by (simp add: t3) ultimately have "a-c \ b-c \ a-c \ b-c" by (rule nn_comparable) hence "a \ b \ a \ b" by (simp add: local.le_diff_eq) thus ?thesis by (simp add: t1) qed ultimately show ?thesis using assms by auto qed lemma negative_imp_inverse_negative: "a < 0 \ inverse a < 0" by (insert positive_imp_inverse_positive [of "-a"], simp add: nonzero_inverse_minus_eq less_imp_not_eq) lemma inverse_positive_imp_positive: assumes inv_gt_0: "0 < inverse a" and nz: "a \ 0" shows "0 < a" proof - have "0 < inverse (inverse a)" using inv_gt_0 by (rule positive_imp_inverse_positive) thus "0 < a" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma inverse_negative_imp_negative: assumes inv_less_0: "inverse a < 0" and nz: "a \ 0" shows "a < 0" proof- have "inverse (inverse a) < 0" using inv_less_0 by (rule negative_imp_inverse_negative) thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq) qed lemma linordered_field_no_lb: "\x. \y. y < x" proof fix x::'a have m1: "- (1::'a) < 0" by simp from add_strict_right_mono[OF m1, where c=x] have "(- 1) + x < x" by simp thus "\y. y < x" by blast qed lemma linordered_field_no_ub: "\x. \y. y > x" proof fix x::'a have m1: " (1::'a) > 0" by simp from add_strict_right_mono[OF m1, where c=x] have "1 + x > x" by simp thus "\y. y > x" by blast qed lemma less_imp_inverse_less: assumes less: "a < b" and apos: "0 < a" shows "inverse b < inverse a" using assms by (metis local.dual_order.strict_iff_order local.inverse_inverse_eq local.inverse_le_imp_le local.positive_imp_inverse_positive) lemma inverse_less_imp_less: "inverse a < inverse b \ 0 < a \ b < a" using local.inverse_le_imp_le local.order.strict_iff_order by blast text\Both premises are essential. Consider -1 and 1.\ lemma inverse_less_iff_less [simp]: "0 < a \ 0 < b \ inverse a < inverse b \ b < a" by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) lemma le_imp_inverse_le: "a \ b \ 0 < a \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less) lemma inverse_le_iff_le [simp]: "0 < a \ 0 < b \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) text\These results refer to both operands being negative. The opposite-sign case is trivial, since inverse preserves signs.\ lemma inverse_le_imp_le_neg: "inverse a \ inverse b \ b < 0 \ b \ a" by (metis local.inverse_le_imp_le local.inverse_minus_eq local.neg_0_less_iff_less local.neg_le_iff_le) lemma inverse_less_imp_less_neg: "inverse a < inverse b \ b < 0 \ b < a" using local.dual_order.strict_iff_order local.inverse_le_imp_le_neg by blast lemma inverse_less_iff_less_neg [simp]: "a < 0 \ b < 0 \ inverse a < inverse b \ b < a" by (metis local.antisym_conv2 local.inverse_less_imp_less_neg local.negative_imp_inverse_negative local.nonzero_inverse_inverse_eq local.order.strict_implies_order) lemma le_imp_inverse_le_neg: "a \ b \ b < 0 \ inverse b \ inverse a" by (force simp add: le_less less_imp_inverse_less_neg) lemma inverse_le_iff_le_neg [simp]: "a < 0 \ b < 0 \ inverse a \ inverse b \ b \ a" by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) lemma one_less_inverse: "0 < a \ a < 1 \ 1 < inverse a" using less_imp_inverse_less [of a 1, unfolded inverse_1] . lemma one_le_inverse: "0 < a \ a \ 1 \ 1 \ inverse a" using le_imp_inverse_le [of a 1, unfolded inverse_1] . lemma pos_le_divide_eq [field_simps]: assumes "0 < c" shows "a \ b / c \ a * c \ b" using assms by (metis local.divide_eq_imp local.divide_inverse_commute local.dual_order.order_iff_strict local.dual_order.strict_iff_order local.mult_right_mono local.mult_strict_left_mono local.nonzero_divide_eq_eq local.order.strict_implies_order local.positive_imp_inverse_positive) lemma pos_less_divide_eq [field_simps]: assumes "0 < c" shows "a < b / c \ a * c < b" using assms local.dual_order.strict_iff_order local.nonzero_divide_eq_eq local.pos_le_divide_eq by auto lemma neg_less_divide_eq [field_simps]: assumes "c < 0" shows "a < b / c \ b < a * c" by (metis assms local.minus_divide_divide local.mult_minus_right local.neg_0_less_iff_less local.neg_less_iff_less local.pos_less_divide_eq) lemma neg_le_divide_eq [field_simps]: assumes "c < 0" shows "a \ b / c \ b \ a * c" by (metis assms local.dual_order.order_iff_strict local.dual_order.strict_iff_order local.neg_less_divide_eq local.nonzero_divide_eq_eq) lemma pos_divide_le_eq [field_simps]: assumes "0 < c" shows "b / c \ a \ b \ a * c" by (metis assms local.dual_order.strict_iff_order local.nonzero_eq_divide_eq local.pos_le_divide_eq) lemma pos_divide_less_eq [field_simps]: assumes "0 < c" shows "b / c < a \ b < a * c" by (metis assms local.minus_divide_left local.mult_minus_left local.neg_less_iff_less local.pos_less_divide_eq) lemma neg_divide_le_eq [field_simps]: assumes "c < 0" shows "b / c \ a \ a * c \ b" by (metis assms local.minus_divide_left local.mult_minus_left local.neg_le_divide_eq local.neg_le_iff_le) lemma neg_divide_less_eq [field_simps]: assumes "c < 0" shows "b / c < a \ a * c < b" using assms local.dual_order.strict_iff_order local.neg_divide_le_eq by auto text\The following \field_simps\ rules are necessary, as minus is always moved atop of division but we want to get rid of division.\ lemma pos_le_minus_divide_eq [field_simps]: "0 < c \ a \ - (b / c) \ a * c \ - b" unfolding minus_divide_left by (rule pos_le_divide_eq) lemma neg_le_minus_divide_eq [field_simps]: "c < 0 \ a \ - (b / c) \ - b \ a * c" unfolding minus_divide_left by (rule neg_le_divide_eq) lemma pos_less_minus_divide_eq [field_simps]: "0 < c \ a < - (b / c) \ a * c < - b" unfolding minus_divide_left by (rule pos_less_divide_eq) lemma neg_less_minus_divide_eq [field_simps]: "c < 0 \ a < - (b / c) \ - b < a * c" unfolding minus_divide_left by (rule neg_less_divide_eq) lemma pos_minus_divide_less_eq [field_simps]: "0 < c \ - (b / c) < a \ - b < a * c" unfolding minus_divide_left by (rule pos_divide_less_eq) lemma neg_minus_divide_less_eq [field_simps]: "c < 0 \ - (b / c) < a \ a * c < - b" unfolding minus_divide_left by (rule neg_divide_less_eq) lemma pos_minus_divide_le_eq [field_simps]: "0 < c \ - (b / c) \ a \ - b \ a * c" unfolding minus_divide_left by (rule pos_divide_le_eq) lemma neg_minus_divide_le_eq [field_simps]: "c < 0 \ - (b / c) \ a \ a * c \ - b" unfolding minus_divide_left by (rule neg_divide_le_eq) lemma frac_less_eq: "y \ 0 \ z \ 0 \ x / y < w / z \ (x * z - w * y) / (y * z) < 0" by (subst less_iff_diff_less_0) (simp add: diff_frac_eq ) lemma frac_le_eq: "y \ 0 \ z \ 0 \ x / y \ w / z \ (x * z - w * y) / (y * z) \ 0" by (subst le_iff_diff_le_0) (simp add: diff_frac_eq ) text\Lemmas \sign_simps\ is a first attempt to automate proofs of positivity/negativity needed for \field_simps\. Have not added \sign_simps\ to \field_simps\ because the former can lead to case explosions.\ lemma divide_pos_pos[simp]: "0 < x \ 0 < y \ 0 < x / y" by(simp add:field_simps) lemma divide_nonneg_pos: "0 \ x \ 0 < y \ 0 \ x / y" by(simp add:field_simps) lemma divide_neg_pos: "x < 0 \ 0 < y \ x / y < 0" by(simp add:field_simps) lemma divide_nonpos_pos: "x \ 0 \ 0 < y \ x / y \ 0" by(simp add:field_simps) lemma divide_pos_neg: "0 < x \ y < 0 \ x / y < 0" by(simp add:field_simps) lemma divide_nonneg_neg: "0 \ x \ y < 0 \ x / y \ 0" by(simp add:field_simps) lemma divide_neg_neg: "x < 0 \ y < 0 \ 0 < x / y" by(simp add:field_simps) lemma divide_nonpos_neg: "x \ 0 \ y < 0 \ 0 \ x / y" by(simp add:field_simps) lemma divide_strict_right_mono: "a < b \ 0 < c \ a / c < b / c" by (simp add: less_imp_not_eq2 divide_inverse mult_strict_right_mono positive_imp_inverse_positive) lemma divide_strict_right_mono_neg: "b < a \ c < 0 \ a / c < b / c" by (simp add: local.neg_less_divide_eq) text\The last premise ensures that \<^term>\a\ and \<^term>\b\ have the same sign\ lemma divide_strict_left_mono: "b < a \ 0 < c \ 0 < a*b \ c / a < c / b" by (metis local.divide_neg_pos local.dual_order.strict_iff_order local.frac_less_eq local.less_iff_diff_less_0 local.mult_not_zero local.mult_strict_left_mono) lemma divide_left_mono: "b \ a \ 0 \ c \ 0 < a*b \ c / a \ c / b" using local.divide_cancel_left local.divide_strict_left_mono local.dual_order.order_iff_strict by auto lemma divide_strict_left_mono_neg: "a < b \ c < 0 \ 0 < a*b \ c / a < c / b" by (metis local.divide_strict_left_mono local.minus_divide_left local.neg_0_less_iff_less local.neg_less_iff_less mult_commute) lemma mult_imp_div_pos_le: "0 < y \ x \ z * y \ x / y \ z" by (subst pos_divide_le_eq, assumption+) lemma mult_imp_le_div_pos: "0 < y \ z * y \ x \ z \ x / y" by(simp add:field_simps) lemma mult_imp_div_pos_less: "0 < y \ x < z * y \ x / y < z" by(simp add:field_simps) lemma mult_imp_less_div_pos: "0 < y \ z * y < x \ z < x / y" by(simp add:field_simps) lemma frac_le: "0 \ x \ x \ y \ 0 < w \ w \ z \ x / z \ y / w" using local.mult_imp_div_pos_le local.mult_imp_le_div_pos local.mult_mono by auto lemma frac_less: "0 \ x \ x < y \ 0 < w \ w \ z \ x / z < y / w" proof- assume a1: "w \ z" assume a2: "0 < w" assume a3: "0 \ x" assume a4: "x < y" have f5: "a = 0 \ (b = c / a) = (b * a = c)" for a b c::'a by (meson local.nonzero_eq_divide_eq) have f6: "0 < z" using a2 a1 by (meson local.order.ordering_axioms order.strict_trans2) have "z \ 0" using a2 a1 by (meson local.leD) moreover have "x / z \ y / w" using a1 a2 a3 a4 local.frac_eq_eq local.mult_less_le_imp_less by fastforce ultimately have "x / z \ y / w" using f5 by (metis (no_types)) thus ?thesis using a4 a3 a2 a1 by (meson local.frac_le local.order.not_eq_order_implies_strict local.order.strict_implies_order) qed lemma frac_less2: "0 < x \ x \ y \ 0 < w \ w < z \ x / z < y / w" by (metis local.antisym_conv2 local.divide_cancel_left local.dual_order.strict_implies_order local.frac_le local.frac_less) lemma less_half_sum: "a < b \ a < (a+b) / (1+1)" by (metis local.add_pos_pos local.add_strict_left_mono local.mult_imp_less_div_pos local.semiring_normalization_rules(4) local.zero_less_one mult_commute) lemma gt_half_sum: "a < b \ (a+b)/(1+1) < b" by (metis local.add_pos_pos local.add_strict_left_mono local.mult_imp_div_pos_less local.semiring_normalization_rules(24) local.semiring_normalization_rules(4) local.zero_less_one mult_commute) subclass unbounded_dense_order proof fix x y :: 'a have less_add_one: "a < a + 1" for a::'a by auto from less_add_one show "\y. x < y" by blast from less_add_one have "x + (- 1) < (x + 1) + (- 1)" by (rule add_strict_right_mono) hence "x - 1 < x + 1 - 1" by simp hence "x - 1 < x" by (simp add: algebra_simps) thus "\y. y < x" .. show "x < y \ \z>x. z < y" by (blast intro!: less_half_sum gt_half_sum) qed lemma dense_le_bounded: fixes x y z :: 'a assumes "x < y" and *: "\w. \ x < w ; w < y \ \ w \ z" shows "y \ z" proof (rule dense_le) fix w assume "w < y" from dense[OF \x < y\] obtain u where "x < u" "u < y" by safe have "u \ w \ w \ u" using \u < y\ \w < y\ comparable local.order.strict_implies_order by blast thus "w \ z" using "*" \u < y\ \w < y\ \x < u\ local.dual_order.trans local.order.strict_trans2 by blast qed subclass field_abs_sgn .. lemma nonzero_abs_inverse: "a \ 0 \ \inverse a\ = inverse \a\" by (rule abs_inverse) lemma nonzero_abs_divide: "b \ 0 \ \a / b\ = \a\ / \b\" by (rule abs_divide) lemma field_le_epsilon: assumes e: "\e. 0 < e \ x \ y + e" shows "x \ y" proof (rule dense_le) fix t assume "t < x" hence "0 < x - t" by (simp add: less_diff_eq) from e [OF this] have "x + 0 \ x + (y - t)" by (simp add: algebra_simps) hence "0 \ y - t" by (simp only: add_le_cancel_left) thus "t \ y" by (simp add: algebra_simps) qed lemma inverse_positive_iff_positive [simp]: "(0 < inverse a) = (0 < a)" using local.positive_imp_inverse_positive by fastforce lemma inverse_negative_iff_negative [simp]: "(inverse a < 0) = (a < 0)" using local.negative_imp_inverse_negative by fastforce lemma inverse_nonnegative_iff_nonnegative [simp]: "0 \ inverse a \ 0 \ a" by (simp add: local.dual_order.order_iff_strict) lemma inverse_nonpositive_iff_nonpositive [simp]: "inverse a \ 0 \ a \ 0" using local.inverse_nonnegative_iff_nonnegative local.neg_0_le_iff_le by fastforce lemma one_less_inverse_iff: "1 < inverse x \ 0 < x \ x < 1" using less_trans[of 1 x 0 for x] by (metis local.dual_order.strict_trans local.inverse_1 local.inverse_less_imp_less local.inverse_positive_iff_positive local.one_less_inverse local.zero_less_one) lemma one_le_inverse_iff: "1 \ inverse x \ 0 < x \ x \ 1" by (metis local.dual_order.strict_trans1 local.inverse_1 local.inverse_le_imp_le local.inverse_positive_iff_positive local.one_le_inverse local.zero_less_one) lemma inverse_less_1_iff: "inverse x < 1 \ x \ 0 \ 1 < x" proof (rule) assume invx1: "inverse x < 1" have "inverse x \ 0 \ inverse x \ 0" using comparable invx1 local.order.strict_implies_order local.zero_less_one by blast then consider (leq0) "inverse x \ 0" | (pos) "inverse x > 0" | (zero) "inverse x = 0" using local.antisym_conv1 by blast thus "x \ 0 \ 1 < x" by (metis invx1 local.eq_refl local.inverse_1 inverse_less_imp_less inverse_nonpositive_iff_nonpositive inverse_positive_iff_positive) next assume "x \ 0 \ 1 < x" then consider (neg) "x \ 0" | (g1) "1 < x" by auto thus "inverse x < 1" by (metis local.dual_order.not_eq_order_implies_strict local.dual_order.strict_trans local.inverse_1 local.inverse_negative_iff_negative local.inverse_zero local.less_imp_inverse_less local.zero_less_one) qed lemma inverse_le_1_iff: "inverse x \ 1 \ x \ 0 \ 1 \ x" by (metis local.dual_order.order_iff_strict local.inverse_1 local.inverse_le_iff_le local.inverse_less_1_iff local.one_le_inverse_iff) text\Simplify expressions such as \0 < 1/x\ to \0 < x\\ lemma zero_le_divide_1_iff [simp]: "0 \ 1 / a \ 0 \ a" using local.dual_order.order_iff_strict local.inverse_eq_divide local.inverse_positive_iff_positive by auto lemma zero_less_divide_1_iff [simp]: "0 < 1 / a \ 0 < a" by (simp add: local.dual_order.strict_iff_order) lemma divide_le_0_1_iff [simp]: "1 / a \ 0 \ a \ 0" by (smt local.abs_0 local.abs_1 local.abs_divide local.abs_neg local.abs_nn local.divide_cancel_left local.le_minus_iff local.minus_divide_right local.zero_neq_one) lemma divide_less_0_1_iff [simp]: "1 / a < 0 \ a < 0" using local.dual_order.strict_iff_order by auto lemma divide_right_mono: "a \ b \ 0 \ c \ a/c \ b/c" using local.divide_cancel_right local.divide_strict_right_mono local.dual_order.order_iff_strict by blast lemma divide_right_mono_neg: "a \ b \ c \ 0 \ b / c \ a / c" by (metis local.divide_cancel_right local.divide_strict_right_mono_neg local.dual_order.strict_implies_order local.eq_refl local.le_imp_less_or_eq) lemma divide_left_mono_neg: "a \ b \ c \ 0 \ 0 < a * b \ c / a \ c / b" by (metis local.divide_left_mono local.minus_divide_left local.neg_0_le_iff_le local.neg_le_iff_le mult_commute) lemma divide_nonneg_nonneg [simp]: "0 \ x \ 0 \ y \ 0 \ x / y" using local.divide_eq_0_iff local.divide_nonneg_pos local.dual_order.order_iff_strict by blast lemma divide_nonpos_nonpos: "x \ 0 \ y \ 0 \ 0 \ x / y" using local.divide_nonpos_neg local.dual_order.order_iff_strict by auto lemma divide_nonneg_nonpos: "0 \ x \ y \ 0 \ x / y \ 0" by (metis local.divide_eq_0_iff local.divide_nonneg_neg local.dual_order.order_iff_strict) lemma divide_nonpos_nonneg: "x \ 0 \ 0 \ y \ x / y \ 0" using local.divide_nonpos_pos local.dual_order.order_iff_strict by auto text \Conditional Simplification Rules: No Case Splits\ lemma le_divide_eq_1_pos [simp]: "0 < a \ (1 \ b/a) = (a \ b)" by (simp add: local.pos_le_divide_eq) lemma le_divide_eq_1_neg [simp]: "a < 0 \ (1 \ b/a) = (b \ a)" by (metis local.le_divide_eq_1_pos local.minus_divide_divide local.neg_0_less_iff_less local.neg_le_iff_le) lemma divide_le_eq_1_pos [simp]: "0 < a \ (b/a \ 1) = (b \ a)" using local.pos_divide_le_eq by auto lemma divide_le_eq_1_neg [simp]: "a < 0 \ (b/a \ 1) = (a \ b)" by (metis local.divide_le_eq_1_pos local.minus_divide_divide local.neg_0_less_iff_less local.neg_le_iff_le) lemma less_divide_eq_1_pos [simp]: "0 < a \ (1 < b/a) = (a < b)" by (simp add: local.dual_order.strict_iff_order) lemma less_divide_eq_1_neg [simp]: "a < 0 \ (1 < b/a) = (b < a)" using local.dual_order.strict_iff_order by auto lemma divide_less_eq_1_pos [simp]: "0 < a \ (b/a < 1) = (b < a)" using local.divide_le_eq_1_pos local.dual_order.strict_iff_order by auto lemma divide_less_eq_1_neg [simp]: "a < 0 \ b/a < 1 \ a < b" using local.dual_order.strict_iff_order by auto lemma abs_div_pos: "0 < y \ \x\ / y = \x / y\" by (simp add: local.abs_pos) lemma zero_le_divide_abs_iff [simp]: "(0 \ a / \b\) = (0 \ a | b = 0)" proof assume assm: "0 \ a / \b\" have absb: "abs b \ 0" by (fact abs_nn) thus "0 \ a \ b = 0" using absb assm local.abs_eq_0_iff local.mult_nonneg_nonneg by fastforce next assume "0 \ a \ b = 0" then consider (a) "0 \ a" | (b) "b = 0" by atomize_elim auto thus "0 \ a / \b\" by (metis local.abs_eq_0_iff local.abs_nn local.divide_eq_0_iff local.divide_nonneg_nonneg) qed lemma divide_le_0_abs_iff [simp]: "(a / \b\ \ 0) = (a \ 0 | b = 0)" by (metis local.minus_divide_left local.neg_0_le_iff_le local.zero_le_divide_abs_iff) text\For creating values between \<^term>\u\ and \<^term>\v\.\ lemma scaling_mono: assumes "u \ v" and "0 \ r" and "r \ s" shows "u + r * (v - u) / s \ v" proof - have "r/s \ 1" using assms by (metis local.divide_le_eq_1_pos local.division_ring_divide_zero local.dual_order.order_iff_strict local.dual_order.trans local.zero_less_one) hence "(r/s) * (v - u) \ 1 * (v - u)" using assms(1) local.diff_ge_0_iff_ge local.mult_right_mono by blast thus ?thesis by (simp add: field_simps) qed end (* class nice_ordered_field *) code_identifier code_module Ordered_Fields \ (SML) Arith and (OCaml) Arith and (Haskell) Arith -subsection\Ordered Complex\ - -declare Conjugate.less_eq_complex_def[simp del] -declare Conjugate.less_complex_def[simp del] - subsection \Ordering on complex numbers\ instantiation complex :: nice_ordered_field begin instance proof intro_classes note defs = less_eq_complex_def less_complex_def abs_complex_def fix x y z a b c :: complex show "a \ 0 \ \a\ = - a" unfolding defs by (simp add: cmod_eq_Re complex_is_Real_iff) show "0 \ a \ \a\ = a" unfolding defs by (metis abs_of_nonneg cmod_eq_Re comp_apply complex.exhaust_sel complex_of_real_def zero_complex.simps(1) zero_complex.simps(2)) show "a < b \ 0 < c \ c * a < c * b" unfolding defs by auto show "0 < (1::complex)" unfolding defs by simp show "0 < a \ 0 < inverse a" unfolding defs by auto define ra ia rb ib rc ic where "ra = Re a" "ia = Im a" "rb = Re b" "ib = Im b" "rc = Re c" "ic = Im c" note ri = this[symmetric] hence "a = Complex ra ia" "b = Complex rb ib" "c = Complex rc ic" by auto note ri = this ri have "rb \ ra" if "1 / ra \ (if rb = 0 then 0 else 1 / rb)" and "ia = 0" and "0 < ra" and "ib = 0" proof(cases "rb = 0") case True thus ?thesis using that(3) by auto next case False thus ?thesis by (smt nice_ordered_field_class.frac_less2 that(1) that(3)) qed thus "inverse a \ inverse b \ 0 < a \ b \ a" unfolding defs ri by (auto simp: power2_eq_square) show "(\a. a < b \ a \ c) \ b \ c" unfolding defs ri by (metis complex.sel(1) complex.sel(2) dense less_le_not_le nice_ordered_field_class.linordered_field_no_lb not_le_imp_less) show "0 \ a \ 0 \ b \ a \ b \ b \ a" unfolding defs by auto show "0 \ \x\" unfolding defs by auto qed end lemma less_eq_complexI: "Re x \ Re y \ Im x = Im y \ x\y" unfolding less_eq_complex_def by simp lemma less_complexI: "Re x < Re y \ Im x = Im y \ x y \ complex_of_real x \ complex_of_real y" unfolding less_eq_complex_def by auto lemma complex_of_real_mono_iff[simp]: "complex_of_real x \ complex_of_real y \ x \ y" unfolding less_eq_complex_def by auto lemma complex_of_real_strict_mono_iff[simp]: "complex_of_real x < complex_of_real y \ x < y" unfolding less_complex_def by auto lemma complex_of_real_nn_iff[simp]: "0 \ complex_of_real y \ 0 \ y" unfolding less_eq_complex_def by auto lemma complex_of_real_pos_iff[simp]: "0 < complex_of_real y \ 0 < y" unfolding less_complex_def by auto lemma Re_mono: "x \ y \ Re x \ Re y" unfolding less_eq_complex_def by simp lemma comp_Im_same: "x \ y \ Im x = Im y" unfolding less_eq_complex_def by simp lemma Re_strict_mono: "x < y \ Re x < Re y" unfolding less_complex_def by simp lemma complex_of_real_cmod: assumes "x \ 0" shows "complex_of_real (cmod x) = x" by (metis Reals_cases abs_of_nonneg assms comp_Im_same complex_is_Real_iff complex_of_real_nn_iff norm_of_real zero_complex.simps(2)) end diff --git a/thys/Count_Complex_Roots/Count_Complex_Roots.thy b/thys/Count_Complex_Roots/Count_Complex_Roots.thy --- a/thys/Count_Complex_Roots/Count_Complex_Roots.thy +++ b/thys/Count_Complex_Roots/Count_Complex_Roots.thy @@ -1,3446 +1,3447 @@ (* Author: Wenda Li *) section \Procedures to count the number of complex roots\ theory Count_Complex_Roots imports Winding_Number_Eval.Winding_Number_Eval Extended_Sturm More_Polynomials Budan_Fourier.Sturm_Multiple_Roots begin subsection \Misc\ (*refined version of the library one with the same name by dropping unnecessary assumptions*) corollary path_image_part_circlepath_subset: assumes "r\0" shows "path_image(part_circlepath z r st tt) \ sphere z r" proof (cases "st\tt") case True then show ?thesis by (auto simp: assms path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) next case False then have "path_image(part_circlepath z r tt st) \ sphere z r" by (auto simp: assms path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) moreover have "path_image(part_circlepath z r tt st) = path_image(part_circlepath z r st tt)" using path_image_reversepath by fastforce ultimately show ?thesis by auto qed (*refined version of the library one with the same name by dropping unnecessary assumptions*) proposition in_path_image_part_circlepath: assumes "w \ path_image(part_circlepath z r st tt)" "0 \ r" shows "norm(w - z) = r" proof - have "w \ {c. dist z c = r}" by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) thus ?thesis by (simp add: dist_norm norm_minus_commute) qed lemma infinite_ball: fixes a :: "'a::euclidean_space" assumes "r > 0" shows "infinite (ball a r)" using uncountable_ball[OF assms, THEN uncountable_infinite] . lemma infinite_cball: fixes a :: "'a::euclidean_space" assumes "r > 0" shows "infinite (cball a r)" using uncountable_cball[OF assms, THEN uncountable_infinite,of a] . (*FIXME: to generalise*) lemma infinite_sphere: fixes a :: complex assumes "r > 0" shows "infinite (sphere a r)" proof - have "uncountable (path_image (circlepath a r))" apply (rule simple_path_image_uncountable) using simple_path_circlepath assms by simp then have "uncountable (sphere a r)" using assms by simp from uncountable_infinite[OF this] show ?thesis . qed lemma infinite_halfspace_Im_gt: "infinite {x. Im x > b}" apply (rule connected_uncountable[THEN uncountable_infinite,of _ "(b+1)* \" "(b+2)*\"]) by (auto intro!:convex_connected simp add: convex_halfspace_Im_gt) lemma (in ring_1) Ints_minus2: "- a \ \ \ a \ \" using Ints_minus[of "-a"] by auto lemma dvd_divide_Ints_iff: "b dvd a \ b=0 \ of_int a / of_int b \ (\ :: 'a :: {field,ring_char_0} set)" proof assume asm:"b dvd a \ b=0" let ?thesis = "of_int a / of_int b \ (\ :: 'a :: {field,ring_char_0} set)" have ?thesis when "b dvd a" proof - obtain c where "a=b * c" using \b dvd a\ unfolding dvd_def by auto then show ?thesis by (auto simp add:field_simps) qed moreover have ?thesis when "b=0" using that by auto ultimately show ?thesis using asm by auto next assume "of_int a / of_int b \ (\ :: 'a :: {field,ring_char_0} set)" from Ints_cases[OF this] obtain c where *:"(of_int::_ \ 'a) c= of_int a / of_int b" by metis have "b dvd a" when "b\0" proof - have "(of_int::_ \ 'a) a = of_int b * of_int c" using that * by auto then have "a = b * c" using of_int_eq_iff by fastforce then show ?thesis unfolding dvd_def by auto qed then show " b dvd a \ b = 0" by auto qed lemma of_int_div_field: assumes "d dvd n" shows "(of_int::_\'a::field_char_0) (n div d) = of_int n / of_int d" apply (subst (2) dvd_mult_div_cancel[OF assms,symmetric]) by (auto simp add:field_simps) lemma powr_eq_1_iff: assumes "a>0" shows "(a::real) powr b =1 \ a=1 \ b=0" proof assume "a powr b = 1" have "b * ln a = 0" using \a powr b = 1\ ln_powr[of a b] assms by auto then have "b=0 \ ln a =0" by auto then show "a = 1 \ b = 0" using assms by auto qed (insert assms, auto) lemma tan_inj_pi: "- (pi/2) < x \ x < pi/2 \ - (pi/2) < y \ y < pi/2 \ tan x = tan y \ x = y" by (metis arctan_tan) (*TODO: can we avoid fcompose in the proof?*) lemma finite_ReZ_segments_poly_circlepath: "finite_ReZ_segments (poly p \ circlepath z0 r) 0" proof (cases "\t\({0..1} - {1/2}). Re ((poly p \ circlepath z0 r) t) = 0") case True have "isCont (Re \ poly p \ circlepath z0 r) (1/2)" by (auto intro!:continuous_intros simp:circlepath) moreover have "(Re \ poly p \ circlepath z0 r)\ 1/2 \ 0" proof - have "\\<^sub>F x in at (1 / 2). (Re \ poly p \ circlepath z0 r) x = 0" unfolding eventually_at_le apply (rule exI[where x="1/2"]) unfolding dist_real_def abs_diff_le_iff by (auto intro!:True[rule_format, unfolded comp_def]) then show ?thesis by (rule tendsto_eventually) qed ultimately have "Re ((poly p \ circlepath z0 r) (1/2)) = 0" unfolding comp_def by (simp add: LIM_unique continuous_within) then have "\t\{0..1}. Re ((poly p \ circlepath z0 r) t) = 0" using True by blast then show ?thesis apply (rule_tac finite_ReZ_segments_constI[THEN finite_ReZ_segments_congE]) by auto next case False define q1 q2 where "q1=fcompose p [:(z0+r)*\,z0-r:] [:\,1:]" and "q2=([:\, 1:] ^ degree p)" define q1R q1I where "q1R=map_poly Re q1" and "q1I=map_poly Im q1" define q2R q2I where "q2R=map_poly Re q2" and "q2I=map_poly Im q2" define qq where "qq=q1R*q2R + q1I*q2I" have poly_eq:"Re ((poly p \ circlepath z0 r) t) = 0 \ poly qq (tan (pi * t)) = 0" when "0\t" "t\1" "t\1/2" for t proof - define tt where "tt=tan (pi * t)" have "Re ((poly p \ circlepath z0 r) t) = 0 \ Re (poly q1 tt / poly q2 tt) = 0" unfolding comp_def apply (subst poly_circlepath_tan_eq[of t p z0 r,folded q1_def q2_def tt_def]) using that by simp_all also have "... \ poly q1R tt * poly q2R tt + poly q1I tt * poly q2I tt = 0" unfolding q1I_def q1R_def q2R_def q2I_def by (simp add:Re_complex_div_eq_0 Re_poly_of_real Im_poly_of_real) also have "... \ poly qq tt = 0" unfolding qq_def by simp finally show ?thesis unfolding tt_def . qed have "finite {t. Re ((poly p \ circlepath z0 r) t) = 0 \ 0 \ t \ t \ 1}" proof - define P where "P=(\t. Re ((poly p \ circlepath z0 r) t) = 0)" define A where "A= ({0..1}::real set)" define S where "S={t\A-{1,1/2}. P t}" have "finite {t. poly qq (tan (pi * t)) = 0 \ 0 \ t \ t < 1 \ t\1/2}" proof - define A where "A={t::real. 0 \ t \ t < 1 \ t \ 1 / 2}" have "finite ((\t. tan (pi * t)) -` {x. poly qq x=0} \ A)" proof (rule finite_vimage_IntI) have "x = y" when "tan (pi * x) = tan (pi * y)" "x\A" "y\A" for x y proof - define x' where "x'=(if x<1/2 then x else x-1)" define y' where "y'=(if y<1/2 then y else y-1)" have "x'*pi = y'*pi" proof (rule tan_inj_pi) have *:"- 1 / 2 < x'" "x' < 1 / 2" "- 1 / 2 < y'" "y' < 1 / 2" using that(2,3) unfolding x'_def y'_def A_def by simp_all show "- (pi / 2) < x' * pi" "x' * pi < pi / 2" "- (pi / 2) < y' * pi" "y'*pi < pi / 2" using mult_strict_right_mono[OF *(1),of pi] mult_strict_right_mono[OF *(2),of pi] mult_strict_right_mono[OF *(3),of pi] mult_strict_right_mono[OF *(4),of pi] by auto next have "tan (x' * pi) = tan (x * pi)" unfolding x'_def using tan_periodic_int[of _ "- 1",simplified] by (auto simp add:algebra_simps) also have "... = tan (y * pi)" using \tan (pi * x) = tan (pi * y)\ by (auto simp:algebra_simps) also have "... = tan (y' * pi)" unfolding y'_def using tan_periodic_int[of _ "- 1",simplified] by (auto simp add:algebra_simps) finally show "tan (x' * pi) = tan (y' * pi)" . qed then have "x'=y'" by auto then show ?thesis using that(2,3) unfolding x'_def y'_def A_def by (auto split:if_splits) qed then show "inj_on (\t. tan (pi * t)) A" unfolding inj_on_def by blast next have "qq\0" proof (rule ccontr) assume "\ qq \ 0" then have "Re ((poly p \ circlepath z0 r) t) = 0" when "t\{0..1} - {1/2}" for t apply (subst poly_eq) using that by auto then show False using False by blast qed then show "finite {x. poly qq x = 0}" by (simp add: poly_roots_finite) qed then show ?thesis by (elim rev_finite_subset) (auto simp:A_def) qed moreover have "{t. poly qq (tan (pi * t)) = 0 \ 0 \ t \ t < 1 \ t\1/2} = S" unfolding S_def P_def A_def using poly_eq by force ultimately have "finite S" by blast then have "finite (S \ (if P 1 then {1} else {}) \ (if P (1/2) then {1/2} else {}))" by auto moreover have "(S \ (if P 1 then {1} else {}) \ (if P (1/2) then {1/2} else {})) = {t. P t \ 0 \ t \ t \ 1}" proof - have "1\A" "1/2 \A" unfolding A_def by auto then have "(S \ (if P 1 then {1} else {}) \ (if P (1/2) then {1/2} else {})) = {t\A. P t}" unfolding S_def apply auto by (metis eq_divide_eq_numeral1(1) zero_neq_numeral)+ also have "... = {t. P t \ 0 \ t \ t \ 1}" unfolding A_def by auto finally show ?thesis . qed ultimately have "finite {t. P t \ 0 \ t \ t \ 1}" by auto then show ?thesis unfolding P_def by simp qed then show ?thesis apply (rule_tac finite_imp_finite_ReZ_segments) by auto qed subsection \Some useful conformal/@{term bij_betw} properties\ lemma bij_betw_plane_ball:"bij_betw (\x. (\-x)/(\+x)) {x. Im x>0} (ball 0 1)" proof (rule bij_betw_imageI) have neq:"\ + x \0" when "Im x>0" for x using that by (metis add_less_same_cancel2 add_uminus_conv_diff diff_0 diff_add_cancel imaginary_unit.simps(2) not_one_less_zero uminus_complex.sel(2)) then show "inj_on (\x. (\ - x) / (\ + x)) {x. 0 < Im x}" unfolding inj_on_def by (auto simp add:divide_simps algebra_simps) have "cmod ((\ - x) / (\ + x)) < 1" when "0 < Im x " for x proof - have "cmod (\ - x) < cmod (\ + x)" unfolding norm_lt inner_complex_def using that by (auto simp add:algebra_simps) then show ?thesis unfolding norm_divide using neq[OF that] by auto qed moreover have "x \ (\x. (\ - x) / (\ + x)) ` {x. 0 < Im x}" when "cmod x < 1" for x proof (rule rev_image_eqI[of "\*(1-x)/(1+x)"]) have "1 + x\0" "\ * 2 + \ * (x * 2) \0" subgoal using that by (metis complex_mod_triangle_sub norm_one norm_zero not_le pth_7(1)) subgoal using that by (metis \1 + x \ 0\ complex_i_not_zero div_mult_self4 mult_2 mult_zero_right nonzero_mult_div_cancel_left nonzero_mult_div_cancel_right one_add_one zero_neq_numeral) done then show "x = (\ - \ * (1 - x) / (1 + x)) / (\ + \ * (1 - x) / (1 + x))" by (auto simp add:field_simps) show " \ * (1 - x) / (1 + x) \ {x. 0 < Im x}" apply (auto simp:Im_complex_div_gt_0 algebra_simps) using that unfolding cmod_def by (auto simp:power2_eq_square) qed ultimately show "(\x. (\ - x) / (\ + x)) ` {x. 0 < Im x} = ball 0 1" by auto qed lemma bij_betw_axis_sphere:"bij_betw (\x. (\-x)/(\+x)) {x. Im x=0} (sphere 0 1 - {-1})" proof (rule bij_betw_imageI) have neq:"\ + x \0" when "Im x=0" for x using that by (metis add_diff_cancel_left' imaginary_unit.simps(2) minus_complex.simps(2) right_minus_eq zero_complex.simps(2) zero_neq_one) then show "inj_on (\x. (\ - x) / (\ + x)) {x. Im x = 0}" unfolding inj_on_def by (auto simp add:divide_simps algebra_simps) have "cmod ((\ - x) / (\ + x)) = 1" "(\ - x) / (\ + x) \ - 1" when "Im x = 0" for x proof - have "cmod (\ + x) = cmod (\ - x)" using that unfolding cmod_def by auto then show "cmod ((\ - x) / (\ + x)) = 1" unfolding norm_divide using neq[OF that] by auto show "(\ - x) / (\ + x) \ - 1" using neq[OF that] by (auto simp add:divide_simps) qed moreover have "x \ (\x. (\ - x) / (\ + x)) ` {x. Im x = 0}" when "cmod x = 1" "x\-1" for x proof (rule rev_image_eqI[of "\*(1-x)/(1+x)"]) have "1 + x\0" "\ * 2 + \ * (x * 2) \0" subgoal using that(2) by algebra subgoal using that by (metis \1 + x \ 0\ complex_i_not_zero div_mult_self4 mult_2 mult_zero_right nonzero_mult_div_cancel_left nonzero_mult_div_cancel_right one_add_one zero_neq_numeral) done then show "x = (\ - \ * (1 - x) / (1 + x)) / (\ + \ * (1 - x) / (1 + x))" by (auto simp add:field_simps) show " \ * (1 - x) / (1 + x) \ {x. Im x = 0}" apply (auto simp:algebra_simps Im_complex_div_eq_0) using that(1) unfolding cmod_def by (auto simp:power2_eq_square) qed ultimately show "(\x. (\ - x) / (\ + x)) ` {x. Im x = 0} = sphere 0 1 - {- 1}" by force qed lemma bij_betw_ball_uball: assumes "r>0" shows "bij_betw (\x. complex_of_real r*x + z0) (ball 0 1) (ball z0 r)" proof (rule bij_betw_imageI) show "inj_on (\x. complex_of_real r * x + z0) (ball 0 1)" unfolding inj_on_def using assms by simp have "dist z0 (complex_of_real r * x + z0) < r" when "cmod x<1" for x using that assms by (auto simp:dist_norm norm_mult abs_of_pos) moreover have "x \ (\x. complex_of_real r * x + z0) ` ball 0 1" when "dist z0 x < r" for x apply (rule rev_image_eqI[of "(x-z0)/r"]) using that assms by (auto simp add: dist_norm norm_divide norm_minus_commute) ultimately show "(\x. complex_of_real r * x + z0) ` ball 0 1 = ball z0 r" by auto qed lemma bij_betw_sphere_usphere: assumes "r>0" shows "bij_betw (\x. complex_of_real r*x + z0) (sphere 0 1) (sphere z0 r)" proof (rule bij_betw_imageI) show "inj_on (\x. complex_of_real r * x + z0) (sphere 0 1)" unfolding inj_on_def using assms by simp have "dist z0 (complex_of_real r * x + z0) = r" when "cmod x=1" for x using that assms by (auto simp:dist_norm norm_mult abs_of_pos) moreover have "x \ (\x. complex_of_real r * x + z0) ` sphere 0 1" when "dist z0 x = r" for x apply (rule rev_image_eqI[of "(x-z0)/r"]) using that assms by (auto simp add: dist_norm norm_divide norm_minus_commute) ultimately show "(\x. complex_of_real r * x + z0) ` sphere 0 1 = sphere z0 r" by auto qed lemma proots_ball_plane_eq: defines "q1\[:\,-1:]" and "q2\[:\,1:]" assumes "p\0" shows "proots_count p (ball 0 1) = proots_count (fcompose p q1 q2) {x. 0 < Im x}" unfolding q1_def q2_def proof (rule proots_fcompose_bij_eq[OF _ \p\0\]) show "\x\{x. 0 < Im x}. poly [:\, 1:] x \ 0" apply simp by (metis add_less_same_cancel2 imaginary_unit.simps(2) not_one_less_zero plus_complex.simps(2) zero_complex.simps(2)) show "infinite (UNIV::complex set)" by (simp add: infinite_UNIV_char_0) qed (use bij_betw_plane_ball in auto) lemma proots_sphere_axis_eq: defines "q1\[:\,-1:]" and "q2\[:\,1:]" assumes "p\0" shows "proots_count p (sphere 0 1 - {- 1}) = proots_count (fcompose p q1 q2) {x. 0 = Im x}" unfolding q1_def q2_def proof (rule proots_fcompose_bij_eq[OF _ \p\0\]) show "\x\{x. 0 = Im x}. poly [:\, 1:] x \ 0" by (simp add: Complex_eq_0 plus_complex.code) show "infinite (UNIV::complex set)" by (simp add: infinite_UNIV_char_0) qed (use bij_betw_axis_sphere in auto) lemma proots_card_ball_plane_eq: defines "q1\[:\,-1:]" and "q2\[:\,1:]" assumes "p\0" shows "card (proots_within p (ball 0 1)) = card (proots_within (fcompose p q1 q2) {x. 0 < Im x})" unfolding q1_def q2_def proof (rule proots_card_fcompose_bij_eq[OF _ \p\0\]) show "\x\{x. 0 < Im x}. poly [:\, 1:] x \ 0" apply simp by (metis add_less_same_cancel2 imaginary_unit.simps(2) not_one_less_zero plus_complex.simps(2) zero_complex.simps(2)) qed (use bij_betw_plane_ball infinite_UNIV_char_0 in auto) lemma proots_card_sphere_axis_eq: defines "q1\[:\,-1:]" and "q2\[:\,1:]" assumes "p\0" shows "card (proots_within p (sphere 0 1 - {- 1})) = card (proots_within (fcompose p q1 q2) {x. 0 = Im x})" unfolding q1_def q2_def proof (rule proots_card_fcompose_bij_eq[OF _ \p\0\]) show "\x\{x. 0 = Im x}. poly [:\, 1:] x \ 0" by (simp add: Complex_eq_0 plus_complex.code) qed (use bij_betw_axis_sphere infinite_UNIV_char_0 in auto) lemma proots_uball_eq: fixes z0::complex and r::real defines "q\[:z0, of_real r:]" assumes "p\0" and "r>0" shows "proots_count p (ball z0 r) = proots_count (p \\<^sub>p q) (ball 0 1)" proof - show ?thesis apply (rule proots_pcompose_bij_eq[OF _ \p\0\]) subgoal unfolding q_def using bij_betw_ball_uball[OF \r>0\,of z0] by (auto simp:algebra_simps) subgoal unfolding q_def using \r>0\ by auto done qed lemma proots_card_uball_eq: fixes z0::complex and r::real defines "q\[:z0, of_real r:]" assumes "r>0" shows "card (proots_within p (ball z0 r)) = card (proots_within (p \\<^sub>p q) (ball 0 1))" proof - have ?thesis when "p=0" proof - have "card (ball z0 r) = 0" "card (ball (0::complex) 1) = 0" using infinite_ball[OF \r>0\,of z0] infinite_ball[of 1 "0::complex"] by auto then show ?thesis using that by auto qed moreover have ?thesis when "p\0" apply (rule proots_card_pcompose_bij_eq[OF _ \p\0\]) subgoal unfolding q_def using bij_betw_ball_uball[OF \r>0\,of z0] by (auto simp:algebra_simps) subgoal unfolding q_def using \r>0\ by auto done ultimately show ?thesis by blast qed lemma proots_card_usphere_eq: fixes z0::complex and r::real defines "q\[:z0, of_real r:]" assumes "r>0" shows "card (proots_within p (sphere z0 r)) = card (proots_within (p \\<^sub>p q) (sphere 0 1))" proof - have ?thesis when "p=0" proof - have "card (sphere z0 r) = 0" "card (sphere (0::complex) 1) = 0" using infinite_sphere[OF \r>0\,of z0] infinite_sphere[of 1 "0::complex"] by auto then show ?thesis using that by auto qed moreover have ?thesis when "p\0" apply (rule proots_card_pcompose_bij_eq[OF _ \p\0\]) subgoal unfolding q_def using bij_betw_sphere_usphere[OF \r>0\,of z0] by (auto simp:algebra_simps) subgoal unfolding q_def using \r>0\ by auto done ultimately show "card (proots_within p (sphere z0 r)) = card (proots_within (p \\<^sub>p q) (sphere 0 1))" by blast qed subsection \Combining two real polynomials into a complex one\ definition cpoly_of:: "real poly \ real poly \ complex poly" where "cpoly_of pR pI = map_poly of_real pR + smult \ (map_poly of_real pI)" lemma cpoly_of_eq_0_iff[iff]: "cpoly_of pR pI = 0 \ pR = 0 \ pI = 0" proof - have "pR = 0 \ pI = 0" when "cpoly_of pR pI = 0" proof - have "complex_of_real (coeff pR n) + \ * complex_of_real (coeff pI n) = 0" for n using that unfolding poly_eq_iff cpoly_of_def by (auto simp:coeff_map_poly) then have "coeff pR n = 0 \ coeff pI n = 0" for n by (metis Complex_eq Im_complex_of_real Re_complex_of_real complex.sel(1) complex.sel(2) of_real_0) then show ?thesis unfolding poly_eq_iff by auto qed then show ?thesis by (auto simp:cpoly_of_def) qed lemma cpoly_of_decompose: "p = cpoly_of (map_poly Re p) (map_poly Im p)" unfolding cpoly_of_def apply (induct p) by (auto simp add:map_poly_pCons map_poly_map_poly complex_eq) lemma cpoly_of_dist_right: "cpoly_of (pR*g) (pI*g) = cpoly_of pR pI * (map_poly of_real g)" unfolding cpoly_of_def by (simp add: distrib_right) lemma poly_cpoly_of_real: "poly (cpoly_of pR pI) (of_real x) = Complex (poly pR x) (poly pI x)" unfolding cpoly_of_def by (simp add: Complex_eq of_real_poly_map_poly) lemma poly_cpoly_of_real_iff: shows "poly (cpoly_of pR pI) (of_real t) =0 \ poly pR t = 0 \ poly pI t=0 " unfolding poly_cpoly_of_real using Complex_eq_0 by blast lemma order_cpoly_gcd_eq: assumes "pR\0 \ pI\0" shows "order t (cpoly_of pR pI) = order t (gcd pR pI)" proof - define g where "g = gcd pR pI" have [simp]:"g\0" unfolding g_def using assms by auto obtain pr pi where pri: "pR = pr * g" "pI = pi * g" "coprime pr pi" unfolding g_def using assms(1) gcd_coprime_exists \g \ 0\ g_def by blast then have "pr \0 \ pi \0" using assms mult_zero_left by blast have "order t (cpoly_of pR pI) = order t (cpoly_of pr pi * (map_poly of_real g))" unfolding pri cpoly_of_dist_right by simp also have "... = order t (cpoly_of pr pi) + order t g" apply (subst order_mult) using \pr \0 \ pi \0\ by (auto simp:map_poly_order_of_real) also have "... = order t g" proof - have "poly (cpoly_of pr pi) t \0" unfolding poly_cpoly_of_real_iff using \coprime pr pi\ coprime_poly_0 by blast then have "order t (cpoly_of pr pi) = 0" by (simp add: order_0I) then show ?thesis by auto qed finally show ?thesis unfolding g_def . qed subsection \Number of roots on a (bounded or unbounded) segment\ \ \1 dimensional hyperplane\ definition unbounded_line::"'a::real_vector \ 'a \ 'a set" where "unbounded_line a b = ({x. \u::real. x= (1 - u) *\<^sub>R a + u *\<^sub>R b})" definition proots_line_card:: "complex poly \ complex \ complex \ nat" where "proots_line_card p st tt = card (proots_within p (open_segment st tt))" definition proots_unbounded_line_card:: "complex poly \ complex \ complex \ nat" where "proots_unbounded_line_card p st tt = card (proots_within p (unbounded_line st tt))" definition proots_unbounded_line :: "complex poly \ complex \ complex \ nat" where "proots_unbounded_line p st tt = proots_count p (unbounded_line st tt)" lemma card_proots_open_segments: assumes "poly p st \0" "poly p tt \ 0" shows "card (proots_within p (open_segment st tt)) = (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in changes_itv_smods 0 1 g (pderiv g))" (is "?L = ?R") proof - define pc pR pI g where "pc = pcompose p [:st, tt-st:]" and "pR = map_poly Re pc" and "pI = map_poly Im pc" and "g = gcd pR pI" have poly_iff:"poly g t=0 \ poly pc t =0" for t proof - have "poly g t = 0 \ poly pR t =0 \ poly pI t =0" unfolding g_def using poly_gcd_iff by auto also have "... \ poly pc t =0" proof - have "cpoly_of pR pI = pc" unfolding pc_def pR_def pI_def using cpoly_of_decompose by auto then show ?thesis using poly_cpoly_of_real_iff by blast qed finally show ?thesis by auto qed have "?R = changes_itv_smods 0 1 g (pderiv g)" unfolding pc_def g_def pI_def pR_def by (auto simp add:Let_def) also have "... = card {t. poly g t = 0 \ 0 < t \ t < 1}" proof - have "poly g 0 \ 0" using poly_iff[of 0] assms unfolding pc_def by (auto simp add:poly_pcompose) moreover have "poly g 1 \ 0" using poly_iff[of 1] assms unfolding pc_def by (auto simp add:poly_pcompose) ultimately show ?thesis using sturm_interval[of 0 1 g] by auto qed - also have "... = card {t::real. poly pc t = 0 \ 0 < t \ t < 1}" + also have "... = card {t::real. poly pc (of_real t) = 0 \ 0 < t \ t < 1}" unfolding poly_iff by simp also have "... = ?L" proof (cases "st=tt") case True then show ?thesis unfolding pc_def poly_pcompose using \poly p tt \ 0\ by auto next case False define ff where "ff = (\t::real. st + t*(tt-st))" define ll where "ll = {t. poly pc (complex_of_real t) = 0 \ 0 < t \ t < 1}" have "ff ` ll = proots_within p (open_segment st tt)" proof (rule equalityI) show "ff ` ll \ proots_within p (open_segment st tt)" unfolding ll_def ff_def pc_def poly_pcompose by (auto simp add:in_segment False scaleR_conv_of_real algebra_simps) next show "proots_within p (open_segment st tt) \ ff ` ll" proof clarify fix x assume asm:"x \ proots_within p (open_segment st tt)" then obtain u where "0R st + u *\<^sub>R tt" by (auto simp add:in_segment) then have "poly p ((1 - u) *\<^sub>R st + u *\<^sub>R tt) = 0" using asm by simp then have "u \ ll" unfolding ll_def pc_def poly_pcompose by (simp add:scaleR_conv_of_real algebra_simps \0 \u<1\) moreover have "x = ff u" unfolding ff_def using u by (auto simp add:algebra_simps scaleR_conv_of_real) ultimately show "x \ ff ` ll" by (rule rev_image_eqI[of "u"]) qed qed moreover have "inj_on ff ll" unfolding ff_def using False inj_on_def by fastforce ultimately show ?thesis unfolding ll_def using card_image[of ff] by fastforce qed finally show ?thesis by simp qed lemma unbounded_line_closed_segment: "closed_segment a b \ unbounded_line a b" unfolding unbounded_line_def closed_segment_def by auto lemma card_proots_unbounded_line: assumes "st\tt" shows "card (proots_within p (unbounded_line st tt)) = (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in nat (changes_R_smods g (pderiv g)))" (is "?L = ?R") proof - define pc pR pI g where "pc = pcompose p [:st, tt-st:]" and "pR = map_poly Re pc" and "pI = map_poly Im pc" and "g = gcd pR pI" have poly_iff:"poly g t=0 \ poly pc t =0" for t proof - have "poly g t = 0 \ poly pR t =0 \ poly pI t =0" unfolding g_def using poly_gcd_iff by auto also have "... \ poly pc t =0" proof - have "cpoly_of pR pI = pc" unfolding pc_def pR_def pI_def using cpoly_of_decompose by auto then show ?thesis using poly_cpoly_of_real_iff by blast qed finally show ?thesis by auto qed have "?R = nat (changes_R_smods g (pderiv g))" unfolding pc_def g_def pI_def pR_def by (auto simp add:Let_def) also have "... = card {t. poly g t = 0}" using sturm_R[of g] by simp also have "... = card {t::real. poly pc t = 0}" unfolding poly_iff by simp also have "... = ?L" proof (cases "st=tt") case True then show ?thesis unfolding pc_def poly_pcompose unbounded_line_def using assms by (auto simp add:proots_within_def) next case False define ff where "ff = (\t::real. st + t*(tt-st))" define ll where "ll = {t. poly pc (complex_of_real t) = 0}" have "ff ` ll = proots_within p (unbounded_line st tt)" proof (rule equalityI) show "ff ` ll \ proots_within p (unbounded_line st tt)" unfolding ll_def ff_def pc_def poly_pcompose by (auto simp add:unbounded_line_def False scaleR_conv_of_real algebra_simps) next show "proots_within p (unbounded_line st tt) \ ff ` ll" proof clarify fix x assume asm:"x \ proots_within p (unbounded_line st tt)" then obtain u where u:"x = (1 - u) *\<^sub>R st + u *\<^sub>R tt" by (auto simp add:unbounded_line_def) then have "poly p ((1 - u) *\<^sub>R st + u *\<^sub>R tt) = 0" using asm by simp then have "u \ ll" unfolding ll_def pc_def poly_pcompose by (simp add:scaleR_conv_of_real algebra_simps unbounded_line_def) moreover have "x = ff u" unfolding ff_def using u by (auto simp add:algebra_simps scaleR_conv_of_real) ultimately show "x \ ff ` ll" by (rule rev_image_eqI[of "u"]) qed qed moreover have "inj_on ff ll" unfolding ff_def using False inj_on_def by fastforce ultimately show ?thesis unfolding ll_def using card_image[of ff] by metis qed finally show ?thesis by simp qed lemma proots_unbounded_line: assumes "st\tt" "p\0" shows "(proots_count p (unbounded_line st tt)) = (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in nat (changes_R_smods_ext g (pderiv g)))" (is "?L = ?R") proof - define pc pR pI g where "pc = pcompose p [:st, tt-st:]" and "pR = map_poly Re pc" and "pI = map_poly Im pc" and "g = gcd pR pI" have [simp]: "g\0" "pc\0" proof - show "pc\0" using assms(1) assms(2) pc_def pcompose_eq_0 by fastforce then have "pR\0 \ pI\0" unfolding pR_def pI_def by (metis cpoly_of_decompose map_poly_0) then show "g\0" unfolding g_def by simp qed have order_eq:"order t g = order t pc" for t apply (subst order_cpoly_gcd_eq[of pR pI,folded g_def,symmetric]) subgoal using \g\0\ unfolding g_def by simp subgoal unfolding pR_def pI_def by (simp add:cpoly_of_decompose[symmetric]) done have "?R = nat (changes_R_smods_ext g (pderiv g))" unfolding pc_def g_def pI_def pR_def by (auto simp add:Let_def) also have "... = proots_count g UNIV" using sturm_ext_R[OF \g\0\] by auto also have "... = proots_count (map_poly complex_of_real g) (of_real ` UNIV)" apply (subst proots_count_of_real) by auto also have "... = proots_count (map_poly complex_of_real g) {x. Im x = 0}" apply (rule arg_cong2[where f=proots_count]) using Reals_def complex_is_Real_iff by auto also have "... = proots_count pc {x. Im x = 0}" apply (rule proots_count_cong) apply (metis (mono_tags) Im_complex_of_real Re_complex_of_real \g \ 0\ complex_surj map_poly_order_of_real mem_Collect_eq order_eq) by auto also have "... = proots_count p (unbounded_line st tt)" proof - have "poly [:st, tt - st:] ` {x. Im x = 0} = unbounded_line st tt" unfolding unbounded_line_def apply safe subgoal for _ x apply (rule_tac x="Re x" in exI) apply (simp add:algebra_simps) by (simp add: mult.commute scaleR_complex.code times_complex.code) subgoal for _ u apply (rule rev_image_eqI[of "of_real u"]) by (auto simp:scaleR_conv_of_real algebra_simps) done then show ?thesis unfolding pc_def apply (subst proots_pcompose) using \p\0\ \st\tt\ by auto qed finally show ?thesis by simp qed lemma proots_unbounded_line_card_code[code]: "proots_unbounded_line_card p st tt = (if st\tt then (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in nat (changes_R_smods g (pderiv g))) else Code.abort (STR ''proots_unbounded_line_card fails due to invalid hyperplanes.'') (\_. proots_unbounded_line_card p st tt))" unfolding proots_unbounded_line_card_def using card_proots_unbounded_line[of st tt p] by auto lemma proots_unbounded_line_code[code]: "proots_unbounded_line p st tt = ( if st\tt then if p\0 then (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in nat (changes_R_smods_ext g (pderiv g))) else Code.abort (STR ''proots_unbounded_line fails due to p=0'') (\_. proots_unbounded_line p st tt) else Code.abort (STR ''proots_unbounded_line fails due to invalid hyperplanes.'') (\_. proots_unbounded_line p st tt) )" unfolding proots_unbounded_line_def using proots_unbounded_line by auto subsection \Checking if there a polynomial root on a closed segment\ definition no_proots_line::"complex poly \ complex \ complex \ bool" where "no_proots_line p st tt = (proots_within p (closed_segment st tt) = {})" (*TODO: the proof can probably be simplified using Lemma card_proots_open_segments*) lemma no_proots_line_code[code]: "no_proots_line p st tt = (if poly p st \0 \ poly p tt \ 0 then (let pc = pcompose p [:st, tt - st:]; pR = map_poly Re pc; pI = map_poly Im pc; g = gcd pR pI in if changes_itv_smods 0 1 g (pderiv g) = 0 then True else False) else False)" (is "?L = ?R") proof (cases "poly p st \0 \ poly p tt \ 0") case False thus ?thesis unfolding no_proots_line_def by auto next case True then have "poly p st \ 0" "poly p tt \ 0" by auto define pc pR pI g where "pc = pcompose p [:st, tt-st:]" and "pR = map_poly Re pc" and "pI = map_poly Im pc" and "g = gcd pR pI" have poly_iff:"poly g t=0 \ poly pc t =0" for t proof - have "poly g t = 0 \ poly pR t =0 \ poly pI t =0" unfolding g_def using poly_gcd_iff by auto also have "... \ poly pc t =0" proof - have "cpoly_of pR pI = pc" unfolding pc_def pR_def pI_def using cpoly_of_decompose by auto then show ?thesis using poly_cpoly_of_real_iff by blast qed finally show ?thesis by auto qed have "?R = (changes_itv_smods 0 1 g (pderiv g) = 0)" using True unfolding pc_def g_def pI_def pR_def by (auto simp add:Let_def) also have "... = (card {x. poly g x = 0 \ 0 < x \ x < 1} = 0)" proof - have "poly g 0 \ 0" using poly_iff[of 0] True unfolding pc_def by (auto simp add:poly_pcompose) moreover have "poly g 1 \ 0" using poly_iff[of 1] True unfolding pc_def by (auto simp add:poly_pcompose) ultimately show ?thesis using sturm_interval[of 0 1 g] by auto qed - also have "... = ({x. poly g x = 0 \ 0 < x \ x < 1} = {})" + also have "... = ({x. poly g (of_real x) = 0 \ 0 < x \ x < 1} = {})" proof - have "g\0" proof (rule ccontr) assume "\ g \ 0" then have "poly pc 0 =0" using poly_iff[of 0] by auto then show False using True unfolding pc_def by (auto simp add:poly_pcompose) qed from poly_roots_finite[OF this] have "finite {x. poly g x = 0 \ 0 < x \ x < 1}" by auto then show ?thesis using card_eq_0_iff by auto qed also have "... = ?L" proof - - have "(\t. poly g t = 0 \ 0 < t \ t < 1) \ (\t::real. poly pc t = 0 \ 0 < t \ t < 1)" + have "(\t. poly g (of_real t) = 0 \ 0 < t \ t < 1) \ + (\t::real. poly pc (of_real t) = 0 \ 0 < t \ t < 1)" using poly_iff by auto also have "... \ (\x. x \ closed_segment st tt \ poly p x = 0)" proof assume "\t. poly pc (complex_of_real t) = 0 \ 0 < t \ t < 1" then obtain t where *:"poly pc (of_real t) = 0" and "0 < t" "t < 1" by auto define x where "x=poly [:st, tt - st:] t" have "x \ closed_segment st tt" using \0 \t<1\ unfolding x_def in_segment by (intro exI[where x=t],auto simp add: algebra_simps scaleR_conv_of_real) moreover have "poly p x=0" using * unfolding pc_def x_def by (auto simp add:poly_pcompose) ultimately show "\x. x \ closed_segment st tt \ poly p x = 0" by auto next assume "\x. x \ closed_segment st tt \ poly p x = 0" then obtain x where "x \ closed_segment st tt" "poly p x = 0" by auto then obtain t::real where *:"x = (1 - t) *\<^sub>R st + t *\<^sub>R tt" and "0\t" "t\1" unfolding in_segment by auto then have "x=poly [:st, tt - st:] t" by (auto simp add: algebra_simps scaleR_conv_of_real) then have "poly pc (complex_of_real t) = 0" using \poly p x=0\ unfolding pc_def by (auto simp add:poly_pcompose) moreover have "t\0" "t\1" using True * \poly p x=0\ by auto then have "00\t\ \t\1\ by auto ultimately show "\t. poly pc (complex_of_real t) = 0 \ 0 < t \ t < 1" by auto qed finally show ?thesis unfolding no_proots_line_def proots_within_def by blast qed finally show ?thesis by simp qed subsection \Counting roots in a rectangle\ definition proots_rectangle ::"complex poly \ complex \ complex \ nat" where "proots_rectangle p lb ub = proots_count p (box lb ub)" lemma closed_segment_imp_Re_Im: fixes x::complex assumes "x\closed_segment lb ub" shows "Re lb \ Re ub \ Re lb \ Re x \ Re x \ Re ub" "Im lb \ Im ub \ Im lb \ Im x \ Im x \ Im ub" proof - obtain u where x_u:"x=(1 - u) *\<^sub>R lb + u *\<^sub>R ub " and "0 \ u" "u \ 1" using assms unfolding closed_segment_def by auto have "Re lb \ Re x" when "Re lb \ Re ub" proof - have "Re x = Re ((1 - u) *\<^sub>R lb + u *\<^sub>R ub)" using x_u by blast also have "... = Re (lb + u *\<^sub>R (ub - lb))" by (auto simp add:algebra_simps) also have "... = Re lb + u * (Re ub - Re lb)" by auto also have "... \ Re lb" using \u\0\ \Re lb \ Re ub\ by auto finally show ?thesis . qed moreover have "Im lb \ Im x" when "Im lb \ Im ub" proof - have "Im x = Im ((1 - u) *\<^sub>R lb + u *\<^sub>R ub)" using x_u by blast also have "... = Im (lb + u *\<^sub>R (ub - lb))" by (auto simp add:algebra_simps) also have "... = Im lb + u * (Im ub - Im lb)" by auto also have "... \ Im lb" using \u\0\ \Im lb \ Im ub\ by auto finally show ?thesis . qed moreover have "Re x \ Re ub" when "Re lb \ Re ub" proof - have "Re x = Re ((1 - u) *\<^sub>R lb + u *\<^sub>R ub)" using x_u by blast also have "... = (1 - u) * Re lb + u * Re ub" by auto also have "... \ (1 - u) * Re ub + u * Re ub" using \u\1\ \Re lb \ Re ub\ by (auto simp add: mult_left_mono) also have "... = Re ub" by (auto simp add:algebra_simps) finally show ?thesis . qed moreover have "Im x \ Im ub" when "Im lb \ Im ub" proof - have "Im x = Im ((1 - u) *\<^sub>R lb + u *\<^sub>R ub)" using x_u by blast also have "... = (1 - u) * Im lb + u * Im ub" by auto also have "... \ (1 - u) * Im ub + u * Im ub" using \u\1\ \Im lb \ Im ub\ by (auto simp add: mult_left_mono) also have "... = Im ub" by (auto simp add:algebra_simps) finally show ?thesis . qed ultimately show "Re lb \ Re ub \ Re lb \ Re x \ Re x \ Re ub" "Im lb \ Im ub \ Im lb \ Im x \ Im x \ Im ub" by auto qed lemma closed_segment_degen_complex: "\Re lb = Re ub; Im lb \ Im ub \ \ x \ closed_segment lb ub \ Re x = Re lb \ Im lb \ Im x \ Im x \ Im ub " "\Im lb = Im ub; Re lb \ Re ub \ \ x \ closed_segment lb ub \ Im x = Im lb \ Re lb \ Re x \ Re x \ Re ub " proof - show "x \ closed_segment lb ub \ Re x = Re lb \ Im lb \ Im x \ Im x \ Im ub" when "Re lb = Re ub" "Im lb \ Im ub" proof show "Re x = Re lb \ Im lb \ Im x \ Im x \ Im ub" when "x \ closed_segment lb ub" using closed_segment_imp_Re_Im[OF that] \Re lb = Re ub\ \Im lb \ Im ub\ by fastforce next assume asm:"Re x = Re lb \ Im lb \ Im x \ Im x \ Im ub" define u where "u=(Im x - Im lb)/ (Im ub - Im lb)" have "x = (1 - u) *\<^sub>R lb + u *\<^sub>R ub" unfolding u_def using asm \Re lb = Re ub\ \Im lb \ Im ub\ apply (intro complex_eqI) apply (auto simp add:field_simps) apply (cases "Im ub - Im lb =0") apply (auto simp add:field_simps) done moreover have "0\u" "u\1" unfolding u_def using \Im lb \ Im ub\ asm by (cases "Im ub - Im lb =0",auto simp add:field_simps)+ ultimately show "x \ closed_segment lb ub" unfolding closed_segment_def by auto qed show "x \ closed_segment lb ub \ Im x = Im lb \ Re lb \ Re x \ Re x \ Re ub" when "Im lb = Im ub" "Re lb \ Re ub" proof show "Im x = Im lb \ Re lb \ Re x \ Re x \ Re ub" when "x \ closed_segment lb ub" using closed_segment_imp_Re_Im[OF that] \Im lb = Im ub\ \Re lb \ Re ub\ by fastforce next assume asm:"Im x = Im lb \ Re lb \ Re x \ Re x \ Re ub" define u where "u=(Re x - Re lb)/ (Re ub - Re lb)" have "x = (1 - u) *\<^sub>R lb + u *\<^sub>R ub" unfolding u_def using asm \Im lb = Im ub\ \Re lb \ Re ub\ apply (intro complex_eqI) apply (auto simp add:field_simps) apply (cases "Re ub - Re lb =0") apply (auto simp add:field_simps) done moreover have "0\u" "u\1" unfolding u_def using \Re lb \ Re ub\ asm by (cases "Re ub - Re lb =0",auto simp add:field_simps)+ ultimately show "x \ closed_segment lb ub" unfolding closed_segment_def by auto qed qed lemma complex_box_ne_empty: fixes a b::complex shows "cbox a b \ {} \ (Re a \ Re b \ Im a \ Im b)" "box a b \ {} \ (Re a < Re b \ Im a < Im b)" by (auto simp add:box_ne_empty Basis_complex_def) lemma proots_rectangle_code1: "proots_rectangle p lb ub = (if Re lb < Re ub \ Im lb < Im ub then if p\0 then if no_proots_line p lb (Complex (Re ub) (Im lb)) \ no_proots_line p (Complex (Re ub) (Im lb)) ub \ no_proots_line p ub (Complex (Re lb) (Im ub)) \ no_proots_line p (Complex (Re lb) (Im ub)) lb then ( let p1 = pcompose p [:lb, Complex (Re ub - Re lb) 0:]; pR1 = map_poly Re p1; pI1 = map_poly Im p1; gc1 = gcd pR1 pI1; p2 = pcompose p [:Complex (Re ub) (Im lb), Complex 0 (Im ub - Im lb):]; pR2 = map_poly Re p2; pI2 = map_poly Im p2; gc2 = gcd pR2 pI2; p3 = pcompose p [:ub, Complex (Re lb - Re ub) 0:]; pR3 = map_poly Re p3; pI3 = map_poly Im p3; gc3 = gcd pR3 pI3; p4 = pcompose p [:Complex (Re lb) (Im ub), Complex 0 (Im lb - Im ub):]; pR4 = map_poly Re p4; pI4 = map_poly Im p4; gc4 = gcd pR4 pI4 in nat (- (changes_alt_itv_smods 0 1 (pR1 div gc1) (pI1 div gc1) + changes_alt_itv_smods 0 1 (pR2 div gc2) (pI2 div gc2) + changes_alt_itv_smods 0 1 (pR3 div gc3) (pI3 div gc3) + changes_alt_itv_smods 0 1 (pR4 div gc4) (pI4 div gc4)) div 4) ) else Code.abort (STR ''proots_rectangle fails when there is a root on the border.'') (\_. proots_rectangle p lb ub) else Code.abort (STR ''proots_rectangle fails when p=0.'') (\_. proots_rectangle p lb ub) else 0)" proof - have ?thesis when "\ (Re lb < Re ub \ Im lb < Im ub)" proof - have "box lb ub = {}" using complex_box_ne_empty[of lb ub] that by auto then have "proots_rectangle p lb ub = 0" unfolding proots_rectangle_def by auto then show ?thesis by (simp add:that) qed moreover have ?thesis when "Re lb < Re ub \ Im lb < Im ub" "p=0" using that by simp moreover have ?thesis when "Re lb < Re ub" "Im lb < Im ub" "p\0" and no_proots: "no_proots_line p lb (Complex (Re ub) (Im lb))" "no_proots_line p (Complex (Re ub) (Im lb)) ub" "no_proots_line p ub (Complex (Re lb) (Im ub))" "no_proots_line p (Complex (Re lb) (Im ub)) lb" proof - define l1 where "l1 = linepath lb (Complex (Re ub) (Im lb))" define l2 where "l2 = linepath (Complex (Re ub) (Im lb)) ub" define l3 where "l3 = linepath ub (Complex (Re lb) (Im ub))" define l4 where "l4 = linepath (Complex (Re lb) (Im ub)) lb" define rec where "rec = l1 +++ l2 +++ l3 +++ l4" have valid[simp]:"valid_path rec" and loop[simp]:"pathfinish rec = pathstart rec" unfolding rec_def l1_def l2_def l3_def l4_def by auto have path_no_proots:"path_image rec \ proots p = {}" unfolding rec_def l1_def l2_def l3_def l4_def apply (subst path_image_join,simp_all del:Complex_eq)+ using no_proots[unfolded no_proots_line_def] by (auto simp del:Complex_eq) define g1 where "g1 = poly p o l1" define g2 where "g2 = poly p o l2" define g3 where "g3 = poly p o l3" define g4 where "g4 = poly p o l4" have [simp]: "path g1" "path g2" "path g3" "path g4" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g3" "pathfinish g3 = pathstart g4" "pathfinish g4 = pathstart g1" unfolding g1_def g2_def g3_def g4_def l1_def l2_def l3_def l4_def by (auto intro!: path_continuous_image continuous_intros simp add:pathfinish_compose pathstart_compose) have [simp]: "finite_ReZ_segments g1 0" "finite_ReZ_segments g2 0" "finite_ReZ_segments g3 0" "finite_ReZ_segments g4 0" unfolding g1_def l1_def g2_def l2_def g3_def l3_def g4_def l4_def poly_linepath_comp by (rule finite_ReZ_segments_poly_of_real)+ define p1 pR1 pI1 gc1 p2 pR2 pI2 gc2 p3 pR3 pI3 gc3 p4 pR4 pI4 gc4 where "p1 = pcompose p [:lb, Complex (Re ub - Re lb) 0:]" and "pR1 = map_poly Re p1" and "pI1 = map_poly Im p1" and "gc1 = gcd pR1 pI1" and "p2 = pcompose p [:Complex (Re ub) (Im lb), Complex 0 (Im ub - Im lb):]" and "pR2 = map_poly Re p2" and "pI2 = map_poly Im p2" and "gc2 = gcd pR2 pI2" and "p3 = pcompose p [:ub, Complex (Re lb - Re ub) 0:]" and "pR3 = map_poly Re p3" and "pI3 = map_poly Im p3" and "gc3 = gcd pR3 pI3" and "p4 = pcompose p [:Complex (Re lb) (Im ub), Complex 0 (Im lb - Im ub):]" and "pR4 = map_poly Re p4" and "pI4 = map_poly Im p4" and "gc4 = gcd pR4 pI4" have "gc1\0" "gc2\0" "gc3\0" "gc4\0" proof - show "gc1\0" proof (rule ccontr) assume "\ gc1 \ 0" then have "pI1 = 0" "pR1 = 0" unfolding gc1_def by auto then have "p1 = 0" unfolding pI1_def pR1_def by (metis cpoly_of_decompose map_poly_0) then have "p=0" unfolding p1_def using \Re lb < Re ub\ by (auto elim!:pcompose_eq_0 simp add:Complex_eq_0) then show False using \p\0\ by simp qed show "gc2\0" proof (rule ccontr) assume "\ gc2 \ 0" then have "pI2 = 0" "pR2 = 0" unfolding gc2_def by auto then have "p2 = 0" unfolding pI2_def pR2_def by (metis cpoly_of_decompose map_poly_0) then have "p=0" unfolding p2_def using \Im lb < Im ub\ by (auto elim!:pcompose_eq_0 simp add:Complex_eq_0) then show False using \p\0\ by simp qed show "gc3\0" proof (rule ccontr) assume "\ gc3 \ 0" then have "pI3 = 0" "pR3 = 0" unfolding gc3_def by auto then have "p3 = 0" unfolding pI3_def pR3_def by (metis cpoly_of_decompose map_poly_0) then have "p=0" unfolding p3_def using \Re lb < Re ub\ by (auto elim!:pcompose_eq_0 simp add:Complex_eq_0) then show False using \p\0\ by simp qed show "gc4\0" proof (rule ccontr) assume "\ gc4 \ 0" then have "pI4 = 0" "pR4 = 0" unfolding gc4_def by auto then have "p4 = 0" unfolding pI4_def pR4_def by (metis cpoly_of_decompose map_poly_0) then have "p=0" unfolding p4_def using \Im lb < Im ub\ by (auto elim!:pcompose_eq_0 simp add:Complex_eq_0) then show False using \p\0\ by simp qed qed define sms where "sms = (changes_alt_itv_smods 0 1 (pR1 div gc1) (pI1 div gc1) + changes_alt_itv_smods 0 1 (pR2 div gc2) (pI2 div gc2) + changes_alt_itv_smods 0 1 (pR3 div gc3) (pI3 div gc3) + changes_alt_itv_smods 0 1 (pR4 div gc4) (pI4 div gc4))" have "proots_rectangle p lb ub = (\r\proots p. winding_number rec r * (order r p))" proof - have "winding_number rec x * of_nat (order x p) = 0" when "x\proots p - proots_within p (box lb ub)" for x proof - have *:"cbox lb ub = box lb ub \ path_image rec" proof - have "x\cbox lb ub" when "x\box lb ub \ path_image rec" for x using that \Re lb \Im lb unfolding box_def cbox_def Basis_complex_def rec_def l1_def l2_def l3_def l4_def apply (auto simp add:path_image_join closed_segment_degen_complex) apply (subst (asm) closed_segment_commute,simp add: closed_segment_degen_complex)+ done moreover have "x\box lb ub \ path_image rec" when "x\cbox lb ub" for x using that unfolding box_def cbox_def Basis_complex_def rec_def l1_def l2_def l3_def l4_def apply (auto simp add:path_image_join closed_segment_degen_complex) apply (subst (asm) (1 2) closed_segment_commute,simp add:closed_segment_degen_complex)+ done ultimately show ?thesis by auto qed moreover have "x\path_image rec" using path_no_proots that by auto ultimately have "x\cbox lb ub" using that by simp from winding_number_zero_outside[OF valid_path_imp_path[OF valid] _ loop this,simplified] * have "winding_number rec x = 0" by auto then show ?thesis by auto qed moreover have "of_nat (order x p) = winding_number rec x * of_nat (order x p)" when "x \ proots_within p (box lb ub)" for x proof - have "x\box lb ub" using that unfolding proots_within_def by auto then have order_asms:"Re lb < Re x" "Re x < Re ub" "Im lb < Im x" "Im x < Im ub" by (auto simp add:box_def Basis_complex_def) have "winding_number rec x = 1" unfolding rec_def l1_def l2_def l3_def l4_def proof eval_winding let ?l1 = "linepath lb (Complex (Re ub) (Im lb))" and ?l2 = "linepath (Complex (Re ub) (Im lb)) ub" and ?l3 = "linepath ub (Complex (Re lb) (Im ub))" and ?l4 = "linepath (Complex (Re lb) (Im ub)) lb" show l1: "x \ path_image ?l1" and l2: "x \ path_image ?l2" and l3: "x \ path_image ?l3" and l4:"x \ path_image ?l4" using no_proots that unfolding no_proots_line_def by auto show "- of_real (cindex_pathE ?l1 x + (cindex_pathE ?l2 x + (cindex_pathE ?l3 x + cindex_pathE ?l4 x))) = 2 * 1" proof - have "(Im x - Im ub) * (Re ub - Re lb) < 0" using mult_less_0_iff order_asms(1) order_asms(2) order_asms(4) by fastforce then have "cindex_pathE ?l3 x = -1" apply (subst cindex_pathE_linepath) using l3 by (auto simp add:algebra_simps order_asms) moreover have "(Im lb - Im x) * (Re ub - Re lb) <0" using mult_less_0_iff order_asms(1) order_asms(2) order_asms(3) by fastforce then have "cindex_pathE ?l1 x = -1" apply (subst cindex_pathE_linepath) using l1 by (auto simp add:algebra_simps order_asms) moreover have "cindex_pathE ?l2 x = 0" apply (subst cindex_pathE_linepath) using l2 order_asms by auto moreover have "cindex_pathE ?l4 x = 0" apply (subst cindex_pathE_linepath) using l4 order_asms by auto ultimately show ?thesis by auto qed qed then show ?thesis by auto qed ultimately show ?thesis using \p\0\ unfolding proots_rectangle_def proots_count_def by (auto intro!: sum.mono_neutral_cong_left[of "proots p" "proots_within p (box lb ub)"]) qed also have "... = 1/(2 * of_real pi * \) *contour_integral rec (\x. deriv (poly p) x / poly p x)" proof - have "contour_integral rec (\x. deriv (poly p) x / poly p x) = 2 * of_real pi * \ * (\x | poly p x = 0. winding_number rec x * of_int (zorder (poly p) x))" proof (rule argument_principle[of UNIV "poly p" "{}" "\_. 1" rec,simplified]) show "connected (UNIV::complex set)" using connected_UNIV[where 'a=complex] . show "path_image rec \ UNIV - {x. poly p x = 0}" using path_no_proots unfolding proots_within_def by auto show "finite {x. poly p x = 0}" by (simp add: poly_roots_finite that(3)) qed also have " ... = 2 * of_real pi * \ * (\x\proots p. winding_number rec x * (order x p))" unfolding proots_within_def apply (auto intro!:sum.cong simp add:order_zorder[OF \p\0\] ) by (metis nat_eq_iff2 of_nat_nat order_root order_zorder that(3)) finally show ?thesis by auto qed also have "... = winding_number (poly p \ rec) 0" proof - have "0 \ path_image (poly p \ rec)" using path_no_proots unfolding path_image_compose proots_within_def by fastforce from winding_number_comp[OF _ poly_holomorphic_on _ _ this,of UNIV,simplified] show ?thesis by auto qed also have winding_eq:"... = - cindex_pathE (poly p \ rec) 0 / 2" proof (rule winding_number_cindex_pathE) show "finite_ReZ_segments (poly p \ rec) 0" unfolding rec_def path_compose_join apply (fold g1_def g2_def g3_def g4_def) by (auto intro!: finite_ReZ_segments_joinpaths path_join_imp) show "valid_path (poly p \ rec)" by (rule valid_path_compose_holomorphic[where S=UNIV]) auto show "0 \ path_image (poly p \ rec)" using path_no_proots unfolding path_image_compose proots_def by fastforce show "pathfinish (poly p \ rec) = pathstart (poly p \ rec)" unfolding rec_def pathstart_compose pathfinish_compose by (auto simp add:l1_def l4_def) qed also have cindex_pathE_eq:"... = of_int (- sms) / of_int 4" proof - have "cindex_pathE (poly p \ rec) 0 = cindex_pathE (g1+++g2+++g3+++g4) 0" unfolding rec_def path_compose_join g1_def g2_def g3_def g4_def by simp also have "... = cindex_pathE g1 0 + cindex_pathE g2 0 + cindex_pathE g3 0 + cindex_pathE g4 0" by (subst cindex_pathE_joinpaths,auto intro!:finite_ReZ_segments_joinpaths)+ also have "... = cindex_polyE 0 1 (pI1 div gc1) (pR1 div gc1) + cindex_polyE 0 1 (pI2 div gc2) (pR2 div gc2) + cindex_polyE 0 1 (pI3 div gc3) (pR3 div gc3) + cindex_polyE 0 1 (pI4 div gc4) (pR4 div gc4)" proof - have "cindex_pathE g1 0 = cindex_polyE 0 1 (pI1 div gc1) (pR1 div gc1)" proof - have *:"g1 = poly p1 o of_real" unfolding g1_def p1_def l1_def poly_linepath_comp by (subst (5) complex_surj[symmetric],simp) then have "cindex_pathE g1 0 = cindexE 0 1 (\t. poly pI1 t / poly pR1 t)" unfolding cindex_pathE_def pR1_def pI1_def by (simp add:Im_poly_of_real Re_poly_of_real) also have "... = cindex_polyE 0 1 pI1 pR1" using cindexE_eq_cindex_polyE by auto also have "... = cindex_polyE 0 1 (pI1 div gc1) (pR1 div gc1)" using \gc1\0\ apply (subst (2) cindex_polyE_mult_cancel[of gc1,symmetric]) by (simp_all add: gc1_def) finally show ?thesis . qed moreover have "cindex_pathE g2 0 = cindex_polyE 0 1 (pI2 div gc2) (pR2 div gc2)" proof - have "g2 = poly p2 o of_real" unfolding g2_def p2_def l2_def poly_linepath_comp by (subst (5) complex_surj[symmetric],simp) then have "cindex_pathE g2 0 = cindexE 0 1 (\t. poly pI2 t / poly pR2 t)" unfolding cindex_pathE_def pR2_def pI2_def by (simp add:Im_poly_of_real Re_poly_of_real) also have "... = cindex_polyE 0 1 pI2 pR2" using cindexE_eq_cindex_polyE by auto also have "... = cindex_polyE 0 1 (pI2 div gc2) (pR2 div gc2)" using \gc2\0\ apply (subst (2) cindex_polyE_mult_cancel[of gc2,symmetric]) by (simp_all add: gc2_def) finally show ?thesis . qed moreover have "cindex_pathE g3 0 = cindex_polyE 0 1 (pI3 div gc3) (pR3 div gc3)" proof - have "g3 = poly p3 o of_real" unfolding g3_def p3_def l3_def poly_linepath_comp by (subst (5) complex_surj[symmetric],simp) then have "cindex_pathE g3 0 = cindexE 0 1 (\t. poly pI3 t / poly pR3 t)" unfolding cindex_pathE_def pR3_def pI3_def by (simp add:Im_poly_of_real Re_poly_of_real) also have "... = cindex_polyE 0 1 pI3 pR3" using cindexE_eq_cindex_polyE by auto also have "... = cindex_polyE 0 1 (pI3 div gc3) (pR3 div gc3)" using \gc3\0\ apply (subst (2) cindex_polyE_mult_cancel[of gc3,symmetric]) by (simp_all add: gc3_def) finally show ?thesis . qed moreover have "cindex_pathE g4 0 = cindex_polyE 0 1 (pI4 div gc4) (pR4 div gc4)" proof - have "g4 = poly p4 o of_real" unfolding g4_def p4_def l4_def poly_linepath_comp by (subst (5) complex_surj[symmetric],simp) then have "cindex_pathE g4 0 = cindexE 0 1 (\t. poly pI4 t / poly pR4 t)" unfolding cindex_pathE_def pR4_def pI4_def by (simp add:Im_poly_of_real Re_poly_of_real) also have "... = cindex_polyE 0 1 pI4 pR4" using cindexE_eq_cindex_polyE by auto also have "... = cindex_polyE 0 1 (pI4 div gc4) (pR4 div gc4)" using \gc4\0\ apply (subst (2) cindex_polyE_mult_cancel[of gc4,symmetric]) by (simp_all add: gc4_def) finally show ?thesis . qed ultimately show ?thesis by auto qed also have "... = sms / 2" proof - have "cindex_polyE 0 1 (pI1 div gc1) (pR1 div gc1) = changes_alt_itv_smods 0 1 (pR1 div gc1) (pI1 div gc1) / 2" apply (rule cindex_polyE_changes_alt_itv_mods) using \gc1\0\ unfolding gc1_def by (auto intro:div_gcd_coprime) moreover have "cindex_polyE 0 1 (pI2 div gc2) (pR2 div gc2) = changes_alt_itv_smods 0 1 (pR2 div gc2) (pI2 div gc2) / 2" apply (rule cindex_polyE_changes_alt_itv_mods) using \gc2\0\ unfolding gc2_def by (auto intro:div_gcd_coprime) moreover have "cindex_polyE 0 1 (pI3 div gc3) (pR3 div gc3) = changes_alt_itv_smods 0 1 (pR3 div gc3) (pI3 div gc3) / 2" apply (rule cindex_polyE_changes_alt_itv_mods) using \gc3\0\ unfolding gc3_def by (auto intro:div_gcd_coprime) moreover have "cindex_polyE 0 1 (pI4 div gc4) (pR4 div gc4) = changes_alt_itv_smods 0 1 (pR4 div gc4) (pI4 div gc4) / 2" apply (rule cindex_polyE_changes_alt_itv_mods) using \gc4\0\ unfolding gc4_def by (auto intro:div_gcd_coprime) ultimately show ?thesis unfolding sms_def by auto qed finally have *:"cindex_pathE (poly p \ rec) 0 = real_of_int sms / 2" . show ?thesis apply (subst *) by auto qed finally have "(of_nat::_\complex) (proots_rectangle p lb ub) = of_int (- sms) / of_int 4" . moreover have "4 dvd sms" proof - have "winding_number (poly p \ rec) 0 \ \" proof (rule integer_winding_number) show "path (poly p \ rec)" by (auto intro!:valid_path_compose_holomorphic[where S=UNIV] valid_path_imp_path) show "pathfinish (poly p \ rec) = pathstart (poly p \ rec)" unfolding rec_def path_compose_join by (auto simp add:l1_def l4_def pathfinish_compose pathstart_compose) show "0 \ path_image (poly p \ rec)" using path_no_proots unfolding path_image_compose proots_def by fastforce qed then have "of_int (- sms) / of_int 4 \ (\::complex set)" by (simp only: winding_eq cindex_pathE_eq) then show ?thesis by (subst (asm) dvd_divide_Ints_iff[symmetric],auto) qed ultimately have "proots_rectangle p lb ub = nat (- sms div 4)" apply (subst (asm) of_int_div_field[symmetric]) by (simp,metis nat_int of_int_eq_iff of_int_of_nat_eq) then show ?thesis unfolding Let_def apply (fold p1_def p2_def p3_def p4_def pI1_def pR1_def pI2_def pR2_def pI3_def pR3_def pI4_def pR4_def gc1_def gc2_def gc3_def gc4_def) apply (fold sms_def) using that by auto qed ultimately show ?thesis by fastforce qed (*further refinement*) lemma proots_rectangle_code2[code]: "proots_rectangle p lb ub = (if Re lb < Re ub \ Im lb < Im ub then if p\0 then if poly p lb \ 0 \ poly p (Complex (Re ub) (Im lb)) \0 \ poly p ub \0 \ poly p (Complex (Re lb) (Im ub)) \0 then (let p1 = pcompose p [:lb, Complex (Re ub - Re lb) 0:]; pR1 = map_poly Re p1; pI1 = map_poly Im p1; gc1 = gcd pR1 pI1; p2 = pcompose p [:Complex (Re ub) (Im lb), Complex 0 (Im ub - Im lb):]; pR2 = map_poly Re p2; pI2 = map_poly Im p2; gc2 = gcd pR2 pI2; p3 = pcompose p [:ub, Complex (Re lb - Re ub) 0:]; pR3 = map_poly Re p3; pI3 = map_poly Im p3; gc3 = gcd pR3 pI3; p4 = pcompose p [:Complex (Re lb) (Im ub), Complex 0 (Im lb - Im ub):]; pR4 = map_poly Re p4; pI4 = map_poly Im p4; gc4 = gcd pR4 pI4 in if changes_itv_smods 0 1 gc1 (pderiv gc1) = 0 \ changes_itv_smods 0 1 gc2 (pderiv gc2) = 0 \ changes_itv_smods 0 1 gc3 (pderiv gc3) = 0 \ changes_itv_smods 0 1 gc4 (pderiv gc4) = 0 then nat (- (changes_alt_itv_smods 0 1 (pR1 div gc1) (pI1 div gc1) + changes_alt_itv_smods 0 1 (pR2 div gc2) (pI2 div gc2) + changes_alt_itv_smods 0 1 (pR3 div gc3) (pI3 div gc3) + changes_alt_itv_smods 0 1 (pR4 div gc4) (pI4 div gc4)) div 4) else Code.abort (STR ''proots_rectangle fails when there is a root on the border.'') (\_. proots_rectangle p lb ub)) else Code.abort (STR ''proots_rectangle fails when there is a root on the border.'') (\_. proots_rectangle p lb ub) else Code.abort (STR ''proots_rectangle fails when p=0.'') (\_. proots_rectangle p lb ub) else 0)" proof - define p1 pR1 pI1 gc1 p2 pR2 pI2 gc2 p3 pR3 pI3 gc3 p4 pR4 pI4 gc4 where "p1 = pcompose p [:lb, Complex (Re ub - Re lb) 0:]" and "pR1 = map_poly Re p1" and "pI1 = map_poly Im p1" and "gc1 = gcd pR1 pI1" and "p2 = pcompose p [:Complex (Re ub) (Im lb), Complex 0 (Im ub - Im lb):]" and "pR2 = map_poly Re p2" and "pI2 = map_poly Im p2" and "gc2 = gcd pR2 pI2" and "p3 = pcompose p [:ub, Complex (Re lb - Re ub) 0:]" and "pR3 = map_poly Re p3" and "pI3 = map_poly Im p3" and "gc3 = gcd pR3 pI3" and "p4 = pcompose p [:Complex (Re lb) (Im ub), Complex 0 (Im lb - Im ub):]" and "pR4 = map_poly Re p4" and "pI4 = map_poly Im p4" and "gc4 = gcd pR4 pI4" define sms where "sms = (- (changes_alt_itv_smods 0 1 (pR1 div gc1) (pI1 div gc1) + changes_alt_itv_smods 0 1 (pR2 div gc2) (pI2 div gc2) + changes_alt_itv_smods 0 1 (pR3 div gc3) (pI3 div gc3) + changes_alt_itv_smods 0 1 (pR4 div gc4) (pI4 div gc4)) div 4)" have more_folds:"p1 = p \\<^sub>p [:lb, Complex (Re ub) (Im lb) - lb:]" "p2 = p \\<^sub>p [:Complex (Re ub) (Im lb), ub - Complex (Re ub) (Im lb):]" "p3 = p \\<^sub>p [:ub, Complex (Re lb) (Im ub) - ub:]" "p4 = p \\<^sub>p [:Complex (Re lb) (Im ub), lb - Complex (Re lb) (Im ub):]" subgoal unfolding p1_def by (subst (10) complex_surj[symmetric],auto simp add:minus_complex.code) subgoal unfolding p2_def by (subst (10) complex_surj[symmetric],auto) subgoal unfolding p3_def by (subst (10) complex_surj[symmetric],auto simp add:minus_complex.code) subgoal unfolding p4_def by (subst (10) complex_surj[symmetric],auto) done show ?thesis apply (subst proots_rectangle_code1) apply (unfold no_proots_line_code Let_def) apply (fold p1_def p2_def p3_def p4_def pI1_def pR1_def pI2_def pR2_def pI3_def pR3_def pI4_def pR4_def gc1_def gc2_def gc3_def gc4_def more_folds) apply (fold sms_def) by presburger qed subsection \Polynomial roots on the upper half-plane\ \ \Roots counted WITH multiplicity\ definition proots_upper ::"complex poly \ nat" where "proots_upper p= proots_count p {z. Im z>0}" \ \Roots counted WITHOUT multiplicity\ definition proots_upper_card::"complex poly \ nat" where "proots_upper_card p = card (proots_within p {x. Im x >0})" lemma Im_Ln_tendsto_at_top: "((\x. Im (Ln (Complex a x))) \ pi/2 ) at_top " proof (cases "a=0") case False define f where "f=(\x. if a>0 then arctan (x/a) else arctan (x/a) + pi)" define g where "g=(\x. Im (Ln (Complex a x)))" have "(f \ pi / 2) at_top" proof (cases "a>0") case True then have "(f \ pi / 2) at_top \ ((\x. arctan (x * inverse a)) \ pi / 2) at_top" unfolding f_def field_class.field_divide_inverse by auto also have "... \ (arctan \ pi / 2) at_top" apply (subst filterlim_at_top_linear_iff[of "inverse a" arctan 0 "nhds (pi/2)",simplified]) using True by auto also have "..." using tendsto_arctan_at_top . finally show ?thesis . next case False then have "(f \ pi / 2) at_top \ ((\x. arctan (x * inverse a) + pi) \ pi / 2) at_top" unfolding f_def field_class.field_divide_inverse by auto also have "... \ ((\x. arctan (x * inverse a)) \ - pi / 2) at_top" apply (subst tendsto_add_const_iff[of "-pi",symmetric]) by auto also have "... \ (arctan \ - pi / 2) at_bot" apply (subst filterlim_at_top_linear_iff[of "inverse a" arctan 0,simplified]) using False \a\0\ by auto also have "..." using tendsto_arctan_at_bot by simp finally show ?thesis . qed moreover have "\\<^sub>F x in at_top. f x = g x" unfolding f_def g_def using \a\0\ apply (subst Im_Ln_eq) subgoal for x using Complex_eq_0 by blast subgoal unfolding eventually_at_top_linorder by auto done ultimately show ?thesis using tendsto_cong[of f g at_top] unfolding g_def by auto next case True show ?thesis apply (rule tendsto_eventually) apply (rule eventually_at_top_linorderI[of 1]) using True by (subst Im_Ln_eq,auto simp add:Complex_eq_0) qed lemma Im_Ln_tendsto_at_bot: "((\x. Im (Ln (Complex a x))) \ - pi/2 ) at_bot " proof (cases "a=0") case False define f where "f=(\x. if a>0 then arctan (x/a) else arctan (x/a) - pi)" define g where "g=(\x. Im (Ln (Complex a x)))" have "(f \ - pi / 2) at_bot" proof (cases "a>0") case True then have "(f \ - pi / 2) at_bot \ ((\x. arctan (x * inverse a)) \ - pi / 2) at_bot" unfolding f_def field_class.field_divide_inverse by auto also have "... \ (arctan \ - pi / 2) at_bot" apply (subst filterlim_at_bot_linear_iff[of "inverse a" arctan 0,simplified]) using True by auto also have "..." using tendsto_arctan_at_bot by simp finally show ?thesis . next case False then have "(f \ - pi / 2) at_bot \ ((\x. arctan (x * inverse a) - pi) \ - pi / 2) at_bot" unfolding f_def field_class.field_divide_inverse by auto also have "... \ ((\x. arctan (x * inverse a)) \ pi / 2) at_bot" apply (subst tendsto_add_const_iff[of "pi",symmetric]) by auto also have "... \ (arctan \ pi / 2) at_top" apply (subst filterlim_at_bot_linear_iff[of "inverse a" arctan 0,simplified]) using False \a\0\ by auto also have "..." using tendsto_arctan_at_top by simp finally show ?thesis . qed moreover have "\\<^sub>F x in at_bot. f x = g x" unfolding f_def g_def using \a\0\ apply (subst Im_Ln_eq) subgoal for x using Complex_eq_0 by blast subgoal unfolding eventually_at_bot_linorder by (auto intro:exI[where x="-1"]) done ultimately show ?thesis using tendsto_cong[of f g at_bot] unfolding g_def by auto next case True show ?thesis apply (rule tendsto_eventually) apply (rule eventually_at_bot_linorderI[of "-1"]) using True by (subst Im_Ln_eq,auto simp add:Complex_eq_0) qed lemma Re_winding_number_tendsto_part_circlepath: shows "((\r. Re (winding_number (part_circlepath z0 r 0 pi ) a)) \ 1/2 ) at_top" proof (cases "Im z0\Im a") case True define g1 where "g1=(\r. part_circlepath z0 r 0 pi)" define g2 where "g2=(\r. part_circlepath z0 r pi (2*pi))" define f1 where "f1=(\r. Re (winding_number (g1 r ) a))" define f2 where "f2=(\r. Re (winding_number (g2 r) a))" have "(f2 \ 1/2 ) at_top" proof - define h1 where "h1 = (\r. Im (Ln (Complex ( Im a-Im z0) (Re z0 - Re a + r))))" define h2 where "h2= (\r. Im (Ln (Complex ( Im a -Im z0) (Re z0 - Re a - r))))" have "\\<^sub>F x in at_top. f2 x = (h1 x - h2 x) / (2 * pi)" proof (rule eventually_at_top_linorderI[of "cmod (a-z0) + 1"]) fix r assume asm:"r \ cmod (a - z0) + 1" have "Im p \ Im a" when "p\path_image (g2 r)" for p proof - obtain t where p_def:"p=z0 + of_real r * exp (\ * of_real t)" and "pi\t" "t\2*pi" using \p\path_image (g2 r)\ unfolding g2_def path_image_part_circlepath[of pi "2*pi",simplified] by auto then have "Im p=Im z0 + sin t * r" by (auto simp add:Im_exp) also have "... \ Im z0" proof - have "sin t\0" using \pi\t\ \t\2*pi\ sin_le_zero by fastforce moreover have "r\0" using asm by (metis add.inverse_inverse add.left_neutral add_uminus_conv_diff diff_ge_0_iff_ge norm_ge_zero order_trans zero_le_one) ultimately have "sin t * r\0" using mult_le_0_iff by blast then show ?thesis by auto qed also have "... \ Im a" using True . finally show ?thesis . qed moreover have "valid_path (g2 r)" unfolding g2_def by auto moreover have "a \ path_image (g2 r)" unfolding g2_def apply (rule not_on_circlepathI) using asm by auto moreover have [symmetric]:"Im (Ln (\ * pathfinish (g2 r) - \ * a)) = h1 r" unfolding h1_def g2_def apply (simp only:pathfinish_pathstart_partcirclepath_simps) apply (subst (4 10) complex_eq) by (auto simp add:algebra_simps Complex_eq) moreover have [symmetric]:"Im (Ln (\ * pathstart (g2 r) - \ * a)) = h2 r" unfolding h2_def g2_def apply (simp only:pathfinish_pathstart_partcirclepath_simps) apply (subst (4 10) complex_eq) by (auto simp add:algebra_simps Complex_eq) ultimately show "f2 r = (h1 r - h2 r) / (2 * pi)" unfolding f2_def apply (subst Re_winding_number_half_lower) by (auto simp add:exp_Euler algebra_simps) qed moreover have "((\x. (h1 x - h2 x) / (2 * pi)) \ 1/2 ) at_top" proof - have "(h1 \ pi/2) at_top" unfolding h1_def apply (subst filterlim_at_top_linear_iff[of 1 _ "Re a - Re z0" ,simplified,symmetric]) using Im_Ln_tendsto_at_top by (simp del:Complex_eq) moreover have "(h2 \ - pi/2) at_top" unfolding h2_def apply (subst filterlim_at_bot_linear_iff[of "- 1" _ "- Re a + Re z0" ,simplified,symmetric]) using Im_Ln_tendsto_at_bot by (simp del:Complex_eq) ultimately have "((\x. h1 x- h2 x) \ pi) at_top" by (auto intro: tendsto_eq_intros) then show ?thesis by (auto intro: tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed moreover have "\\<^sub>F r in at_top. f2 r = 1 - f1 r" proof (rule eventually_at_top_linorderI[of "cmod (a-z0) + 1"]) fix r assume asm:"r \ cmod (a - z0) + 1" have "f1 r + f2 r = Re(winding_number (g1 r +++ g2 r) a)" unfolding f1_def f2_def g1_def g2_def apply (subst winding_number_join) using asm by (auto intro!:not_on_circlepathI) also have "... = Re(winding_number (circlepath z0 r) a)" proof - have "g1 r +++ g2 r = circlepath z0 r" unfolding circlepath_def g1_def g2_def joinpaths_def part_circlepath_def linepath_def by (auto simp add:field_simps) then show ?thesis by auto qed also have "... = 1" proof - have "winding_number (circlepath z0 r) a = 1" apply (rule winding_number_circlepath) using asm by auto then show ?thesis by auto qed finally have "f1 r+f2 r=1" . then show " f2 r = 1 - f1 r" by auto qed ultimately have "((\r. 1 - f1 r) \ 1/2 ) at_top" using tendsto_cong[of f2 "\r. 1 - f1 r" at_top] by auto then have "(f1 \ 1/2 ) at_top" apply (rule_tac tendsto_minus_cancel) apply (subst tendsto_add_const_iff[of 1,symmetric]) by auto then show ?thesis unfolding f1_def g1_def by auto next case False define g where "g=(\r. part_circlepath z0 r 0 pi)" define f where "f=(\r. Re (winding_number (g r) a))" have "(f \ 1/2 ) at_top" proof - define h1 where "h1 = (\r. Im (Ln (Complex ( Im z0-Im a) (Re a - Re z0 + r))))" define h2 where "h2= (\r. Im (Ln (Complex ( Im z0 -Im a ) (Re a - Re z0 - r))))" have "\\<^sub>F x in at_top. f x = (h1 x - h2 x) / (2 * pi)" proof (rule eventually_at_top_linorderI[of "cmod (a-z0) + 1"]) fix r assume asm:"r \ cmod (a - z0) + 1" have "Im p \ Im a" when "p\path_image (g r)" for p proof - obtain t where p_def:"p=z0 + of_real r * exp (\ * of_real t)" and "0\t" "t\pi" using \p\path_image (g r)\ unfolding g_def path_image_part_circlepath[of 0 pi,simplified] by auto then have "Im p=Im z0 + sin t * r" by (auto simp add:Im_exp) moreover have "sin t * r\0" proof - have "sin t\0" using \0\t\ \t\pi\ sin_ge_zero by fastforce moreover have "r\0" using asm by (metis add.inverse_inverse add.left_neutral add_uminus_conv_diff diff_ge_0_iff_ge norm_ge_zero order_trans zero_le_one) ultimately have "sin t * r\0" by simp then show ?thesis by auto qed ultimately show ?thesis using False by auto qed moreover have "valid_path (g r)" unfolding g_def by auto moreover have "a \ path_image (g r)" unfolding g_def apply (rule not_on_circlepathI) using asm by auto moreover have [symmetric]:"Im (Ln (\ * a - \ * pathfinish (g r))) = h1 r" unfolding h1_def g_def apply (simp only:pathfinish_pathstart_partcirclepath_simps) apply (subst (4 9) complex_eq) by (auto simp add:algebra_simps Complex_eq) moreover have [symmetric]:"Im (Ln (\ * a - \ * pathstart (g r))) = h2 r" unfolding h2_def g_def apply (simp only:pathfinish_pathstart_partcirclepath_simps) apply (subst (4 9) complex_eq) by (auto simp add:algebra_simps Complex_eq) ultimately show "f r = (h1 r - h2 r) / (2 * pi)" unfolding f_def apply (subst Re_winding_number_half_upper) by (auto simp add:exp_Euler algebra_simps) qed moreover have "((\x. (h1 x - h2 x) / (2 * pi)) \ 1/2 ) at_top" proof - have "(h1 \ pi/2) at_top" unfolding h1_def apply (subst filterlim_at_top_linear_iff[of 1 _ "- Re a + Re z0" ,simplified,symmetric]) using Im_Ln_tendsto_at_top by (simp del:Complex_eq) moreover have "(h2 \ - pi/2) at_top" unfolding h2_def apply (subst filterlim_at_bot_linear_iff[of "- 1" _ "Re a - Re z0" ,simplified,symmetric]) using Im_Ln_tendsto_at_bot by (simp del:Complex_eq) ultimately have "((\x. h1 x- h2 x) \ pi) at_top" by (auto intro: tendsto_eq_intros) then show ?thesis by (auto intro: tendsto_eq_intros) qed ultimately show ?thesis by (auto dest:tendsto_cong) qed then show ?thesis unfolding f_def g_def by auto qed lemma not_image_at_top_poly_part_circlepath: assumes "degree p>0" shows "\\<^sub>F r in at_top. b\path_image (poly p o part_circlepath z0 r st tt)" proof - have "finite (proots (p-[:b:]))" apply (rule finite_proots) using assms by auto from finite_ball_include[OF this] obtain R::real where "R>0" and R_ball:"proots (p-[:b:]) \ ball z0 R" by auto show ?thesis proof (rule eventually_at_top_linorderI[of R]) fix r assume "r\R" show "b\path_image (poly p o part_circlepath z0 r st tt)" unfolding path_image_compose proof clarify fix x assume asm:"b = poly p x" "x \ path_image (part_circlepath z0 r st tt)" then have "x\proots (p-[:b:])" unfolding proots_def by auto then have "x\ball z0 r" using R_ball \r\R\ by auto then have "cmod (x- z0) < r" by (simp add: dist_commute dist_norm) moreover have "cmod (x - z0) = r" using asm(2) in_path_image_part_circlepath \R>0\ \r\R\ by auto ultimately show False by auto qed qed qed lemma not_image_poly_part_circlepath: assumes "degree p>0" shows "\r>0. b\path_image (poly p o part_circlepath z0 r st tt)" proof - have "finite (proots (p-[:b:]))" apply (rule finite_proots) using assms by auto from finite_ball_include[OF this] obtain r::real where "r>0" and r_ball:"proots (p-[:b:]) \ ball z0 r" by auto have "b\path_image (poly p o part_circlepath z0 r st tt)" unfolding path_image_compose proof clarify fix x assume asm:"b = poly p x" "x \ path_image (part_circlepath z0 r st tt)" then have "x\proots (p-[:b:])" unfolding proots_def by auto then have "x\ball z0 r" using r_ball by auto then have "cmod (x- z0) < r" by (simp add: dist_commute dist_norm) moreover have "cmod (x - z0) = r" using asm(2) in_path_image_part_circlepath \r>0\ by auto ultimately show False by auto qed then show ?thesis using \r>0\ by blast qed lemma Re_winding_number_poly_part_circlepath: assumes "degree p>0" shows "((\r. Re (winding_number (poly p o part_circlepath z0 r 0 pi) 0)) \ degree p/2 ) at_top" using assms proof (induct rule:poly_root_induct_alt) case 0 then show ?case by auto next case (no_proots p) then have False using Fundamental_Theorem_Algebra.fundamental_theorem_of_algebra constant_degree neq0_conv by blast then show ?case by auto next case (root a p) define g where "g = (\r. part_circlepath z0 r 0 pi)" define q where "q=[:- a, 1:] * p" define w where "w = (\r. winding_number (poly q \ g r) 0)" have ?case when "degree p=0" proof - obtain pc where pc_def:"p=[:pc:]" using \degree p = 0\ degree_eq_zeroE by blast then have "pc\0" using root(2) by auto have "\\<^sub>F r in at_top. Re (w r) = Re (winding_number (g r) a)" proof (rule eventually_at_top_linorderI[of "cmod (( pc * a) / pc - z0) + 1"]) fix r::real assume asm:"cmod ((pc * a) / pc - z0) + 1 \ r" have "w r = winding_number ((\x. pc*x - pc*a) \ (g r)) 0" unfolding w_def pc_def g_def q_def apply auto by (metis (no_types, opaque_lifting) add.right_neutral mult.commute mult_zero_right poly_0 poly_pCons uminus_add_conv_diff) also have "... = winding_number (g r) a " apply (subst winding_number_comp_linear[where b="-pc*a",simplified]) subgoal using \pc\0\ . subgoal unfolding g_def by auto subgoal unfolding g_def apply (rule not_on_circlepathI) using asm by auto subgoal using \pc\0\ by (auto simp add:field_simps) done finally have "w r = winding_number (g r) a " . then show "Re (w r) = Re (winding_number (g r) a)" by simp qed moreover have "((\r. Re (winding_number (g r) a)) \ 1/2) at_top" using Re_winding_number_tendsto_part_circlepath unfolding g_def by auto ultimately have "((\r. Re (w r)) \ 1/2) at_top" by (auto dest!:tendsto_cong) moreover have "degree ([:- a, 1:] * p) = 1" unfolding pc_def using \pc\0\ by auto ultimately show ?thesis unfolding w_def g_def comp_def q_def by simp qed moreover have ?case when "degree p>0" proof - have "\\<^sub>F r in at_top. 0 \ path_image (poly q \ g r)" unfolding g_def apply (rule not_image_at_top_poly_part_circlepath) unfolding q_def using root.prems by blast then have "\\<^sub>F r in at_top. Re (w r) = Re (winding_number (g r) a) + Re (winding_number (poly p \ g r) 0)" proof (rule eventually_mono) fix r assume asm:"0 \ path_image (poly q \ g r)" define cc where "cc= 1 / (of_real (2 * pi) * \)" define pf where "pf=(\w. deriv (poly p) w / poly p w)" define af where "af=(\w. 1/(w-a))" have "w r = cc * contour_integral (g r) (\w. deriv (poly q) w / poly q w)" unfolding w_def apply (subst winding_number_comp[of UNIV,simplified]) using asm unfolding g_def cc_def by auto also have "... = cc * contour_integral (g r) (\w. deriv (poly p) w / poly p w + 1/(w-a))" proof - have "contour_integral (g r) (\w. deriv (poly q) w / poly q w) = contour_integral (g r) (\w. deriv (poly p) w / poly p w + 1/(w-a))" proof (rule contour_integral_eq) fix x assume "x \ path_image (g r)" have "deriv (poly q) x = deriv (poly p) x * (x-a) + poly p x" proof - have "poly q = (\x. (x-a) * poly p x)" apply (rule ext) unfolding q_def by (auto simp add:algebra_simps) then show ?thesis apply simp apply (subst deriv_mult[of "\x. x- a" _ "poly p"]) by (auto intro:derivative_intros) qed moreover have "poly p x\0 \ x-a\0" proof (rule ccontr) assume "\ (poly p x \ 0 \ x - a \ 0)" then have "poly q x=0" unfolding q_def by auto then have "0\poly q ` path_image (g r)" using \x \ path_image (g r)\ by auto then show False using \0 \ path_image (poly q \ g r)\ unfolding path_image_compose by auto qed ultimately show "deriv (poly q) x / poly q x = deriv (poly p) x / poly p x + 1 / (x - a)" unfolding q_def by (auto simp add:field_simps) qed then show ?thesis by auto qed also have "... = cc * contour_integral (g r) (\w. deriv (poly p) w / poly p w) + cc * contour_integral (g r) (\w. 1/(w-a))" proof (subst contour_integral_add) have "continuous_on (path_image (g r)) (\w. deriv (poly p) w)" unfolding deriv_pderiv by (intro continuous_intros) moreover have "\w\path_image (g r). poly p w \ 0" using asm unfolding q_def path_image_compose by auto ultimately show "(\w. deriv (poly p) w / poly p w) contour_integrable_on g r" unfolding g_def by (auto intro!: contour_integrable_continuous_part_circlepath continuous_intros) show "(\w. 1 / (w - a)) contour_integrable_on g r" apply (rule contour_integrable_inversediff) subgoal unfolding g_def by auto subgoal using asm unfolding q_def path_image_compose by auto done qed (auto simp add:algebra_simps) also have "... = winding_number (g r) a + winding_number (poly p o g r) 0" proof - have "winding_number (poly p o g r) 0 = cc * contour_integral (g r) (\w. deriv (poly p) w / poly p w)" apply (subst winding_number_comp[of UNIV,simplified]) using \0 \ path_image (poly q \ g r)\ unfolding path_image_compose q_def g_def cc_def by auto moreover have "winding_number (g r) a = cc * contour_integral (g r) (\w. 1/(w-a))" apply (subst winding_number_valid_path) using \0 \ path_image (poly q \ g r)\ unfolding path_image_compose q_def g_def cc_def by auto ultimately show ?thesis by auto qed finally show "Re (w r) = Re (winding_number (g r) a) + Re (winding_number (poly p \ g r) 0)" by auto qed moreover have "((\r. Re (winding_number (g r) a) + Re (winding_number (poly p \ g r) 0)) \ degree q / 2) at_top" proof - have "((\r. Re (winding_number (g r) a)) \1 / 2) at_top" unfolding g_def by (rule Re_winding_number_tendsto_part_circlepath) moreover have "((\r. Re (winding_number (poly p \ g r) 0)) \ degree p / 2) at_top" unfolding g_def by (rule root(1)[OF that]) moreover have "degree q = degree p + 1" unfolding q_def apply (subst degree_mult_eq) using that by auto ultimately show ?thesis by (simp add: tendsto_add add_divide_distrib) qed ultimately have "((\r. Re (w r)) \ degree q/2) at_top" by (auto dest!:tendsto_cong) then show ?thesis unfolding w_def q_def g_def by blast qed ultimately show ?case by blast qed lemma Re_winding_number_poly_linepth: fixes pp::"complex poly" defines "g \ (\r. poly pp o linepath (-r) (of_real r))" assumes "lead_coeff pp=1" and no_real_zero:"\x\proots pp. Im x\0" shows "((\r. 2*Re (winding_number (g r) 0) + cindex_pathE (g r) 0 ) \ 0 ) at_top" proof - define p where "p=map_poly Re pp" define q where "q=map_poly Im pp" define f where "f=(\t. poly q t / poly p t)" have sgnx_top:"sgnx (poly p) at_top = 1" unfolding sgnx_poly_at_top sgn_pos_inf_def p_def using \lead_coeff pp=1\ by (subst lead_coeff_map_poly_nz,auto) have not_g_image:"0 \ path_image (g r)" for r proof (rule ccontr) assume "\ 0 \ path_image (g r)" then obtain x where "poly pp x=0" "x\closed_segment (- of_real r) (of_real r)" unfolding g_def path_image_compose of_real_linepath by auto then have "Im x=0" "x\proots pp" using closed_segment_imp_Re_Im(2) unfolding proots_def by fastforce+ then show False using \\x\proots pp. Im x\0\ by auto qed have arctan_f_tendsto:"((\r. (arctan (f r) - arctan (f (-r))) / pi) \ 0) at_top" proof (cases "degree p>0") case True have "degree p>degree q" proof - have "degree p=degree pp" unfolding p_def using \lead_coeff pp=1\ by (auto intro:map_poly_degree_eq) moreover then have "degree qlead_coeff pp=1\ True by (auto intro!:map_poly_degree_less) ultimately show ?thesis by auto qed then have "(f \ 0) at_infinity" unfolding f_def using poly_divide_tendsto_0_at_infinity by auto then have "(f \ 0) at_bot" "(f \ 0) at_top" by (auto elim!:filterlim_mono simp add:at_top_le_at_infinity at_bot_le_at_infinity) then have "((\r. arctan (f r))\ 0) at_top" "((\r. arctan (f (-r)))\ 0) at_top" apply - subgoal by (auto intro:tendsto_eq_intros) subgoal apply (subst tendsto_compose_filtermap[of _ uminus,unfolded comp_def]) by (auto intro:tendsto_eq_intros simp add:at_bot_mirror[symmetric]) done then show ?thesis by (auto intro:tendsto_eq_intros) next case False obtain c where "f=(\r. c)" proof - have "degree p=0" using False by auto moreover have "degree q\degree p" proof - have "degree p=degree pp" unfolding p_def using \lead_coeff pp=1\ by (auto intro:map_poly_degree_eq) moreover have "degree q\degree pp" unfolding q_def by simp ultimately show ?thesis by auto qed ultimately have "degree q=0" by simp then obtain pa qa where "p=[:pa:]" "q=[:qa:]" using \degree p=0\ by (meson degree_eq_zeroE) then show ?thesis using that unfolding f_def by auto qed then show ?thesis by auto qed have [simp]:"valid_path (g r)" "path (g r)" "finite_ReZ_segments (g r) 0" for r proof - show "valid_path (g r)" unfolding g_def apply (rule valid_path_compose_holomorphic[where S=UNIV]) by (auto simp add:of_real_linepath) then show "path (g r)" using valid_path_imp_path by auto show "finite_ReZ_segments (g r) 0" unfolding g_def of_real_linepath using finite_ReZ_segments_poly_linepath by simp qed have g_f_eq:"Im (g r t) / Re (g r t) = (f o (\x. 2*r*x - r)) t" for r t proof - have "Im (g r t) / Re (g r t) = Im ((poly pp o of_real o (\x. 2*r*x - r)) t) / Re ((poly pp o of_real o (\x. 2*r*x - r)) t)" unfolding g_def linepath_def comp_def by (auto simp add:algebra_simps) also have "... = (f o (\x. 2*r*x - r)) t" unfolding comp_def by (simp only:Im_poly_of_real diff_0_right Re_poly_of_real f_def q_def p_def) finally show ?thesis . qed have ?thesis when "proots p={}" proof - have "\\<^sub>Fr in at_top. 2 * Re (winding_number (g r) 0) + cindex_pathE (g r) 0 = (arctan (f r) - arctan (f (-r))) / pi" proof (rule eventually_at_top_linorderI[of 1]) fix r::real assume "r\1" have image_pos:"\p\path_image (g r). 0 (\p\path_image (g r). 0 < Re p)" then obtain t where "poly p t\0" unfolding g_def path_image_compose of_real_linepath p_def using Re_poly_of_real apply (simp add:closed_segment_def) by (metis not_less of_real_def real_vector.scale_scale scaleR_left_diff_distrib) moreover have False when "poly p t<0" proof - have "sgnx (poly p) (at_right t) = -1" using sgnx_poly_nz that by auto then obtain x where "x>t" "poly p x = 0" using sgnx_at_top_IVT[of p t] sgnx_top by auto then have "x\proots p" unfolding proots_def by auto then show False using \proots p={}\ by auto qed moreover have False when "poly p t=0" using \proots p={}\ that unfolding proots_def by auto ultimately show False by linarith qed have "Re (winding_number (g r) 0) = (Im (Ln (pathfinish (g r))) - Im (Ln (pathstart (g r)))) / (2 * pi)" apply (rule Re_winding_number_half_right[of "g r" 0,simplified]) subgoal using image_pos by auto subgoal by (auto simp add:not_g_image) done also have "... = (arctan (f r) - arctan (f (-r)))/(2*pi)" proof - have "Im (Ln (pathfinish (g r))) = arctan (f r)" proof - have "Re (pathfinish (g r)) > 0" by (auto intro: image_pos[rule_format]) then have "Im (Ln (pathfinish (g r))) = arctan (Im (pathfinish (g r)) / Re (pathfinish (g r)))" by (subst Im_Ln_eq,auto) also have "... = arctan (f r)" unfolding path_defs by (subst g_f_eq,auto) finally show ?thesis . qed moreover have "Im (Ln (pathstart (g r))) = arctan (f (-r))" proof - have "Re (pathstart (g r)) > 0" by (auto intro: image_pos[rule_format]) then have "Im (Ln (pathstart (g r))) = arctan (Im (pathstart (g r)) / Re (pathstart (g r)))" by (subst Im_Ln_eq,auto) also have "... = arctan (f (-r))" unfolding path_defs by (subst g_f_eq,auto) finally show ?thesis . qed ultimately show ?thesis by auto qed finally have "Re (winding_number (g r) 0) = (arctan (f r) - arctan (f (-r)))/(2*pi)" . moreover have "cindex_pathE (g r) 0 = 0" proof - have "cindex_pathE (g r) 0 = cindex_pathE (poly pp o of_real o (\x. 2*r*x - r)) 0" unfolding g_def linepath_def comp_def by (auto simp add:algebra_simps) also have "... = cindexE 0 1 (f o (\x. 2*r*x - r)) " unfolding cindex_pathE_def comp_def by (simp only:Im_poly_of_real diff_0_right Re_poly_of_real f_def q_def p_def) also have "... = cindexE (-r) r f" apply (subst cindexE_linear_comp[of "2*r" 0 1 _ "-r",simplified]) using \r\1\ by auto also have "... = 0" proof - have "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" when "x\{-r..r}" for x proof - have "poly p x\0" using \proots p={}\ unfolding proots_def by auto then show "jumpF f (at_left x) =0" "jumpF f (at_right x) = 0" unfolding f_def by (auto intro!: jumpF_not_infinity continuous_intros) qed then show ?thesis unfolding cindexE_def by auto qed finally show ?thesis . qed ultimately show "2 * Re (winding_number (g r) 0) + cindex_pathE (g r) 0 = (arctan (f r) - arctan (f (-r))) / pi" unfolding path_defs by (auto simp add:field_simps) qed with arctan_f_tendsto show ?thesis by (auto dest:tendsto_cong) qed moreover have ?thesis when "proots p\{}" proof - define max_r where "max_r=Max (proots p)" define min_r where "min_r=Min (proots p)" have "max_r \proots p" "min_r \proots p" "min_r\max_r" and min_max_bound:"\p\proots p. p\{min_r..max_r}" proof - have "p\0" proof - have "(0::real) \ 1" by simp then show ?thesis by (metis (full_types) \p \ map_poly Re pp\ assms(2) coeff_0 coeff_map_poly one_complex.simps(1) zero_complex.sel(1)) qed then have "finite (proots p)" by auto then show "max_r \proots p" "min_r \proots p" using Min_in Max_in that unfolding max_r_def min_r_def by fast+ then show "\p\proots p. p\{min_r..max_r}" using Min_le Max_ge \finite (proots p)\ unfolding max_r_def min_r_def by auto then show "min_r\max_r" using \max_r\proots p\ by auto qed have "\\<^sub>Fr in at_top. 2 * Re (winding_number (g r) 0) + cindex_pathE (g r) 0 = (arctan (f r) - arctan (f (-r))) / pi" proof (rule eventually_at_top_linorderI[of "max (norm max_r) (norm min_r) + 1"]) fix r assume r_asm:"max (norm max_r) (norm min_r) + 1 \ r" then have "r\0" "min_r>-r" "max_r{0..1}" "v\{0..1}" "u\v" unfolding u_def v_def using r_asm \min_r\max_r\ by (auto simp add:field_simps) define g1 where "g1=subpath 0 u (g r)" define g2 where "g2=subpath u v (g r)" define g3 where "g3=subpath v 1 (g r)" have "path g1" "path g2" "path g3" "valid_path g1" "valid_path g2" "valid_path g3" unfolding g1_def g2_def g3_def using uv by (auto intro!:path_subpath valid_path_subpath) define wc_add where "wc_add = (\g. 2*Re (winding_number g 0) + cindex_pathE g 0)" have "wc_add (g r) = wc_add g1 + wc_add g2 + wc_add g3" proof - have "winding_number (g r) 0 = winding_number g1 0 + winding_number g2 0 + winding_number g3 0" unfolding g1_def g2_def g3_def using \u\{0..1}\ \v\{0..1}\ not_g_image by (subst winding_number_subpath_combine,simp_all)+ moreover have "cindex_pathE (g r) 0 = cindex_pathE g1 0 + cindex_pathE g2 0 + cindex_pathE g3 0" unfolding g1_def g2_def g3_def using \u\{0..1}\ \v\{0..1}\ \u\v\ not_g_image by (subst cindex_pathE_subpath_combine,simp_all)+ ultimately show ?thesis unfolding wc_add_def by auto qed moreover have "wc_add g2=0" proof - have "2 * Re (winding_number g2 0) = - cindex_pathE g2 0" unfolding g2_def apply (rule winding_number_cindex_pathE_aux) subgoal using uv by (auto intro:finite_ReZ_segments_subpath) subgoal using uv by (auto intro:valid_path_subpath) subgoal using Path_Connected.path_image_subpath_subset \\r. path (g r)\ not_g_image uv by blast subgoal unfolding subpath_def v_def g_def linepath_def using r_asm \max_r \proots p\ by (auto simp add:field_simps Re_poly_of_real p_def) subgoal unfolding subpath_def u_def g_def linepath_def using r_asm \min_r \proots p\ by (auto simp add:field_simps Re_poly_of_real p_def) done then show ?thesis unfolding wc_add_def by auto qed moreover have "wc_add g1=- arctan (f (-r)) / pi" proof - have g1_pq: "Re (g1 t) = poly p (min_r*t+r*t-r)" "Im (g1 t) = poly q (min_r*t+r*t-r)" "Im (g1 t)/Re (g1 t) = (f o (\x. (min_r+r)*x - r)) t" for t proof - have "g1 t = poly pp (of_real (min_r*t+r*t-r))" using \r\0\ unfolding g1_def g_def linepath_def subpath_def u_def p_def by (auto simp add:field_simps) then show "Re (g1 t) = poly p (min_r*t+r*t-r)" "Im (g1 t) = poly q (min_r*t+r*t-r)" unfolding p_def q_def by (simp only:Re_poly_of_real Im_poly_of_real)+ then show "Im (g1 t)/Re (g1 t) = (f o (\x. (min_r+r)*x - r)) t" unfolding f_def by (auto simp add:algebra_simps) qed have "Re(g1 1)=0" using \r\0\ Re_poly_of_real \min_r\proots p\ unfolding g1_def subpath_def u_def g_def linepath_def by (auto simp add:field_simps p_def) have "0 \ path_image g1" by (metis (full_types) path_image_subpath_subset \\r. path (g r)\ atLeastAtMost_iff g1_def le_less not_g_image subsetCE uv(1) zero_le_one) have wc_add_pos:"wc_add h = - arctan (poly q (- r) / poly p (-r)) / pi" when Re_pos:"\x\{0..<1}. 0 < (Re \ h) x" and hp:"\t. Re (h t) = c*poly p (min_r*t+r*t-r)" and hq:"\t. Im (h t) = c*poly q (min_r*t+r*t-r)" and [simp]:"c\0" (*and hpq:"\t. Im (h t)/Re (h t) = (f o (\x. (min_r+r)*x - r)) t"*) and "Re (h 1) = 0" and "valid_path h" and h_img:"0 \ path_image h" for h c proof - define f where "f=(\t. c*poly q t / (c*poly p t))" define farg where "farg= (if 0 < Im (h 1) then pi / 2 else - pi / 2)" have "Re (winding_number h 0) = (Im (Ln (pathfinish h)) - Im (Ln (pathstart h))) / (2 * pi)" apply (rule Re_winding_number_half_right[of h 0,simplified]) subgoal using that \Re (h 1) = 0\ unfolding path_image_def by (auto simp add:le_less) subgoal using \valid_path h\ . subgoal using h_img . done also have "... = (farg - arctan (f (-r))) / (2 * pi)" proof - have "Im (Ln (pathfinish h)) = farg" using \Re(h 1)=0\ unfolding farg_def path_defs apply (subst Im_Ln_eq) subgoal using h_img unfolding path_defs by fastforce subgoal by simp done moreover have "Im (Ln (pathstart h)) = arctan (f (-r))" proof - have "pathstart h \ 0" using h_img by (metis pathstart_in_path_image) then have "Im (Ln (pathstart h)) = arctan (Im (pathstart h) / Re (pathstart h))" using Re_pos[rule_format,of 0] by (simp add: Im_Ln_eq path_defs) also have "... = arctan (f (-r))" unfolding f_def path_defs hp[rule_format] hq[rule_format] by simp finally show ?thesis . qed ultimately show ?thesis by auto qed finally have "Re (winding_number h 0) = (farg - arctan (f (-r))) / (2 * pi)" . moreover have "cindex_pathE h 0 = (-farg/pi)" proof - have "cindex_pathE h 0 = cindexE 0 1 (f \ (\x. (min_r + r) * x - r))" unfolding cindex_pathE_def using \c\0\ by (auto simp add:hp hq f_def comp_def algebra_simps) also have "... = cindexE (-r) min_r f" apply (subst cindexE_linear_comp[where b="-r",simplified]) using r_asm by auto also have "... = - jumpF f (at_left min_r)" proof - define right where "right = {x. jumpF f (at_right x) \ 0 \ - r \ x \ x < min_r}" define left where "left = {x. jumpF f (at_left x) \ 0 \ - r < x \ x \ min_r}" have *:"jumpF f (at_right x) =0" "jumpF f (at_left x) =0" when "x\{-r..min_r" using min_max_bound[rule_format,of x] that by auto then show False using \x\{-r.. by auto qed then show "jumpF f (at_right x) =0" "jumpF f (at_left x) =0" unfolding f_def by (auto intro!:jumpF_not_infinity continuous_intros) qed then have "right = {}" unfolding right_def by force moreover have "left = (if jumpF f (at_left min_r) = 0 then {} else {min_r})" unfolding left_def le_less using * r_asm by force ultimately show ?thesis unfolding cindexE_def by (fold left_def right_def,auto) qed also have "... = (-farg/pi)" proof - have p_pos:"c*poly p x > 0" when "x \ {- r<..t. min_r*t+r*t-r)" have "(x+r)/(min_r+r) \ {0..<1}" using that r_asm by (auto simp add:field_simps) then have "0 < c*poly p (hh ((x+r)/(min_r+r)))" apply (drule_tac Re_pos[rule_format]) unfolding comp_def hp[rule_format] hq[rule_format] hh_def . moreover have "hh ((x+r)/(min_r+r)) = x" unfolding hh_def using \min_r>-r\ apply (auto simp add:divide_simps) by (auto simp add:algebra_simps) ultimately show ?thesis by simp qed have "c*poly q min_r \0" using no_real_zero \c\0\ by (metis Im_complex_of_real UNIV_I \min_r \ proots p\ cpoly_of_decompose mult_eq_0_iff p_def poly_cpoly_of_real_iff proots_within_iff q_def) moreover have ?thesis when "c*poly q min_r > 0" proof - have "0 < Im (h 1)" unfolding hq[rule_format] hp[rule_format] using that by auto moreover have "jumpF f (at_left min_r) = 1/2" proof - have "((\t. c*poly p t) has_sgnx 1) (at_left min_r)" unfolding has_sgnx_def apply (rule eventually_at_leftI[of "-r"]) using p_pos \min_r>-r\ by auto then have "filterlim f at_top (at_left min_r)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ "c*poly q min_r"]) using that \min_r\proots p\ by (auto intro!:tendsto_eq_intros) then show ?thesis unfolding jumpF_def by auto qed ultimately show ?thesis unfolding farg_def by auto qed moreover have ?thesis when "c*poly q min_r < 0" proof - have "0 > Im (h 1)" unfolding hq[rule_format] hp[rule_format] using that by auto moreover have "jumpF f (at_left min_r) = - 1/2" proof - have "((\t. c*poly p t) has_sgnx 1) (at_left min_r)" unfolding has_sgnx_def apply (rule eventually_at_leftI[of "-r"]) using p_pos \min_r>-r\ by auto then have "filterlim f at_bot (at_left min_r)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ "c*poly q min_r"]) using that \min_r\proots p\ by (auto intro!:tendsto_eq_intros) then show ?thesis unfolding jumpF_def by auto qed ultimately show ?thesis unfolding farg_def by auto qed ultimately show ?thesis by linarith qed finally show ?thesis . qed ultimately show ?thesis unfolding wc_add_def f_def by (auto simp add:field_simps) qed have "\x\{0..<1}. (Re \ g1) x \ 0" proof (rule ccontr) assume "\ (\x\{0..<1}. (Re \ g1) x \ 0)" then obtain t where t_def:"Re (g1 t) =0" "t\{0..<1}" unfolding path_image_def by fastforce define m where "m=min_r*t+r*t-r" have "poly p m=0" proof - have "Re (g1 t) = Re (poly pp (of_real m))" unfolding m_def g1_def g_def linepath_def subpath_def u_def using \r\0\ by (auto simp add:field_simps) then show ?thesis using t_def unfolding Re_poly_of_real p_def by auto qed moreover have "m0" using r_asm by simp then have "(min_r + r)*(t-1)<0" using \t\{0..<1}\ by (simp add: mult_pos_neg) then show ?thesis unfolding m_def by (auto simp add:algebra_simps) qed ultimately show False using min_max_bound unfolding proots_def by auto qed then have "(\x\{0..<1}. 0 < (Re \ g1) x) \ (\x\{0..<1}. (Re \ g1) x < 0)" apply (elim continuous_on_neq_split) using \path g1\ unfolding path_def by (auto intro!:continuous_intros elim:continuous_on_subset) moreover have ?thesis when "\x\{0..<1}. (Re \ g1) x < 0" proof - have "wc_add (uminus o g1) = - arctan (f (- r)) / pi" unfolding f_def apply (rule wc_add_pos[of _ "-1"]) using g1_pq that \min_r \proots p\ \valid_path g1\ \0 \ path_image g1\ by (auto simp add:path_image_compose) moreover have "wc_add (uminus \ g1) = wc_add g1" unfolding wc_add_def cindex_pathE_def apply (subst winding_number_uminus_comp) using \valid_path g1\ \0 \ path_image g1\ by auto ultimately show ?thesis by auto qed moreover have ?thesis when "\x\{0..<1}. (Re \ g1) x > 0" unfolding f_def apply (rule wc_add_pos[of _ "1"]) using g1_pq that \min_r \proots p\ \valid_path g1\ \0 \ path_image g1\ by (auto simp add:path_image_compose) ultimately show ?thesis by blast qed moreover have "wc_add g3 = arctan (f r) / pi" proof - have g3_pq: "Re (g3 t) = poly p ((r-max_r)*t + max_r)" "Im (g3 t) = poly q ((r-max_r)*t + max_r)" "Im (g3 t)/Re (g3 t) = (f o (\x. (r-max_r)*x + max_r)) t" for t proof - have "g3 t = poly pp (of_real ((r-max_r)*t + max_r))" using \r\0\ \max_r unfolding g3_def g_def linepath_def subpath_def v_def p_def by (auto simp add:algebra_simps) then show "Re (g3 t) = poly p ((r-max_r)*t + max_r)" "Im (g3 t) = poly q ((r-max_r)*t + max_r)" unfolding p_def q_def by (simp only:Re_poly_of_real Im_poly_of_real)+ then show "Im (g3 t)/Re (g3 t) = (f o (\x. (r-max_r)*x + max_r)) t" unfolding f_def by (auto simp add:algebra_simps) qed have "Re(g3 0)=0" using \r\0\ Re_poly_of_real \max_r\proots p\ unfolding g3_def subpath_def v_def g_def linepath_def by (auto simp add:field_simps p_def) have "0 \ path_image g3" proof - have "(1::real) \ {0..1}" by auto then show ?thesis using Path_Connected.path_image_subpath_subset \\r. path (g r)\ g3_def not_g_image uv(2) by blast qed have wc_add_pos:"wc_add h = arctan (poly q r / poly p r) / pi" when Re_pos:"\x\{0<..1}. 0 < (Re \ h) x" and hp:"\t. Re (h t) = c*poly p ((r-max_r)*t + max_r)" and hq:"\t. Im (h t) = c*poly q ((r-max_r)*t + max_r)" and [simp]:"c\0" (*and hpq:"\t. Im (h t)/Re (h t) = (f o (\x. (min_r+r)*x - r)) t"*) and "Re (h 0) = 0" and "valid_path h" and h_img:"0 \ path_image h" for h c proof - define f where "f=(\t. c*poly q t / (c*poly p t))" define farg where "farg= (if 0 < Im (h 0) then pi / 2 else - pi / 2)" have "Re (winding_number h 0) = (Im (Ln (pathfinish h)) - Im (Ln (pathstart h))) / (2 * pi)" apply (rule Re_winding_number_half_right[of h 0,simplified]) subgoal using that \Re (h 0) = 0\ unfolding path_image_def by (auto simp add:le_less) subgoal using \valid_path h\ . subgoal using h_img . done also have "... = (arctan (f r) - farg) / (2 * pi)" proof - have "Im (Ln (pathstart h)) = farg" using \Re(h 0)=0\ unfolding farg_def path_defs apply (subst Im_Ln_eq) subgoal using h_img unfolding path_defs by fastforce subgoal by simp done moreover have "Im (Ln (pathfinish h)) = arctan (f r)" proof - have "pathfinish h \ 0" using h_img by (metis pathfinish_in_path_image) then have "Im (Ln (pathfinish h)) = arctan (Im (pathfinish h) / Re (pathfinish h))" using Re_pos[rule_format,of 1] by (simp add: Im_Ln_eq path_defs) also have "... = arctan (f r)" unfolding f_def path_defs hp[rule_format] hq[rule_format] by simp finally show ?thesis . qed ultimately show ?thesis by auto qed finally have "Re (winding_number h 0) = (arctan (f r) - farg) / (2 * pi)" . moreover have "cindex_pathE h 0 = farg/pi" proof - have "cindex_pathE h 0 = cindexE 0 1 (f \ (\x. (r-max_r)*x + max_r))" unfolding cindex_pathE_def using \c\0\ by (auto simp add:hp hq f_def comp_def algebra_simps) also have "... = cindexE max_r r f" apply (subst cindexE_linear_comp) using r_asm by auto also have "... = jumpF f (at_right max_r)" proof - define right where "right = {x. jumpF f (at_right x) \ 0 \ max_r \ x \ x < r}" define left where "left = {x. jumpF f (at_left x) \ 0 \ max_r < x \ x \ r}" have *:"jumpF f (at_right x) =0" "jumpF f (at_left x) =0" when "x\{max_r<..r}" for x proof - have False when "poly p x =0" proof - have "x\max_r" using min_max_bound[rule_format,of x] that by auto then show False using \x\{max_r<..r}\ by auto qed then show "jumpF f (at_right x) =0" "jumpF f (at_left x) =0" unfolding f_def by (auto intro!:jumpF_not_infinity continuous_intros) qed then have "left = {}" unfolding left_def by force moreover have "right = (if jumpF f (at_right max_r) = 0 then {} else {max_r})" unfolding right_def le_less using * r_asm by force ultimately show ?thesis unfolding cindexE_def by (fold left_def right_def,auto) qed also have "... = farg/pi" proof - have p_pos:"c*poly p x > 0" when "x \ {max_r<..t. (r-max_r)*t + max_r)" have "(x-max_r)/(r-max_r) \ {0<..1}" using that r_asm by (auto simp add:field_simps) then have "0 < c*poly p (hh ((x-max_r)/(r-max_r)))" apply (drule_tac Re_pos[rule_format]) unfolding comp_def hp[rule_format] hq[rule_format] hh_def . moreover have "hh ((x-max_r)/(r-max_r)) = x" unfolding hh_def using \max_r by (auto simp add:divide_simps) ultimately show ?thesis by simp qed have "c*poly q max_r \0" using no_real_zero \c\0\ by (metis Im_complex_of_real UNIV_I \max_r \ proots p\ cpoly_of_decompose mult_eq_0_iff p_def poly_cpoly_of_real_iff proots_within_iff q_def) moreover have ?thesis when "c*poly q max_r > 0" proof - have "0 < Im (h 0)" unfolding hq[rule_format] hp[rule_format] using that by auto moreover have "jumpF f (at_right max_r) = 1/2" proof - have "((\t. c*poly p t) has_sgnx 1) (at_right max_r)" unfolding has_sgnx_def apply (rule eventually_at_rightI[of _ "r"]) using p_pos \max_r by auto then have "filterlim f at_top (at_right max_r)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ "c*poly q max_r"]) using that \max_r\proots p\ by (auto intro!:tendsto_eq_intros) then show ?thesis unfolding jumpF_def by auto qed ultimately show ?thesis unfolding farg_def by auto qed moreover have ?thesis when "c*poly q max_r < 0" proof - have "0 > Im (h 0)" unfolding hq[rule_format] hp[rule_format] using that by auto moreover have "jumpF f (at_right max_r) = - 1/2" proof - have "((\t. c*poly p t) has_sgnx 1) (at_right max_r)" unfolding has_sgnx_def apply (rule eventually_at_rightI[of _ "r"]) using p_pos \max_r by auto then have "filterlim f at_bot (at_right max_r)" unfolding f_def apply (subst filterlim_divide_at_bot_at_top_iff[of _ "c*poly q max_r"]) using that \max_r\proots p\ by (auto intro!:tendsto_eq_intros) then show ?thesis unfolding jumpF_def by auto qed ultimately show ?thesis unfolding farg_def by auto qed ultimately show ?thesis by linarith qed finally show ?thesis . qed ultimately show ?thesis unfolding wc_add_def f_def by (auto simp add:field_simps) qed have "\x\{0<..1}. (Re \ g3) x \ 0" proof (rule ccontr) assume "\ (\x\{0<..1}. (Re \ g3) x \ 0)" then obtain t where t_def:"Re (g3 t) =0" "t\{0<..1}" unfolding path_image_def by fastforce define m where "m=(r-max_r)*t + max_r" have "poly p m=0" proof - have "Re (g3 t) = Re (poly pp (of_real m))" unfolding m_def g3_def g_def linepath_def subpath_def v_def using \r\0\ by (auto simp add:algebra_simps) then show ?thesis using t_def unfolding Re_poly_of_real p_def by auto qed moreover have "m>max_r" proof - have "r-max_r>0" using r_asm by simp then have "(r - max_r)*t>0" using \t\{0<..1}\ by (simp add: mult_pos_neg) then show ?thesis unfolding m_def by (auto simp add:algebra_simps) qed ultimately show False using min_max_bound unfolding proots_def by auto qed then have "(\x\{0<..1}. 0 < (Re \ g3) x) \ (\x\{0<..1}. (Re \ g3) x < 0)" apply (elim continuous_on_neq_split) using \path g3\ unfolding path_def by (auto intro!:continuous_intros elim:continuous_on_subset) moreover have ?thesis when "\x\{0<..1}. (Re \ g3) x < 0" proof - have "wc_add (uminus o g3) = arctan (f r) / pi" unfolding f_def apply (rule wc_add_pos[of _ "-1"]) using g3_pq that \max_r \proots p\ \valid_path g3\ \0 \ path_image g3\ by (auto simp add:path_image_compose) moreover have "wc_add (uminus \ g3) = wc_add g3" unfolding wc_add_def cindex_pathE_def apply (subst winding_number_uminus_comp) using \valid_path g3\ \0 \ path_image g3\ by auto ultimately show ?thesis by auto qed moreover have ?thesis when "\x\{0<..1}. (Re \ g3) x > 0" unfolding f_def apply (rule wc_add_pos[of _ "1"]) using g3_pq that \max_r \proots p\ \valid_path g3\ \0 \ path_image g3\ by (auto simp add:path_image_compose) ultimately show ?thesis by blast qed ultimately have "wc_add (g r) = (arctan (f r) - arctan (f (-r))) / pi " by (auto simp add:field_simps) then show "2 * Re (winding_number (g r) 0) + cindex_pathE (g r) 0 = (arctan (f r) - arctan (f (- r))) / pi" unfolding wc_add_def . qed with arctan_f_tendsto show ?thesis by (auto dest:tendsto_cong) qed ultimately show ?thesis by auto qed lemma proots_upper_cindex_eq: assumes "lead_coeff p=1" and no_real_roots:"\x\proots p. Im x\0" shows "proots_upper p = (degree p - cindex_poly_ubd (map_poly Im p) (map_poly Re p)) /2" proof (cases "degree p = 0") case True then obtain c where "p=[:c:]" using degree_eq_zeroE by blast then have p_def:"p=[:1:]" using \lead_coeff p=1\ by simp have "proots_count p {x. Im x>0} = 0" unfolding p_def proots_count_def by auto moreover have "cindex_poly_ubd (map_poly Im p) (map_poly Re p) = 0" apply (subst cindex_poly_ubd_code) unfolding p_def by (auto simp add:map_poly_pCons changes_R_smods_def changes_poly_neg_inf_def changes_poly_pos_inf_def) ultimately show ?thesis using True unfolding proots_upper_def by auto next case False then have "degree p>0" "p\0" by auto define w1 where "w1=(\r. Re (winding_number (poly p \ (\x. complex_of_real (linepath (- r) (of_real r) x))) 0))" define w2 where "w2=(\r. Re (winding_number (poly p \ part_circlepath 0 r 0 pi) 0))" define cp where "cp=(\r. cindex_pathE (poly p \ (\x. of_real (linepath (- r) (of_real r) x))) 0)" define ci where "ci=(\r. cindexE (-r) r (\x. poly (map_poly Im p) x/poly (map_poly Re p) x))" define cubd where "cubd =cindex_poly_ubd (map_poly Im p) (map_poly Re p)" obtain R where "proots p \ ball 0 R" and "R>0" using \p\0\ finite_ball_include[of "proots p" 0] by auto have "((\r. w1 r +w2 r+ cp r / 2 -ci r/2) \ real (degree p) / 2 - of_int cubd / 2) at_top" proof - have t1:"((\r. 2 * w1 r + cp r) \ 0) at_top" using Re_winding_number_poly_linepth[OF assms] unfolding w1_def cp_def by auto have t2:"(w2 \ real (degree p) / 2) at_top" using Re_winding_number_poly_part_circlepath[OF \degree p>0\,of 0] unfolding w2_def by auto have t3:"(ci \ of_int cubd) at_top" apply (rule tendsto_eventually) using cindex_poly_ubd_eventually[of "map_poly Im p" "map_poly Re p"] unfolding ci_def cubd_def by auto from tendsto_add[OF tendsto_add[OF tendsto_mult_left[OF t3,of "-1/2",simplified] tendsto_mult_left[OF t1,of "1/2",simplified]] t2] show ?thesis by (simp add:algebra_simps) qed moreover have "\\<^sub>F r in at_top. w1 r +w2 r+ cp r / 2 -ci r/2 = proots_count p {x. Im x>0}" proof (rule eventually_at_top_linorderI[of R]) fix r assume "r\R" then have r_ball:"proots p \ ball 0 r" and "r>0" using \R>0\ \proots p \ ball 0 R\ by auto define ll where "ll=linepath (- complex_of_real r) r" define rr where "rr=part_circlepath 0 r 0 pi" define lr where "lr = ll +++ rr" have img_ll:"path_image ll \ - proots p" and img_rr: "path_image rr \ - proots p" subgoal unfolding ll_def using \0 < r\ closed_segment_degen_complex(2) no_real_roots by auto subgoal unfolding rr_def using in_path_image_part_circlepath \0 < r\ r_ball by fastforce done have [simp]:"valid_path (poly p o ll)" "valid_path (poly p o rr)" "valid_path ll" "valid_path rr" "pathfinish rr=pathstart ll" "pathfinish ll = pathstart rr" proof - show "valid_path (poly p o ll)" "valid_path (poly p o rr)" unfolding ll_def rr_def by (auto intro:valid_path_compose_holomorphic) then show "valid_path ll" "valid_path rr" unfolding ll_def rr_def by auto show "pathfinish rr=pathstart ll" "pathfinish ll = pathstart rr" unfolding ll_def rr_def by auto qed have "proots_count p {x. Im x>0} = (\x\proots p. winding_number lr x * of_nat (order x p))" unfolding proots_count_def of_nat_sum proof (rule sum.mono_neutral_cong_left) show "finite (proots p)" "proots_within p {x. 0 < Im x} \ proots p" using \p\0\ by auto next have "winding_number lr x=0" when "x\proots p - proots_within p {x. 0 < Im x}" for x unfolding lr_def ll_def rr_def proof (eval_winding,simp_all) show *:"x \ closed_segment (- complex_of_real r) (complex_of_real r)" using img_ll that unfolding ll_def by auto show "x \ path_image (part_circlepath 0 r 0 pi)" using img_rr that unfolding rr_def by auto have xr_facts:"0 > Im x" "-r0" using that by auto moreover have "Im x\0" using no_real_roots that by blast ultimately show "0 > Im x" by auto next show "cmod xRe x\ < r" using abs_Re_le_cmod[of x] by argo then show "-rr>0\ unfolding cindex_pathE_linepath[OF *] ll_def by (auto simp add: mult_pos_neg) moreover have "cindex_pathE rr x=-1" unfolding rr_def using r_ball that by (auto intro!: cindex_pathE_circlepath_upper) ultimately show "-cindex_pathE (linepath (- of_real r) (of_real r)) x = cindex_pathE (part_circlepath 0 r 0 pi) x" unfolding ll_def rr_def by auto qed then show "\i\proots p - proots_within p {x. 0 < Im x}. winding_number lr i * of_nat (order i p) = 0" by auto next fix x assume x_asm:"x \ proots_within p {x. 0 < Im x}" have "winding_number lr x=1" unfolding lr_def ll_def rr_def proof (eval_winding,simp_all) show *:"x \ closed_segment (- complex_of_real r) (complex_of_real r)" using img_ll x_asm unfolding ll_def by auto show "x \ path_image (part_circlepath 0 r 0 pi)" using img_rr x_asm unfolding rr_def by auto have xr_facts:"0 < Im x" "-rRe x\ < r" using abs_Re_le_cmod[of x] by argo then show "-rr>0\ unfolding cindex_pathE_linepath[OF *] ll_def by (auto simp add: mult_less_0_iff) moreover have "cindex_pathE rr x=-1" unfolding rr_def using r_ball x_asm by (auto intro!: cindex_pathE_circlepath_upper) ultimately show "- of_real (cindex_pathE (linepath (- of_real r) (of_real r)) x) - of_real (cindex_pathE (part_circlepath 0 r 0 pi) x) = 2" unfolding ll_def rr_def by auto qed then show "of_nat (order x p) = winding_number lr x * of_nat (order x p)" by auto qed also have "... = 1/(2*pi*\) * contour_integral lr (\x. deriv (poly p) x / poly p x)" apply (subst argument_principle_poly[of p lr]) using \p\0\ img_ll img_rr unfolding lr_def ll_def rr_def by (auto simp add:path_image_join) also have "... = winding_number (poly p \ lr) 0" apply (subst winding_number_comp[of UNIV "poly p" lr 0]) using \p\0\ img_ll img_rr unfolding lr_def ll_def rr_def by (auto simp add:path_image_join path_image_compose) also have "... = Re (winding_number (poly p \ lr) 0)" proof - have "winding_number (poly p \ lr) 0 \ Ints" apply (rule integer_winding_number) using \p\0\ img_ll img_rr unfolding lr_def by (auto simp add:path_image_join path_image_compose path_compose_join pathstart_compose pathfinish_compose valid_path_imp_path) then show ?thesis by (simp add: complex_eqI complex_is_Int_iff) qed also have "... = Re (winding_number (poly p \ ll) 0) + Re (winding_number (poly p \ rr) 0)" unfolding lr_def path_compose_join using img_ll img_rr apply (subst winding_number_join) by (auto simp add:valid_path_imp_path path_image_compose pathstart_compose pathfinish_compose) also have "... = w1 r +w2 r" unfolding w1_def w2_def ll_def rr_def of_real_linepath by auto finally have "of_nat (proots_count p {x. 0 < Im x}) = complex_of_real (w1 r + w2 r)" . then have "proots_count p {x. 0 < Im x} = w1 r + w2 r" using of_real_eq_iff by fastforce moreover have "cp r = ci r" proof - define f where "f=(\x. Im (poly p (of_real x)) / Re (poly p x))" have "cp r = cindex_pathE (poly p \ (\x. 2*r*x - r)) 0" unfolding cp_def linepath_def by (auto simp add:algebra_simps) also have "... = cindexE 0 1 (f o (\x. 2*r*x - r))" unfolding cp_def ci_def cindex_pathE_def f_def comp_def by auto also have "... = cindexE (-r) r f" apply (subst cindexE_linear_comp[of "2*r" 0 1 f "-r",simplified]) using \r>0\ by auto also have "... = ci r" unfolding ci_def f_def Im_poly_of_real Re_poly_of_real by simp finally show ?thesis . qed ultimately show "w1 r + w2 r + cp r / 2 - ci r / 2 = real (proots_count p {x. 0 < Im x})" by auto qed ultimately have "((\r::real. real (proots_count p {x. 0 < Im x})) \ real (degree p) / 2 - of_int cubd / 2) at_top" by (auto dest: tendsto_cong) then show ?thesis apply (subst (asm) tendsto_const_iff) unfolding cubd_def proots_upper_def by auto qed lemma cindexE_roots_on_horizontal_border: fixes a::complex and s::real defines "g\linepath a (a + of_real s)" assumes pqr:"p = q * r" and r_monic:"lead_coeff r=1" and r_proots:"\x\proots r. Im x=Im a" shows "cindexE lb ub (\t. Im ((poly p \ g) t) / Re ((poly p \ g) t)) = cindexE lb ub (\t. Im ((poly q \ g) t) / Re ((poly q \ g) t))" using assms proof (induct r arbitrary:p rule:poly_root_induct_alt) case 0 then have False by (metis Im_complex_of_real UNIV_I imaginary_unit.simps(2) proots_within_0 zero_neq_one) then show ?case by simp next case (no_proots r) then obtain b where "b\0" "r=[:b:]" using fundamental_theorem_of_algebra_alt by blast then have "r=1" using \lead_coeff r = 1\ by simp with \p = q * r\ show ?case by simp next case (root b r) then have ?case when "s=0" using that(1) unfolding cindex_pathE_def by (simp add:cindexE_constI) moreover have ?case when "s\0" proof - define qrg where "qrg = poly (q*r) \ g" have "cindexE lb ub (\t. Im ((poly p \ g) t) / Re ((poly p \ g) t)) = cindexE lb ub (\t. Im (qrg t * (g t - b)) / Re (qrg t * (g t - b)))" unfolding qrg_def \p = q * ([:- b, 1:] * r)\ comp_def by (simp add:algebra_simps) also have "... = cindexE lb ub (\t. ((Re a + t * s - Re b )* Im (qrg t)) / ((Re a + t * s - Re b )* Re (qrg t)))" proof - have "Im b = Im a" using \\x\proots ([:- b, 1:] * r). Im x = Im a\ by auto then show ?thesis unfolding cindex_pathE_def g_def linepath_def by (simp add:algebra_simps) qed also have "... = cindexE lb ub (\t. Im (qrg t) / Re (qrg t))" proof (rule cindexE_cong[of "{t. Re a + t * s - Re b = 0}"]) show "finite {t. Re a + t * s - Re b = 0}" proof (cases "Re a= Re b") case True then have "{t. Re a + t * s - Re b = 0} = {0}" using \s\0\ by auto then show ?thesis by auto next case False then have "{t. Re a + t * s - Re b = 0} = {(Re b - Re a) / s}" using \s\0\ by (auto simp add:field_simps) then show ?thesis by auto qed next fix x assume asm:"x \ {t. Re a + t * s - Re b = 0}" define tt where "tt=Re a + x * s - Re b" have "tt\0" using asm unfolding tt_def by auto then show "tt * Im (qrg x) / (tt * Re (qrg x)) = Im (qrg x) / Re (qrg x)" by auto qed also have "... = cindexE lb ub (\t. Im ((poly q \ g) t) / Re ((poly q \ g) t))" unfolding qrg_def proof (rule root(1)) show "lead_coeff r = 1" by (metis lead_coeff_mult lead_coeff_pCons(1) mult_cancel_left2 one_poly_eq_simps(2) root.prems(2) zero_neq_one) qed (use root in simp_all) finally show ?thesis . qed ultimately show ?case by auto qed lemma poly_decompose_by_proots: fixes p ::"'a::idom poly" assumes "p\0" shows "\q r. p = q * r \ lead_coeff q=1 \ (\x\proots q. P x) \ (\x\proots r. \P x)" using assms proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) then show ?case apply (rule_tac x=1 in exI) apply (rule_tac x=p in exI) by (simp add:proots_def) next case (root a p) then obtain q r where pqr:"p = q * r" and leadq:"lead_coeff q=1" and qball:"\a\proots q. P a" and rball:"\x\proots r. \ P x" using mult_zero_right by blast have ?case when "P a" apply (rule_tac x="[:- a, 1:] * q" in exI) apply (rule_tac x=r in exI) using pqr qball rball that leadq unfolding lead_coeff_mult by (auto simp add:algebra_simps) moreover have ?case when "\ P a" apply (rule_tac x="q" in exI) apply (rule_tac x="[:- a, 1:] *r" in exI) using pqr qball rball that leadq unfolding lead_coeff_mult by (auto simp add:algebra_simps) ultimately show ?case by blast qed lemma proots_upper_cindex_eq': assumes "lead_coeff p=1" shows "proots_upper p = (degree p - proots_count p {x. Im x=0} - cindex_poly_ubd (map_poly Im p) (map_poly Re p)) /2" proof - have "p\0" using assms by auto from poly_decompose_by_proots[OF this,of "\x. Im x\0"] obtain q r where pqr:"p = q * r" and leadq:"lead_coeff q=1" and qball: "\x\proots q. Im x \0" and rball:"\x\proots r. Im x =0" by auto have "real_of_int (proots_upper p) = proots_upper q + proots_upper r" using \p\0\ unfolding proots_upper_def pqr by (auto simp add:proots_count_times) also have "... = proots_upper q" proof - have "proots_within r {z. 0 < Im z} = {}" using rball by auto then have "proots_upper r =0 " unfolding proots_upper_def proots_count_def by simp then show ?thesis by auto qed also have "... = (degree q - cindex_poly_ubd (map_poly Im q) (map_poly Re q)) / 2" by (rule proots_upper_cindex_eq[OF leadq qball]) also have "... = (degree p - proots_count p {x. Im x=0} - cindex_poly_ubd (map_poly Im p) (map_poly Re p)) /2" proof - have "degree q = degree p - proots_count p {x. Im x=0}" proof - have "degree p= degree q + degree r" unfolding pqr apply (rule degree_mult_eq) using \p \ 0\ pqr by auto moreover have "degree r = proots_count p {x. Im x=0}" unfolding degree_proots_count proots_count_def proof (rule sum.cong) fix x assume "x \ proots_within p {x. Im x = 0}" then have "Im x=0" by auto then have "order x q = 0" using qball order_0I by blast then show "order x r = order x p" using \p\0\ unfolding pqr by (simp add: order_mult) next show "proots r = proots_within p {x. Im x = 0}" unfolding pqr proots_within_times using qball rball by auto qed ultimately show ?thesis by auto qed moreover have "cindex_poly_ubd (map_poly Im q) (map_poly Re q) = cindex_poly_ubd (map_poly Im p) (map_poly Re p)" proof - define iq rq ip rp where "iq = map_poly Im q" and "rq=map_poly Re q" and "ip=map_poly Im p" and "rp = map_poly Re p" have "cindexE (- x) x (\x. poly iq x / poly rq x) = cindexE (- x) x (\x. poly ip x / poly rp x)" for x proof - have "lead_coeff r = 1" using pqr leadq \lead_coeff p=1\ by (simp add: coeff_degree_mult) then have "cindexE (- x) x (\t. Im (poly p (t *\<^sub>R 1)) / Re (poly p (t *\<^sub>R 1))) = cindexE (- x) x (\t. Im (poly q (t *\<^sub>R 1)) / Re (poly q (t *\<^sub>R 1)))" using cindexE_roots_on_horizontal_border[OF pqr,of 0 "-x" x 1 ,unfolded linepath_def comp_def,simplified] rball by simp then show ?thesis unfolding scaleR_conv_of_real iq_def ip_def rq_def rp_def by (simp add:Im_poly_of_real Re_poly_of_real) qed then have "\\<^sub>F r::real in at_top. real_of_int (cindex_poly_ubd iq rq) = cindex_poly_ubd ip rp" using eventually_conj[OF cindex_poly_ubd_eventually[of iq rq] cindex_poly_ubd_eventually[of ip rp]] by (elim eventually_mono,auto) then show ?thesis apply (fold iq_def rq_def ip_def rp_def) by simp qed ultimately show ?thesis by auto qed finally show ?thesis by simp qed (*If we know that the polynomial p is squarefree, we can cope with the case when there're roots on the border.*) lemma proots_within_upper_squarefree: assumes "rsquarefree p" shows "card (proots_within p {x. Im x >0}) = (let pp = smult (inverse (lead_coeff p)) p; pI = map_poly Im pp; pR = map_poly Re pp; g = gcd pR pI in nat ((degree p - changes_R_smods g (pderiv g) - changes_R_smods pR pI) div 2) )" proof - define pp where "pp = smult (inverse (lead_coeff p)) p" define pI where "pI = map_poly Im pp" define pR where "pR = map_poly Re pp" define g where "g = gcd pR pI" have "card (proots_within p {x. Im x >0}) = proots_upper p" unfolding proots_upper_def using card_proots_within_rsquarefree[OF assms] by auto also have "... = proots_upper pp" unfolding proots_upper_def pp_def apply (subst proots_count_smult) using assms by auto also have "... = (degree pp - proots_count pp {x. Im x = 0} - cindex_poly_ubd pI pR) div 2" proof - define rr where "rr = proots_count pp {x. Im x = 0}" define cpp where "cpp = cindex_poly_ubd pI pR" have *:"proots_upper pp = (degree pp - rr - cpp) / 2" apply (rule proots_upper_cindex_eq'[of pp,folded rr_def cpp_def pR_def pI_def]) unfolding pp_def using assms by auto also have "... = (degree pp - rr - cpp) div 2" proof (subst real_of_int_div) define tt where "tt=int (degree pp - rr) - cpp" have "real_of_int tt=2*proots_upper pp" by (simp add:*[folded tt_def]) then show "even tt" by (metis dvd_triv_left even_of_nat of_int_eq_iff of_int_of_nat_eq) qed simp finally show ?thesis unfolding rr_def cpp_def by simp qed also have "... = (degree pp - changes_R_smods g (pderiv g) - cindex_poly_ubd pI pR) div 2" proof - have "rsquarefree pp" using assms rsquarefree_smult_iff unfolding pp_def by (metis inverse_eq_imp_eq inverse_zero leading_coeff_neq_0 rsquarefree_0) from card_proots_within_rsquarefree[OF this] have "proots_count pp {x. Im x = 0} = card (proots_within pp {x. Im x = 0})" by simp also have "... = card (proots_within pp (unbounded_line 0 1))" proof - have "{x. Im x = 0} = unbounded_line 0 1" unfolding unbounded_line_def apply auto subgoal for x apply (rule_tac x="Re x" in exI) by (metis complex_is_Real_iff of_real_Re of_real_def) done then show ?thesis by simp qed also have "... = changes_R_smods g (pderiv g)" unfolding card_proots_unbounded_line[of 0 1 pp,simplified,folded pI_def pR_def] g_def by (auto simp add:Let_def sturm_R[symmetric]) finally have "proots_count pp {x. Im x = 0} = changes_R_smods g (pderiv g)" . moreover have "degree pp \ proots_count pp {x. Im x = 0}" by (metis \rsquarefree pp\ proots_count_leq_degree rsquarefree_0) ultimately show ?thesis by auto qed also have "... = (degree p - changes_R_smods g (pderiv g) - changes_R_smods pR pI) div 2" using cindex_poly_ubd_code unfolding pp_def by simp finally have "card (proots_within p {x. 0 < Im x}) = (degree p - changes_R_smods g (pderiv g) - changes_R_smods pR pI) div 2" . then show ?thesis unfolding Let_def apply (fold pp_def pR_def pI_def g_def) by (simp add: pp_def) qed lemma proots_upper_code1[code]: "proots_upper p = (if p \ 0 then (let pp=smult (inverse (lead_coeff p)) p; pI=map_poly Im pp; pR=map_poly Re pp; g = gcd pI pR in nat ((degree p - nat (changes_R_smods_ext g (pderiv g)) - changes_R_smods pR pI) div 2) ) else Code.abort (STR ''proots_upper fails when p=0.'') (\_. proots_upper p))" proof - define pp where "pp = smult (inverse (lead_coeff p)) p" define pI where "pI = map_poly Im pp" define pR where "pR=map_poly Re pp" define g where "g = gcd pI pR" have ?thesis when "p=0" using that by auto moreover have ?thesis when "p\0" proof - have "pp\0" unfolding pp_def using that by auto define rr where "rr=int (degree pp - proots_count pp {x. Im x = 0}) - cindex_poly_ubd pI pR" have "lead_coeff p\0" using \p\0\ by simp then have "proots_upper pp = rr / 2" unfolding rr_def apply (rule_tac proots_upper_cindex_eq'[of pp, folded pI_def pR_def]) unfolding pp_def lead_coeff_smult by auto then have "proots_upper pp = nat (rr div 2)" by linarith moreover have "rr = degree p - nat (changes_R_smods_ext g (pderiv g)) - changes_R_smods pR pI" proof - have "degree pp = degree p" unfolding pp_def by auto moreover have "proots_count pp {x. Im x = 0} = nat (changes_R_smods_ext g (pderiv g))" proof - have "{x. Im x = 0} = unbounded_line 0 1" unfolding unbounded_line_def by (simp add: complex_eq_iff) then show ?thesis using proots_unbounded_line[of 0 1 pp,simplified, folded pI_def pR_def] \pp\0\ by (auto simp:Let_def g_def gcd.commute) qed moreover have "cindex_poly_ubd pI pR = changes_R_smods pR pI" using cindex_poly_ubd_code by auto ultimately show ?thesis unfolding rr_def by auto qed moreover have "proots_upper pp = proots_upper p" unfolding pp_def proots_upper_def apply (subst proots_count_smult) using that by auto ultimately show ?thesis unfolding Let_def using that apply (fold pp_def pI_def pR_def g_def) by argo qed ultimately show ?thesis by blast qed lemma proots_upper_card_code[code]: "proots_upper_card p = (if p=0 then 0 else (let pf = p div (gcd p (pderiv p)); pp = smult (inverse (lead_coeff pf)) pf; pI = map_poly Im pp; pR = map_poly Re pp; g = gcd pR pI in nat ((degree pf - changes_R_smods g (pderiv g) - changes_R_smods pR pI) div 2) ))" proof (cases "p=0") case True then show ?thesis unfolding proots_upper_card_def using infinite_halfspace_Im_gt by auto next case False define pf pp pI pR g where "pf = p div (gcd p (pderiv p))" and "pp = smult (inverse (lead_coeff pf)) pf" and "pI = map_poly Im pp" and "pR = map_poly Re pp" and "g = gcd pR pI" have "proots_upper_card p = proots_upper_card pf" proof - have "proots_within p {x. 0 < Im x} = proots_within pf {x. 0 < Im x}" unfolding proots_within_def using poly_gcd_pderiv_iff[of p,folded pf_def] by auto then show ?thesis unfolding proots_upper_card_def by auto qed also have "... = nat ((degree pf - changes_R_smods g (pderiv g) - changes_R_smods pR pI) div 2)" using proots_within_upper_squarefree[OF rsquarefree_gcd_pderiv[OF \p\0\] ,unfolded Let_def,folded pf_def,folded pp_def pI_def pR_def g_def] unfolding proots_upper_card_def by blast finally show ?thesis unfolding Let_def apply (fold pf_def,fold pp_def pI_def pR_def g_def) using False by auto qed subsection \Polynomial roots on a general half-plane\ text \the number of roots of polynomial @{term p}, counted with multiplicity, on the left half plane of the vector @{term "b-a"}.\ definition proots_half ::"complex poly \ complex \ complex \ nat" where "proots_half p a b = proots_count p {w. Im ((w-a) / (b-a)) > 0}" lemma proots_half_empty: assumes "a=b" shows "proots_half p a b = 0" unfolding proots_half_def using assms by auto (*TODO: the proof can potentially simplified with some conformal properties.*) lemma proots_half_proots_upper: assumes "a\b" "p\0" shows "proots_half p a b= proots_upper (pcompose p [:a, (b-a):])" proof - define q where "q=[:a, (b - a):]" define f where "f=(\x. (b-a)*x+ a)" have "(\r\proots_within p {w. Im ((w-a) / (b-a)) > 0}. order r p) = (\r\proots_within (p \\<^sub>p q) {z. 0 < Im z}. order r (p \\<^sub>pq))" proof (rule sum.reindex_cong[of f]) have "inj f" using assms unfolding f_def inj_on_def by fastforce then show "inj_on f (proots_within (p \\<^sub>p q) {z. 0 < Im z})" by (elim inj_on_subset,auto) next show "proots_within p {w. Im ((w-a) / (b-a)) > 0} = f ` proots_within (p \\<^sub>p q) {z. 0 < Im z}" proof safe fix x assume x_asm:"x \ proots_within p {w. Im ((w-a) / (b-a)) > 0}" define xx where "xx=(x -a) / (b - a)" have "poly (p \\<^sub>p q) xx = 0" unfolding q_def xx_def poly_pcompose using assms x_asm by auto moreover have "Im xx > 0" unfolding xx_def using x_asm by auto ultimately have "xx \ proots_within (p \\<^sub>p q) {z. 0 < Im z}" by auto then show "x \ f ` proots_within (p \\<^sub>p q) {z. 0 < Im z}" apply (intro rev_image_eqI[of xx]) unfolding f_def xx_def using assms by auto next fix x assume "x \ proots_within (p \\<^sub>p q) {z. 0 < Im z}" then show "f x \ proots_within p {w. 0 < Im ((w-a) / (b - a))}" unfolding f_def q_def using assms apply (auto simp add:poly_pcompose) by (auto simp add:algebra_simps) qed next fix x assume "x \ proots_within (p \\<^sub>p q) {z. 0 < Im z}" show "order (f x) p = order x (p \\<^sub>p q)" using \p\0\ proof (induct p rule:poly_root_induct_alt) case 0 then show ?case by simp next case (no_proots p) have "order (f x) p = 0" apply (rule order_0I) using no_proots by auto moreover have "order x (p \\<^sub>p q) = 0" apply (rule order_0I) unfolding poly_pcompose q_def using no_proots by auto ultimately show ?case by auto next case (root c p) have "order (f x) ([:- c, 1:] * p) = order (f x) [:-c,1:] + order (f x) p" apply (subst order_mult) using root by auto also have "... = order x ([:- c, 1:] \\<^sub>p q) + order x (p \\<^sub>p q)" proof - have "order (f x) [:- c, 1:] = order x ([:- c, 1:] \\<^sub>p q)" proof (cases "f x=c") case True have "[:- c, 1:] \\<^sub>p q = smult (b-a) [:-x,1:]" using True unfolding q_def f_def pcompose_pCons by auto then have "order x ([:- c, 1:] \\<^sub>p q) = order x (smult (b-a) [:-x,1:])" by auto then have "order x ([:- c, 1:] \\<^sub>p q) = 1" apply (subst (asm) order_smult) using assms order_power_n_n[of _ 1,simplified] by auto moreover have "order (f x) [:- c, 1:] = 1" using True order_power_n_n[of _ 1,simplified] by auto ultimately show ?thesis by auto next case False have "order (f x) [:- c, 1:] = 0" apply (rule order_0I) using False unfolding f_def by auto moreover have "order x ([:- c, 1:] \\<^sub>p q) = 0" apply (rule order_0I) using False unfolding f_def q_def poly_pcompose by auto ultimately show ?thesis by auto qed moreover have "order (f x) p = order x (p \\<^sub>p q)" apply (rule root) using root by auto ultimately show ?thesis by auto qed also have "... = order x (([:- c, 1:] * p) \\<^sub>p q)" unfolding pcompose_mult apply (subst order_mult) subgoal unfolding q_def using assms(1) pcompose_eq_0 root.prems by fastforce by simp finally show ?case . qed qed then show ?thesis unfolding proots_half_def proots_upper_def proots_count_def q_def by auto qed lemma proots_half_code1[code]: "proots_half p a b = (if a\b then if p\0 then proots_upper (p \\<^sub>p [:a, b - a:]) else Code.abort (STR ''proots_half fails when p=0.'') (\_. proots_half p a b) else 0)" proof - have ?thesis when "a=b" using proots_half_empty that by auto moreover have ?thesis when "a\b" "p=0" using that by auto moreover have ?thesis when "a\b" "p\0" using proots_half_proots_upper[OF that] that by auto ultimately show ?thesis by auto qed subsection \Polynomial roots within a circle (open ball)\ \ \Roots counted WITH multiplicity\ definition proots_ball::"complex poly \ complex \ real \ nat" where "proots_ball p z0 r = proots_count p (ball z0 r)" \ \Roots counted WITHOUT multiplicity\ definition proots_ball_card ::"complex poly \ complex \ real \ nat" where "proots_ball_card p z0 r = card (proots_within p (ball z0 r))" lemma proots_ball_code1[code]: "proots_ball p z0 r = ( if r \ 0 then 0 else if p\0 then proots_upper (fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:]) else Code.abort (STR ''proots_ball fails when p=0.'') (\_. proots_ball p z0 r) )" proof (cases "p=0 \ r\0") case False have "proots_ball p z0 r = proots_count (p \\<^sub>p [:z0, of_real r:]) (ball 0 1)" unfolding proots_ball_def apply (rule proots_uball_eq[THEN arg_cong]) using False by auto also have "... = proots_upper (fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:])" unfolding proots_upper_def apply (rule proots_ball_plane_eq[THEN arg_cong]) using False pcompose_eq_0[of p "[:z0, of_real r:]"] by auto finally show ?thesis using False by auto qed (auto simp:proots_ball_def ball_empty) lemma proots_ball_card_code1[code]: "proots_ball_card p z0 r = ( if r \ 0 \ p=0 then 0 else proots_upper_card (fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:]) )" proof (cases "p=0 \ r\0") case True moreover have ?thesis when "r\0" proof - have "proots_within p (ball z0 r) = {}" by (simp add: ball_empty that) then show ?thesis unfolding proots_ball_card_def using that by auto qed moreover have ?thesis when "r>0" "p=0" unfolding proots_ball_card_def using that infinite_ball[of r z0] by auto ultimately show ?thesis by argo next case False then have "p\0" "r>0" by auto have "proots_ball_card p z0 r = card (proots_within (p \\<^sub>p [:z0, of_real r:]) (ball 0 1))" unfolding proots_ball_card_def by (rule proots_card_uball_eq[OF \r>0\, THEN arg_cong]) also have "... = proots_upper_card (fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:])" unfolding proots_upper_card_def apply (rule proots_card_ball_plane_eq[THEN arg_cong]) using False pcompose_eq_0[of p "[:z0, of_real r:]"] by auto finally show ?thesis using False by auto qed subsection \Polynomial roots on a circle (sphere)\ \ \Roots counted WITH multiplicity\ definition proots_sphere::"complex poly \ complex \ real \ nat" where "proots_sphere p z0 r = proots_count p (sphere z0 r)" \ \Roots counted WITHOUT multiplicity\ definition proots_sphere_card ::"complex poly \ complex \ real \ nat" where "proots_sphere_card p z0 r = card (proots_within p (sphere z0 r))" lemma proots_sphere_card_code1[code]: "proots_sphere_card p z0 r = ( if r=0 then (if poly p z0=0 then 1 else 0) else if r < 0 \ p=0 then 0 else (if poly p (z0-r) =0 then 1 else 0) + proots_unbounded_line_card (fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:]) 0 1 )" proof - have ?thesis when "r=0" proof - have "proots_within p {z0} = (if poly p z0 = 0 then {z0} else {})" by auto then show ?thesis unfolding proots_sphere_card_def using that by simp qed moreover have ?thesis when "r\0" "r < 0 \ p=0" proof - have ?thesis when "r<0" proof - have "proots_within p (sphere z0 r) = {}" by (auto simp add: ball_empty that) then show ?thesis unfolding proots_sphere_card_def using that by auto qed moreover have ?thesis when "r>0" "p=0" unfolding proots_sphere_card_def using that infinite_sphere[of r z0] by auto ultimately show ?thesis using that by argo qed moreover have ?thesis when "r>0" "p\0" proof - define pp where "pp = p \\<^sub>p [:z0, of_real r:]" define ppp where "ppp=fcompose pp [:\, - 1:] [:\, 1:]" have "pp\0" unfolding pp_def using that pcompose_eq_0 by fastforce have "proots_sphere_card p z0 r = card (proots_within pp (sphere 0 1))" unfolding proots_sphere_card_def pp_def by (rule proots_card_usphere_eq[OF \r>0\, THEN arg_cong]) also have "... = card (proots_within pp {-1} \ proots_within pp (sphere 0 1 - {-1}))" by (simp add: insert_absorb proots_within_union) also have "... = card (proots_within pp {-1}) + card (proots_within pp (sphere 0 1 - {-1}))" apply (rule card_Un_disjoint) using \pp\0\ by auto also have "... = card (proots_within pp {-1}) + card (proots_within ppp {x. 0 = Im x})" using proots_card_sphere_axis_eq[OF \pp\0\,folded ppp_def] by simp also have "... = (if poly p (z0-r) =0 then 1 else 0) + proots_unbounded_line_card ppp 0 1" proof - have "proots_within pp {-1} = (if poly p (z0-r) =0 then {-1} else {})" unfolding pp_def by (auto simp:poly_pcompose) then have "card (proots_within pp {-1}) = (if poly p (z0-r) =0 then 1 else 0)" by auto moreover have "{x. Im x = 0} = unbounded_line 0 1" unfolding unbounded_line_def apply auto by (metis complex_is_Real_iff of_real_Re of_real_def) then have "card (proots_within ppp {x. 0 = Im x}) = proots_unbounded_line_card ppp 0 1" unfolding proots_unbounded_line_card_def by simp ultimately show ?thesis by auto qed finally show ?thesis apply (fold pp_def,fold ppp_def) using that by auto qed ultimately show ?thesis by auto qed subsection \Polynomial roots on a closed ball\ \ \Roots counted WITH multiplicity\ definition proots_cball::"complex poly \ complex \ real \ nat" where "proots_cball p z0 r = proots_count p (cball z0 r)" \ \Roots counted WITHOUT multiplicity\ definition proots_cball_card ::"complex poly \ complex \ real \ nat" where "proots_cball_card p z0 r = card (proots_within p (cball z0 r))" (*FIXME: this surely can be optimised/refined.*) lemma proots_cball_card_code1[code]: "proots_cball_card p z0 r = ( if r=0 then (if poly p z0=0 then 1 else 0) else if r < 0 \ p=0 then 0 else ( let pp=fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:] in (if poly p (z0-r) =0 then 1 else 0) + proots_unbounded_line_card pp 0 1 + proots_upper_card pp ) )" proof - have ?thesis when "r=0" proof - have "proots_within p {z0} = (if poly p z0 = 0 then {z0} else {})" by auto then show ?thesis unfolding proots_cball_card_def using that by simp qed moreover have ?thesis when "r\0" "r < 0 \ p=0" proof - have ?thesis when "r<0" proof - have "proots_within p (cball z0 r) = {}" by (auto simp add: ball_empty that) then show ?thesis unfolding proots_cball_card_def using that by auto qed moreover have ?thesis when "r>0" "p=0" unfolding proots_cball_card_def using that infinite_cball[of r z0] by auto ultimately show ?thesis using that by argo qed moreover have ?thesis when "p\0" "r>0" proof - define pp where "pp=fcompose (p \\<^sub>p [:z0, of_real r:]) [:\,-1:] [:\,1:]" have "proots_cball_card p z0 r = card (proots_within p (sphere z0 r) \ proots_within p (ball z0 r))" unfolding proots_cball_card_def apply (simp add:proots_within_union) by (metis Diff_partition cball_diff_sphere sphere_cball) also have "... = card (proots_within p (sphere z0 r)) + card (proots_within p (ball z0 r))" apply (rule card_Un_disjoint) using \p\0\ by auto also have "... = (if poly p (z0-r) =0 then 1 else 0) + proots_unbounded_line_card pp 0 1 + proots_upper_card pp" using proots_sphere_card_code1[of p z0 r,folded pp_def,unfolded proots_sphere_card_def] proots_ball_card_code1[of p z0 r,folded pp_def,unfolded proots_ball_card_def] that by simp finally show ?thesis apply (fold pp_def) using that by auto qed ultimately show ?thesis by auto qed end diff --git a/thys/Jordan_Normal_Form/Conjugate.thy b/thys/Jordan_Normal_Form/Conjugate.thy --- a/thys/Jordan_Normal_Form/Conjugate.thy +++ b/thys/Jordan_Normal_Form/Conjugate.thy @@ -1,91 +1,91 @@ (* Author: René Thiemann Akihisa Yamada License: BSD *) theory Conjugate - imports HOL.Complex + imports HOL.Complex "HOL-Library.Complex_Order" begin class conjugate = fixes conjugate :: "'a \ 'a" assumes conjugate_id[simp]: "conjugate (conjugate a) = a" and conjugate_cancel_iff[simp]: "conjugate a = conjugate b \ a = b" class conjugatable_ring = ring + conjugate + assumes conjugate_dist_mul: "conjugate (a * b) = conjugate a * conjugate b" and conjugate_dist_add: "conjugate (a + b) = conjugate a + conjugate b" and conjugate_neg: "conjugate (-a) = - conjugate a" and conjugate_zero[simp]: "conjugate 0 = 0" begin lemma conjugate_zero_iff[simp]: "conjugate a = 0 \ a = 0" using conjugate_cancel_iff[of _ 0, unfolded conjugate_zero]. end class conjugatable_field = conjugatable_ring + field lemma sum_conjugate: fixes f :: "'b \ 'a :: conjugatable_ring" assumes finX: "finite X" shows "conjugate (sum f X) = sum (\x. conjugate (f x)) X" using finX by (induct set:finite, auto simp: conjugate_dist_add) class conjugatable_ordered_ring = conjugatable_ring + ordered_comm_monoid_add + assumes conjugate_square_positive: "a * conjugate a \ 0" class conjugatable_ordered_field = conjugatable_ordered_ring + field begin subclass conjugatable_field.. end lemma conjugate_square_0: fixes a :: "'a :: {conjugatable_ordered_ring, semiring_no_zero_divisors}" shows "a * conjugate a = 0 \ a = 0" by auto subsection \Instantiations\ instantiation complex :: conjugatable_ordered_field begin definition [simp]: "conjugate \ cnj" - definition [simp]: "x < y \ Im x = Im y \ Re x < Re y" - definition [simp]: "x \ y \ Im x = Im y \ Re x \ Re y" - instance by (intro_classes, auto simp: complex.expand) +instance + by intro_classes (auto simp: less_eq_complex_def) + end instantiation real :: conjugatable_ordered_field begin definition [simp]: "conjugate (x::real) \ x" instance by (intro_classes, auto) end instantiation rat :: conjugatable_ordered_field begin definition [simp]: "conjugate (x::rat) \ x" instance by (intro_classes, auto) end instantiation int :: conjugatable_ordered_ring begin definition [simp]: "conjugate (x::int) \ x" instance by (intro_classes, auto) end lemma conjugate_square_eq_0 [simp]: fixes x :: "'a :: {conjugatable_ring,semiring_no_zero_divisors}" shows "x * conjugate x = 0 \ x = 0" "conjugate x * x = 0 \ x = 0" by auto lemma conjugate_square_greater_0 [simp]: fixes x :: "'a :: {conjugatable_ordered_ring,ring_no_zero_divisors}" shows "x * conjugate x > 0 \ x \ 0" using conjugate_square_positive[of x] by (auto simp: le_less) lemma conjugate_square_smaller_0 [simp]: fixes x :: "'a :: {conjugatable_ordered_ring,ring_no_zero_divisors}" shows "\ x * conjugate x < 0" using conjugate_square_positive[of x] by auto end diff --git a/thys/LLL_Basis_Reduction/Norms.thy b/thys/LLL_Basis_Reduction/Norms.thy --- a/thys/LLL_Basis_Reduction/Norms.thy +++ b/thys/LLL_Basis_Reduction/Norms.thy @@ -1,667 +1,667 @@ (* Authors: Jose Divasón Maximilian Haslbeck Sebastiaan Joosten René Thiemann Akihisa Yamada License: BSD *) section \Norms\ text \In this theory we provide the basic definitions and properties of several norms of vectors and polynomials. \ theory Norms imports "HOL-Computational_Algebra.Polynomial" "HOL-Library.Adhoc_Overloading" "Jordan_Normal_Form.Conjugate" "Algebraic_Numbers.Resultant" (* only for poly_of_vec *) Missing_Lemmas begin subsection \L-\\\ Norms\ consts linf_norm :: "'a \ 'b" ("\(_)\\<^sub>\") definition linf_norm_vec where "linf_norm_vec v \ max_list (map abs (list_of_vec v) @ [0])" adhoc_overloading linf_norm linf_norm_vec definition linf_norm_poly where "linf_norm_poly f \ max_list (map abs (coeffs f) @ [0])" adhoc_overloading linf_norm linf_norm_poly lemma linf_norm_vec: "\vec n f\\<^sub>\ = max_list (map (abs \ f) [0..vCons a v\\<^sub>\ = max \a\ \v\\<^sub>\" by (auto simp: linf_norm_vec_def max_list_Cons) lemma linf_norm_vec_0 [simp]: "\vec 0 f\\<^sub>\ = 0" by (simp add: linf_norm_vec_def) lemma linf_norm_zero_vec [simp]: "\0\<^sub>v n :: 'a :: ordered_ab_group_add_abs vec\\<^sub>\ = 0" by (induct n, simp add: zero_vec_def, auto simp: zero_vec_Suc) lemma linf_norm_vec_ge_0 [intro!]: fixes v :: "'a :: ordered_ab_group_add_abs vec" shows "\v\\<^sub>\ \ 0" by (induct v, auto simp: max_def) lemma linf_norm_vec_eq_0 [simp]: fixes v :: "'a :: ordered_ab_group_add_abs vec" assumes "v \ carrier_vec n" shows "\v\\<^sub>\ = 0 \ v = 0\<^sub>v n" by (insert assms, induct rule: carrier_vec_induct, auto simp: zero_vec_Suc max_def) lemma linf_norm_vec_greater_0 [simp]: fixes v :: "'a :: ordered_ab_group_add_abs vec" assumes "v \ carrier_vec n" shows "\v\\<^sub>\ > 0 \ v \ 0\<^sub>v n" by (insert assms, induct rule: carrier_vec_induct, auto simp: zero_vec_Suc max_def) lemma linf_norm_poly_0 [simp]: "\0::_ poly\\<^sub>\ = 0" by (simp add: linf_norm_poly_def) lemma linf_norm_pCons [simp]: fixes p :: "'a :: ordered_ab_group_add_abs poly" shows "\pCons a p\\<^sub>\ = max \a\ \p\\<^sub>\" by (cases "p = 0", cases "a = 0", auto simp: linf_norm_poly_def max_list_Cons) lemma linf_norm_poly_ge_0 [intro!]: fixes f :: "'a :: ordered_ab_group_add_abs poly" shows "\f\\<^sub>\ \ 0" by (induct f, auto simp: max_def) lemma linf_norm_poly_eq_0 [simp]: fixes f :: "'a :: ordered_ab_group_add_abs poly" shows "\f\\<^sub>\ = 0 \ f = 0" by (induct f, auto simp: max_def) lemma linf_norm_poly_greater_0 [simp]: fixes f :: "'a :: ordered_ab_group_add_abs poly" shows "\f\\<^sub>\ > 0 \ f \ 0" by (induct f, auto simp: max_def) subsection \Square Norms\ consts sq_norm :: "'a \ 'b" ("\(_)\\<^sup>2") abbreviation "sq_norm_conjugate x \ x * conjugate x" adhoc_overloading sq_norm sq_norm_conjugate subsubsection \Square norms for vectors\ text \We prefer sum\_list over sum because it is not essentially dependent on commutativity, and easier for proving. \ definition "sq_norm_vec v \ \x \ list_of_vec v. \x\\<^sup>2" adhoc_overloading sq_norm sq_norm_vec lemma sq_norm_vec_vCons[simp]: "\vCons a v\\<^sup>2 = \a\\<^sup>2 + \v\\<^sup>2" by (simp add: sq_norm_vec_def) lemma sq_norm_vec_0[simp]: "\vec 0 f\\<^sup>2 = 0" by (simp add: sq_norm_vec_def) lemma sq_norm_vec_as_cscalar_prod: fixes v :: "'a :: conjugatable_ring vec" shows "\v\\<^sup>2 = v \c v" by (induct v, simp_all add: sq_norm_vec_def) lemma sq_norm_zero_vec[simp]: "\0\<^sub>v n :: 'a :: conjugatable_ring vec\\<^sup>2 = 0" by (simp add: sq_norm_vec_as_cscalar_prod) lemmas sq_norm_vec_ge_0 [intro!] = conjugate_square_ge_0_vec[folded sq_norm_vec_as_cscalar_prod] lemmas sq_norm_vec_eq_0 [simp] = conjugate_square_eq_0_vec[folded sq_norm_vec_as_cscalar_prod] lemmas sq_norm_vec_greater_0 [simp] = conjugate_square_greater_0_vec[folded sq_norm_vec_as_cscalar_prod] subsubsection \Square norm for polynomials\ definition sq_norm_poly where "sq_norm_poly p \ \a\coeffs p. \a\\<^sup>2" adhoc_overloading sq_norm sq_norm_poly lemma sq_norm_poly_0 [simp]: "\0::_poly\\<^sup>2 = 0" by (auto simp: sq_norm_poly_def) lemma sq_norm_poly_pCons [simp]: fixes a :: "'a :: conjugatable_ring" shows "\pCons a p\\<^sup>2 = \a\\<^sup>2 + \p\\<^sup>2" by (cases "p = 0"; cases "a = 0", auto simp: sq_norm_poly_def) lemma sq_norm_poly_ge_0 [intro!]: fixes p :: "'a :: conjugatable_ordered_ring poly" shows "\p\\<^sup>2 \ 0" by (unfold sq_norm_poly_def, rule sum_list_nonneg, auto intro!:conjugate_square_positive) lemma sq_norm_poly_eq_0 [simp]: fixes p :: "'a :: {conjugatable_ordered_ring,ring_no_zero_divisors} poly" shows "\p\\<^sup>2 = 0 \ p = 0" proof (induct p) case IH: (pCons a p) show ?case proof (cases "a = 0") case True with IH show ?thesis by simp next case False then have "\a\\<^sup>2 + \p\\<^sup>2 > 0" by (intro add_pos_nonneg, auto) then show ?thesis by auto qed qed simp lemma sq_norm_poly_pos [simp]: fixes p :: "'a :: {conjugatable_ordered_ring,ring_no_zero_divisors} poly" shows "\p\\<^sup>2 > 0 \ p \ 0" by (auto simp: less_le) lemma sq_norm_vec_of_poly [simp]: fixes p :: "'a :: conjugatable_ring poly" shows "\vec_of_poly p\\<^sup>2 = \p\\<^sup>2" apply (unfold sq_norm_poly_def sq_norm_vec_def) apply (fold sum_mset_sum_list) apply auto. lemma sq_norm_poly_of_vec [simp]: fixes v :: "'a :: conjugatable_ring vec" shows "\poly_of_vec v\\<^sup>2 = \v\\<^sup>2" apply (unfold sq_norm_poly_def sq_norm_vec_def coeffs_poly_of_vec) apply (fold rev_map) apply (fold sum_mset_sum_list) apply (unfold mset_rev) apply (unfold sum_mset_sum_list) by (auto intro: sum_list_map_dropWhile0) subsection \Relating Norms\ text \A class where ordering around 0 is linear.\ abbreviation (in ordered_semiring) is_real where "is_real a \ a < 0 \ a = 0 \ 0 < a" class semiring_real_line = ordered_semiring_strict + ordered_semiring_0 + assumes add_pos_neg_is_real: "a > 0 \ b < 0 \ is_real (a + b)" and mult_neg_neg: "a < 0 \ b < 0 \ 0 < a * b" and pos_pos_linear: "0 < a \ 0 < b \ a < b \ a = b \ b < a" and neg_neg_linear: "a < 0 \ b < 0 \ a < b \ a = b \ b < a" begin lemma add_neg_pos_is_real: "a < 0 \ b > 0 \ is_real (a + b)" using add_pos_neg_is_real[of b a] by (simp add: ac_simps) lemma nonneg_linorder_cases [consumes 2, case_names less eq greater]: assumes "0 \ a" and "0 \ b" and "a < b \ thesis" "a = b \ thesis" "b < a \ thesis" shows thesis using assms pos_pos_linear by (auto simp: le_less) lemma nonpos_linorder_cases [consumes 2, case_names less eq greater]: assumes "a \ 0" "b \ 0" and "a < b \ thesis" "a = b \ thesis" "b < a \ thesis" shows thesis using assms neg_neg_linear by (auto simp: le_less) lemma real_linear: assumes "is_real a" and "is_real b" shows "a < b \ a = b \ b < a" using pos_pos_linear neg_neg_linear assms by (auto dest: less_trans[of _ 0]) lemma real_linorder_cases [consumes 2, case_names less eq greater]: assumes real: "is_real a" "is_real b" and cases: "a < b \ thesis" "a = b \ thesis" "b < a \ thesis" shows thesis using real_linear[OF real] cases by auto lemma assumes a: "is_real a" and b: "is_real b" shows real_add_le_cancel_left_pos: "c + a \ c + b \ a \ b" and real_add_less_cancel_left_pos: "c + a < c + b \ a < b" and real_add_le_cancel_right_pos: "a + c \ b + c \ a \ b" and real_add_less_cancel_right_pos: "a + c < b + c \ a < b" using add_strict_left_mono[of b a c] add_strict_left_mono[of a b c] using add_strict_right_mono[of b a c] add_strict_right_mono[of a b c] by (atomize(full), cases rule: real_linorder_cases[OF a b], auto) lemma assumes a: "is_real a" and b: "is_real b" and c: "0 < c" shows real_mult_le_cancel_left_pos: "c * a \ c * b \ a \ b" and real_mult_less_cancel_left_pos: "c * a < c * b \ a < b" and real_mult_le_cancel_right_pos: "a * c \ b * c \ a \ b" and real_mult_less_cancel_right_pos: "a * c < b * c \ a < b" using mult_strict_left_mono[of b a c] mult_strict_left_mono[of a b c] c using mult_strict_right_mono[of b a c] mult_strict_right_mono[of a b c] c by (atomize(full), cases rule: real_linorder_cases[OF a b], auto) lemma assumes a: "is_real a" and b: "is_real b" shows not_le_real: "\ a \ b \ a < b" and not_less_real: "\ a > b \ a \ b" by (atomize(full), cases rule: real_linorder_cases[OF a b], auto simp: less_imp_le) lemma real_mult_eq_0_iff: assumes a: "is_real a" and b: "is_real b" shows "a * b = 0 \ a = 0 \ b = 0" proof- { assume l: "a * b = 0" and "a \ 0" and "b \ 0" with a b have "a < 0 \ 0 < a" and "b < 0 \ 0 < b" by auto then have "False" using mult_pos_pos[of a b] mult_pos_neg[of a b] mult_neg_pos[of a b] mult_neg_neg[of a b] by (auto simp:l) } then show ?thesis by auto qed end lemma real_pos_mult_max: fixes a b c :: "'a :: semiring_real_line" assumes c: "c > 0" and a: "is_real a" and b: "is_real b" shows "c * max a b = max (c * a) (c * b)" by (rule hom_max, simp add: real_mult_le_cancel_left_pos[OF a b c]) class ring_abs_real_line = ordered_ring_abs + semiring_real_line class semiring_1_real_line = semiring_real_line + monoid_mult + zero_less_one begin subclass ordered_semiring_1 by (unfold_locales, auto) lemma power_both_mono: "1 \ a \ m \ n \ a \ b \ a ^ m \ b ^ n" using power_mono[of a b n] power_increasing[of m n a] by (auto simp: order.trans[OF zero_le_one]) lemma power_pos: assumes a0: "0 < a" shows "0 < a ^ n" by (induct n, insert mult_strict_mono[OF a0] a0, auto) lemma power_neg: assumes a0: "a < 0" shows "odd n \ a ^ n < 0" and "even n \ a ^ n > 0" by (atomize(full), induct n, insert a0, auto simp add: mult_pos_neg2 mult_neg_neg) lemma power_ge_0_iff: assumes a: "is_real a" shows "0 \ a ^ n \ 0 \ a \ even n" using a proof (elim disjE) assume "a < 0" with power_neg[OF this, of n] show ?thesis by(cases "even n", auto) next assume "0 < a" with power_pos[OF this] show ?thesis by auto next assume "a = 0" then show ?thesis by (auto simp:power_0_left) qed lemma nonneg_power_less: assumes "0 \ a" and "0 \ b" shows "a^n < b^n \ n > 0 \ a < b" proof (insert assms, induct n arbitrary: a b) case 0 then show ?case by auto next case (Suc n) note a = \0 \ a\ note b = \0 \ b\ show ?case proof (cases "n > 0") case True from a b show ?thesis proof (cases rule: nonneg_linorder_cases) case less then show ?thesis by (auto simp: Suc.hyps[OF a b] True intro!:mult_strict_mono' a b zero_le_power) next case eq then show ?thesis by simp next case greater with Suc.hyps[OF b a] True have "b ^ n < a ^ n" by auto with mult_strict_mono'[OF greater this] b greater show ?thesis by auto qed qed auto qed lemma power_strict_mono: shows "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" by (subst nonneg_power_less, auto) lemma nonneg_power_le: assumes "0 \ a" and "0 \ b" shows "a^n \ b^n \ n = 0 \ a \ b" using assms proof (cases rule: nonneg_linorder_cases) case less with power_strict_mono[OF this, of n] assms show ?thesis by (cases n, auto) next case eq then show ?thesis by auto next case greater with power_strict_mono[OF this, of n] assms show ?thesis by (cases n, auto) qed end subclass (in linordered_idom) semiring_1_real_line apply unfold_locales by (auto simp: mult_strict_left_mono mult_strict_right_mono mult_neg_neg) class ring_1_abs_real_line = ring_abs_real_line + semiring_1_real_line begin subclass ring_1.. lemma abs_cases: assumes "a = 0 \ thesis" and "\a\ > 0 \ thesis" shows thesis using assms by auto lemma abs_linorder_cases[case_names less eq greater]: assumes "\a\ < \b\ \ thesis" and "\a\ = \b\ \ thesis" and "\b\ < \a\ \ thesis" shows thesis apply (cases rule: nonneg_linorder_cases[of "\a\" "\b\"]) using assms by auto lemma [simp]: shows not_le_abs_abs: "\ \a\ \ \b\ \ \a\ < \b\" and not_less_abs_abs: "\ \a\ > \b\ \ \a\ \ \b\" by (atomize(full), cases a b rule: abs_linorder_cases, auto simp: less_imp_le) lemma abs_power_less [simp]: "\a\^n < \b\^n \ n > 0 \ \a\ < \b\" by (subst nonneg_power_less, auto) lemma abs_power_le [simp]: "\a\^n \ \b\^n \ n = 0 \ \a\ \ \b\" by (subst nonneg_power_le, auto) lemma abs_power_pos [simp]: "\a\^n > 0 \ a \ 0 \ n = 0" using power_pos[of "\a\"] by (cases "n", auto) lemma abs_power_nonneg [intro!]: "\a\^n \ 0" by auto lemma abs_power_eq_0 [simp]: "\a\^n = 0 \ a = 0 \ n \ 0" apply (induct n, force) apply (unfold power_Suc) apply (subst real_mult_eq_0_iff, auto). end instance nat :: semiring_1_real_line by (intro_classes, auto) instance int :: ring_1_abs_real_line.. lemma vec_index_vec_of_list [simp]: "vec_of_list xs $ i = xs ! i" by transfer (auto simp: mk_vec_def undef_vec_def dest: empty_nth) lemma vec_of_list_append: "vec_of_list (xs @ ys) = vec_of_list xs @\<^sub>v vec_of_list ys" by (auto simp: nth_append) lemma linf_norm_vec_of_list: "\vec_of_list xs\\<^sub>\ = max_list (map abs xs @ [0])" by (simp add: linf_norm_vec_def) lemma linf_norm_vec_as_Greatest: fixes v :: "'a :: ring_1_abs_real_line vec" shows "\v\\<^sub>\ = (GREATEST a. a \ abs ` set (list_of_vec v) \ {0})" unfolding linf_norm_vec_of_list[of "list_of_vec v", simplified] by (subst max_list_as_Greatest, auto) lemma vec_of_poly_pCons: assumes "f \ 0" shows "vec_of_poly (pCons a f) = vec_of_poly f @\<^sub>v vec_of_list [a]" using assms by (auto simp: vec_eq_iff Suc_diff_le) lemma vec_of_poly_as_vec_of_list: assumes "f \ 0" shows "vec_of_poly f = vec_of_list (rev (coeffs f))" proof (insert assms, induct f) case 0 then show ?case by auto next case (pCons a f) then show ?case by (cases "f = 0", auto simp: vec_of_list_append vec_of_poly_pCons) qed lemma linf_norm_vec_of_poly [simp]: fixes f :: "'a :: ring_1_abs_real_line poly" shows "\vec_of_poly f\\<^sub>\ = \f\\<^sub>\" proof (cases "f = 0") case False then show ?thesis apply (unfold vec_of_poly_as_vec_of_list linf_norm_vec_of_list linf_norm_poly_def) apply (subst (1 2) max_list_as_Greatest, auto). qed simp lemma linf_norm_poly_as_Greatest: fixes f :: "'a :: ring_1_abs_real_line poly" shows "\f\\<^sub>\ = (GREATEST a. a \ abs ` set (coeffs f) \ {0})" using linf_norm_vec_as_Greatest[of "vec_of_poly f"] by simp lemma vec_index_le_linf_norm: fixes v :: "'a :: ring_1_abs_real_line vec" assumes "i < dim_vec v" shows "\v$i\ \ \v\\<^sub>\" apply (unfold linf_norm_vec_def, rule le_max_list) using assms apply (auto simp: in_set_conv_nth intro!: imageI exI[of _ i]). lemma coeff_le_linf_norm: fixes f :: "'a :: ring_1_abs_real_line poly" shows "\coeff f i\ \ \f\\<^sub>\" using vec_index_le_linf_norm[of "degree f - i" "vec_of_poly f"] by (cases "i \ degree f", auto simp: coeff_eq_0) class conjugatable_ring_1_abs_real_line = conjugatable_ring + ring_1_abs_real_line + power + assumes sq_norm_as_sq_abs [simp]: "\a\\<^sup>2 = \a\\<^sup>2" begin subclass conjugatable_ordered_ring by (unfold_locales, simp) end instance int :: conjugatable_ring_1_abs_real_line by (intro_classes, simp add: numeral_2_eq_2) instance rat :: conjugatable_ring_1_abs_real_line by (intro_classes, simp add: numeral_2_eq_2) instance real :: conjugatable_ring_1_abs_real_line by (intro_classes, simp add: numeral_2_eq_2) instance complex :: semiring_1_real_line apply intro_classes - by (auto simp: complex_eq_iff mult_le_cancel_left mult_le_cancel_right mult_neg_neg) + by (auto simp: complex_eq_iff mult_le_cancel_left mult_le_cancel_right mult_neg_neg less_complex_def less_eq_complex_def) text \ Due to the assumption @{thm abs_ge_self} from Groups.thy, @{type complex} cannot be @{class ring_1_abs_real_line}! \ instance complex :: ordered_ab_group_add_abs oops lemma sq_norm_as_sq_abs [simp]: "(sq_norm :: 'a :: conjugatable_ring_1_abs_real_line \ 'a) = power2 \ abs" by auto lemma sq_norm_vec_le_linf_norm: fixes v :: "'a :: {conjugatable_ring_1_abs_real_line} vec" assumes "v \ carrier_vec n" shows "\v\\<^sup>2 \ of_nat n * \v\\<^sub>\\<^sup>2" proof (insert assms, induct rule: carrier_vec_induct) case (Suc n a v) have [dest!]: "\ \a\ \ \v\\<^sub>\ \ of_nat n * \v\\<^sub>\\<^sup>2 \ of_nat n * \a\\<^sup>2" by (rule real_linorder_cases[of "\a\" "\v\\<^sub>\"], insert Suc, auto simp: less_le intro!: power_mono mult_left_mono) from Suc show ?case by (auto simp: ring_distribs max_def intro!:add_mono power_mono) qed simp lemma sq_norm_poly_le_linf_norm: fixes p :: "'a :: {conjugatable_ring_1_abs_real_line} poly" shows "\p\\<^sup>2 \ of_nat (degree p + 1) * \p\\<^sub>\\<^sup>2" using sq_norm_vec_le_linf_norm[of "vec_of_poly p" "degree p + 1"] by (auto simp: carrier_dim_vec) lemma coeff_le_sq_norm: fixes f :: "'a :: {conjugatable_ring_1_abs_real_line} poly" shows "\coeff f i\\<^sup>2 \ \f\\<^sup>2" proof (induct f arbitrary: i) case (pCons a f) show ?case proof (cases i) case (Suc ii) note pCons(2)[of ii] also have "\f\\<^sup>2 \ \a\\<^sup>2 + \f\\<^sup>2" by auto finally show ?thesis unfolding Suc by auto qed auto qed simp lemma max_norm_witness: fixes f :: "'a :: ordered_ring_abs poly" shows "\ i. \f\\<^sub>\ = \coeff f i\" by (induct f, auto simp add: max_def intro: exI[of _ "Suc _"] exI[of _ 0]) lemma max_norm_le_sq_norm: fixes f :: "'a :: conjugatable_ring_1_abs_real_line poly" shows "\f\\<^sub>\\<^sup>2 \ \f\\<^sup>2" proof - from max_norm_witness[of f] obtain i where id: "\f\\<^sub>\ = \coeff f i\" by auto show ?thesis unfolding id using coeff_le_sq_norm[of f i] by auto qed (*TODO MOVE*) lemma (in conjugatable_ring) conjugate_minus: "conjugate (x - y) = conjugate x - conjugate y" by (unfold diff_conv_add_uminus conjugate_dist_add conjugate_neg, rule) lemma conjugate_1[simp]: "(conjugate 1 :: 'a :: {conjugatable_ring, ring_1}) = 1" proof- have "conjugate 1 * 1 = (conjugate 1 :: 'a)" by simp also have "conjugate \ = 1" by simp finally show ?thesis by (unfold conjugate_dist_mul, simp) qed lemma conjugate_of_int [simp]: "(conjugate (of_int x) :: 'a :: {conjugatable_ring,ring_1}) = of_int x" proof (induct x) case (nonneg n) then show ?case by (induct n, auto simp: conjugate_dist_add) next case (neg n) then show ?case apply (induct n, auto simp: conjugate_minus conjugate_neg) by (metis conjugate_1 conjugate_dist_add one_add_one) qed lemma sq_norm_of_int: "\map_vec of_int v :: 'a :: {conjugatable_ring,ring_1} vec\\<^sup>2 = of_int \v\\<^sup>2" unfolding sq_norm_vec_as_cscalar_prod scalar_prod_def unfolding hom_distribs by (rule sum.cong, auto) definition "norm1 p = sum_list (map abs (coeffs p))" lemma norm1_ge_0: "norm1 (f :: 'a :: {abs,ordered_semiring_0,ordered_ab_group_add_abs}poly) \ 0" unfolding norm1_def by (rule sum_list_nonneg, auto) lemma norm2_norm1_main_equality: fixes f :: "nat \ 'a :: linordered_idom" shows "(\i = 0..f i\)\<^sup>2 = (\i = 0..i = 0..j = 0..f i\ * \f j\)" proof (induct n) case (Suc n) have id: "{0 ..< Suc n} = insert n {0 ..< n}" by auto have id: "sum f {0 ..< Suc n} = f n + sum f {0 ..< n}" for f :: "nat \ 'a" unfolding id by (rule sum.insert, auto) show ?case unfolding id power2_sum unfolding Suc by (auto simp: power2_eq_square sum_distrib_left sum.distrib ac_simps) qed auto lemma norm2_norm1_main_inequality: fixes f :: "nat \ 'a :: linordered_idom" shows "(\i = 0.. (\i = 0..f i\)\<^sup>2" unfolding norm2_norm1_main_equality by (auto intro!: sum_nonneg) lemma norm2_le_norm1_int: "\f :: int poly\\<^sup>2 \ (norm1 f)^2" proof - define F where "F = (!) (coeffs f)" define n where "n = length (coeffs f)" have 1: "\f\\<^sup>2 = (\i = 0..i = 0..F i\)" unfolding norm1_def sq_norm_poly_def sum_list_sum_nth F_def n_def by (subst sum.cong, auto) show ?thesis unfolding 1 2 by (rule norm2_norm1_main_inequality) qed lemma sq_norm_smult_vec: "sq_norm ((c :: 'a :: {conjugatable_ring,comm_semiring_0}) \\<^sub>v v) = (c * conjugate c) * sq_norm v" unfolding sq_norm_vec_as_cscalar_prod by (subst scalar_prod_smult_left, force, unfold conjugate_smult_vec, subst scalar_prod_smult_right, force, simp add: ac_simps) lemma vec_le_sq_norm: fixes v :: "'a :: conjugatable_ring_1_abs_real_line vec" assumes "v \ carrier_vec n" "i < n" shows "\v $ i\\<^sup>2 \ \v\\<^sup>2" using assms proof (induction v arbitrary: i) case (Suc n a v i) note IH = Suc show ?case proof (cases i) case (Suc ii) then show ?thesis using IH IH(2)[of ii] le_add_same_cancel2 order_trans by fastforce qed auto qed auto class trivial_conjugatable = conjugate + assumes conjugate_id [simp]: "conjugate x = x" class trivial_conjugatable_ordered_field = conjugatable_ordered_field + trivial_conjugatable class trivial_conjugatable_linordered_field = trivial_conjugatable_ordered_field + linordered_field begin subclass conjugatable_ring_1_abs_real_line by (standard) (auto simp add: semiring_normalization_rules) end instance rat :: trivial_conjugatable_linordered_field by (standard, auto) instance real :: trivial_conjugatable_linordered_field by (standard, auto) lemma scalar_prod_ge_0: "(x :: 'a :: linordered_idom vec) \ x \ 0" unfolding scalar_prod_def by (rule sum_nonneg, auto) lemma cscalar_prod_is_scalar_prod[simp]: "(x :: 'a :: trivial_conjugatable_ordered_field vec) \c y = x \ y" unfolding conjugate_id by (rule arg_cong[of _ _ "scalar_prod x"], auto) lemma scalar_prod_Cauchy: fixes u v::"'a :: {trivial_conjugatable_linordered_field} Matrix.vec" assumes "u \ carrier_vec n" "v \ carrier_vec n" shows "(u \ v)\<^sup>2 \ \u\\<^sup>2 * \v\\<^sup>2 " proof - { assume v_0: "v \ 0\<^sub>v n" have "0 \ (u - r \\<^sub>v v) \ (u - r \\<^sub>v v)" for r by (simp add: scalar_prod_ge_0) also have "(u - r \\<^sub>v v) \ (u - r \\<^sub>v v) = u \ u - r * (u \ v) - r * (u \ v) + r * r * (v \ v)" for r::'a proof - have "(u - r \\<^sub>v v) \ (u - r \\<^sub>v v) = (u - r \\<^sub>v v) \ u - (u - r \\<^sub>v v) \ (r \\<^sub>v v)" using assms by (subst scalar_prod_minus_distrib) auto also have "\ = u \ u - (r \\<^sub>v v) \ u - r * ((u - r \\<^sub>v v) \ v)" using assms by (subst minus_scalar_prod_distrib) auto also have "\ = u \ u - r * (v \ u) - r * (u \ v - r * (v \ v))" using assms by (subst minus_scalar_prod_distrib) auto also have "\ = u \ u - r * (u \ v) - r * (u \ v) + r * r * (v \ v)" using assms comm_scalar_prod by (auto simp add: field_simps) finally show ?thesis by simp qed also have "u \ u - r * (u \ v) - r * (u \ v) + r * r * (v \ v) = sq_norm u - (u \ v)\<^sup>2 / sq_norm v" if "r = (u \ v) / (v \ v)" for r unfolding that by (auto simp add: sq_norm_vec_as_cscalar_prod power2_eq_square) finally have "0 \ \u\\<^sup>2 - (u \ v)\<^sup>2 / \v\\<^sup>2" by auto then have "(u \ v)\<^sup>2 / \v\\<^sup>2 \ \u\\<^sup>2" by auto then have "(u \ v)\<^sup>2 \ \u\\<^sup>2 * \v\\<^sup>2" using pos_divide_le_eq[of "\v\\<^sup>2"] v_0 assms by (auto) } then show ?thesis by (fastforce simp add: assms) qed end diff --git a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy --- a/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy +++ b/thys/Prime_Distribution_Elementary/Partial_Zeta_Bounds.thy @@ -1,446 +1,447 @@ (* File: Partial_Zeta_Bounds.thy Author: Manuel Eberl, TU München Asymptotic bounds on sums over k^(-s) for k=1..n, for fixed s and variable n. *) section \Bounds on partial sums of the $\zeta$ function\ theory Partial_Zeta_Bounds imports Euler_MacLaurin.Euler_MacLaurin_Landau Zeta_Function.Zeta_Function Prime_Number_Theorem.Prime_Number_Theorem_Library Prime_Distribution_Elementary_Library begin (* TODO: This does not necessarily require the full complex zeta function. *) text \ We employ Euler--MacLaurin's summation formula to obtain asymptotic estimates for the partial sums of the Riemann $\zeta(s)$ function for fixed real $a$, i.\,e.\ the function \[f(n) = \sum_{k=1}^n k^{-s}\ .\] We distinguish various cases. The case $s = 1$ is simply the Harmonic numbers and is treated apart from the others. \ lemma harm_asymp_equiv: "sum_upto (\n. 1 / n) \[at_top] ln" proof - have "sum_upto (\n. n powr -1) \[at_top] (\x. ln (\x\ + 1))" proof (rule asymp_equiv_sandwich) have "eventually (\x. sum_upto (\n. n powr -1) x \ {ln (\x\ + 1)..ln \x\ + 1}) at_top" using eventually_ge_at_top[of 1] proof eventually_elim case (elim x) have "sum_upto (\n. real n powr -1) x = harm (nat \x\)" unfolding sum_upto_altdef harm_def by (intro sum.cong) (auto simp: field_simps powr_minus) also have "\ \ {ln (\x\ + 1)..ln \x\ + 1}" using elim harm_le[of "nat \x\"] ln_le_harm[of "nat \x\"] by (auto simp: le_nat_iff le_floor_iff) finally show ?case by simp qed thus "eventually (\x. sum_upto (\n. n powr -1) x \ ln (\x\ + 1)) at_top" "eventually (\x. sum_upto (\n. n powr -1) x \ ln \x\ + 1) at_top" by (eventually_elim; simp)+ qed real_asymp+ also have "\ \[at_top] ln" by real_asymp finally show ?thesis by (simp add: powr_minus field_simps) qed lemma fixes s :: real assumes s: "s > 0" "s \ 1" shows zeta_partial_sum_bigo_pos: "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s)) \ O(\x. real x powr -s)" and zeta_partial_sum_bigo_pos': "(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + Re (zeta s)) +o O(\x. real x powr -s)" proof - define F where "F = (\x. x powr (1 - s) / (1 - s))" define f where "f = (\x. x powr -s)" define f' where "f' = (\x. -s * x powr (-s-1))" define z where "z = Re (zeta s)" interpret euler_maclaurin_nat' F f "(!) [f, f']" 1 0 z "{}" proof have "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s) - real b powr -s / 2) \ Re (zeta s) - 0" proof (intro tendsto_diff) let ?g = "\b. (\i\<^sub>F b in at_top. Re (?g b) = (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)" using eventually_ge_at_top[of 1] proof eventually_elim case (elim b) have "(\k=1..b. real k powr -s) = (\kn. n - 1"]) auto also have "\ - real b powr (1 - s) / (1 - s) = Re (?g b)" by (auto simp: powr_Reals_eq add_ac) finally show ?case .. qed moreover have "(\b. Re (?g b)) \ Re (zeta s)" using hurwitz_zeta_critical_strip[of "of_real s" 1] s by (intro tendsto_intros) (simp add: zeta_def) ultimately show "(\b. (\k=1..b. real k powr -s) - real b powr (1 - s) / (1 - s)) \ Re (zeta s)" by (blast intro: Lim_transform_eventually) qed (use s in real_asymp) thus "(\b. (\k = 1..b. f (real k)) - F (real b) - (\i<2 * 0 + 1. (bernoulli' (Suc i) / fact (Suc i)) *\<^sub>R ([f, f'] ! i) (real b))) \ z" by (simp add: f_def F_def z_def) qed (use \s \ 1\ in \auto intro!: derivative_eq_intros continuous_intros simp flip: has_field_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) { fix n :: nat assume n: "n \ 1" have "(\k=1..n. real k powr -s) = n powr (1 - s) / (1 - s) + z + 1/2 * n powr -s - EM_remainder 1 f' (int n)" using euler_maclaurin_strong_nat'[of n] n by (simp add: F_def f_def) } note * = this have "(\n. (\k=1..n. real k powr -s) - n powr (1 - s) / (1 - s) - z) \ \(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n))" using * by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\n. 1/2 * n powr -s - EM_remainder 1 f' (int n)) \ O(\n. n powr -s)" proof (intro sum_in_bigo) have "(\x. norm (EM_remainder 1 f' (int x))) \ O(\x. real x powr - s)" proof (rule EM_remainder_strong_bigo_nat[where a = 1 and Y = "{}"]) fix x :: real assume "x \ 1" show "norm (f' x) \ s * x powr (-s-1)" using s by (simp add: f'_def) next from s show "((\x. x powr -s) \ 0) at_top" by real_asymp qed (auto simp: f'_def intro!: continuous_intros derivative_eq_intros) thus "(\x. EM_remainder 1 f' (int x)) \ O(\x. real x powr -s)" by simp qed real_asymp+ finally show "(\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - z) \ O(\x. real x powr -s)" . thus"(\n. \k=1..n. real k powr -s) =o (\n. real n powr (1 - s) / (1 - s) + z) +o O(\x. real x powr -s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_tail_bigo: fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n + 1) s)) \ O(\x. real x powr (1 - s))" proof - have [simp]: "complex_of_real (Re (zeta s)) = zeta s" using zeta_real[of s] by (auto elim!: Reals_cases) from s have s': "s > 0" "s \ 1" by auto have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s) + real n powr (1 - s) / (1 - s)) \ O(\x. real x powr (1 - s))" proof (rule sum_in_bigo) have "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) = (\n. (\k=1..n. real k powr -s) - real n powr (1 - s) / (1 - s) - Re (zeta s))" (is "?lhs = ?rhs") proof fix n :: nat have "hurwitz_zeta (1 + real n) s = zeta s - (\kauto simp: zeta_def powr_Reals_eq\) also have "(\kk=1..n. real k powr -s)" by (rule sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show "?lhs n = ?rhs n" by (simp add: add_ac) qed also have "\ \ O(\x. real x powr (-s))" by (rule zeta_partial_sum_bigo_pos) (use s in auto) also have "(\x. real x powr (-s)) \ O(\x. real x powr (1-s))" by real_asymp finally show "(\n. -Re (hurwitz_zeta (real n + 1) s) - real n powr (1 - s) / (1 - s)) \ \" . qed (use s in real_asymp) thus ?thesis by simp qed lemma zeta_tail_bigo': fixes s :: real assumes s: "s > 1" shows "(\n. Re (hurwitz_zeta (real n) s)) \ O(\x. real x powr (1 - s))" proof - have "(\n. Re (hurwitz_zeta (real n) s)) \ \(\n. Re (hurwitz_zeta (real (n - 1) + 1) s))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) (auto simp: of_nat_diff) also have "(\n. Re (hurwitz_zeta (real (n - 1) + 1) s)) \ O(\x. real (x - 1) powr (1 - s))" by (rule landau_o.big.compose[OF zeta_tail_bigo[OF assms]]) real_asymp also have "(\x. real (x - 1) powr (1 - s)) \ O(\x. real x powr (1 - s))" by real_asymp finally show ?thesis . qed lemma fixes s :: real assumes s: "s > 0" shows zeta_partial_sum_bigo_neg: "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" and zeta_partial_sum_bigo_neg': "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" proof - define F where "F = (\x. x powr (1 + s) / (1 + s))" define f where "f = (\x. x powr s)" define f' where "f' = (\x. s * x powr (s - 1))" have "(\i=1..n. f (real i)) - F n = 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)" if n: "n \ 1" for n proof - have "(\i\{1<..n}. f (real i)) - integral {real 1..real n} f = (\k<1. (bernoulli' (Suc k) / fact (Suc k)) *\<^sub>R (([f, f'] ! k) (real n) - ([f, f'] ! k) (real 1))) + EM_remainder' 1 ([f, f'] ! 1) (real 1) (real n)" by (rule euler_maclaurin_strong_raw_nat[where Y = "{}"]) (use \s > 0\ \n \ 1\ in \auto intro!: derivative_eq_intros continuous_intros simp flip: has_field_derivative_iff_has_vector_derivative simp: F_def f_def f'_def nth_Cons split: nat.splits\) also have "(\i\{1<..n}. f (real i)) = (\i\insert 1 {1<..n}. f (real i)) - f 1" using n by (subst sum.insert) auto also from n have "insert 1 {1<..n} = {1..n}" by auto finally have "(\i=1..n. f (real i)) - F n = f 1 + (integral {1..real n} f - F n) + (f (real n) - f 1) / 2 + EM_remainder' 1 f' 1 (real n)" by simp hence "(\i=1..n. f (real i)) - F n = 1 / 2 + (integral {1..real n} f - F n) + f (real n) / 2 + EM_remainder' 1 f' 1 (real n)" using s by (simp add: f_def diff_divide_distrib) also have "(f has_integral (F (real n) - F 1)) {1..real n}" using assms n by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative simp: F_def f_def intro!: derivative_eq_intros continuous_intros) hence "integral {1..real n} f - F n = - F 1" by (simp add: has_integral_iff) also have "1 / 2 + (-F 1) + f (real n) / 2 = 1 / 2 - F 1 + f n / 2" by simp finally show ?thesis . qed hence "(\n. (\i=1..n. f (real i)) - F n) \ \(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n))" by (intro bigthetaI_cong eventually_mono[OF eventually_ge_at_top[of 1]]) also have "(\n. 1 / 2 - F 1 + f n / 2 + EM_remainder' 1 f' 1 (real n)) \ O(\n. real n powr s)" unfolding F_def f_def proof (intro sum_in_bigo) have "(\x. integral {1..real x} (\t. pbernpoly 1 t *\<^sub>R f' t)) \ O(\n. 1 / s * real n powr s)" proof (intro landau_o.big.compose[OF integral_bigo]) have "(\x. pbernpoly 1 x * f' x) \ O(\x. 1 * x powr (s - 1))" by (intro landau_o.big.mult pbernpoly_bigo) (auto simp: f'_def) thus "(\x. pbernpoly 1 x *\<^sub>R f' x) \ O(\x. x powr (s - 1))" by simp from s show "filterlim (\a. 1 / s * a powr s) at_top at_top" by real_asymp next fix a' x :: real assume "a' \ 1" "a' \ x" thus "(\a. pbernpoly 1 a *\<^sub>R f' a) integrable_on {a'..x}" by (intro integrable_EM_remainder') (auto intro!: continuous_intros simp: f'_def) qed (use s in \auto intro!: filterlim_real_sequentially continuous_intros derivative_eq_intros\) thus "(\x. EM_remainder' 1 f' 1 (real x)) \ O(\n. real n powr s)" using \s > 0\ by (simp add: EM_remainder'_def) qed (use \s > 0\ in real_asymp)+ finally show "(\n. (\i=1..n. real i powr s) - n powr (1 + s) / (1 + s)) \ O(\n. n powr s)" by (simp add: f_def F_def) thus "(\n. (\i=1..n. real i powr s)) =o (\n. n powr (1 + s) / (1 + s)) +o O(\n. n powr s)" by (subst set_minus_plus [symmetric]) (simp_all add: fun_diff_def algebra_simps) qed lemma zeta_partial_sum_le_pos: assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - (x powr (1-s) / (1-s) + z)\ \ c * x powr -s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) + z - (real (Suc n) powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" by standard (use \s > 0\ in \real_asymp+\) show "(\n. real n powr - s) \ O(\n. real (Suc n) powr - s)" by real_asymp show "mono_on (\a. a powr - s) {1..} \ mono_on (\x. - (x powr - s)) {1..}" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith hence "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..}" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed thus "mono_on (\a. a powr (1 - s) / (1 - s) + z) {1..} \ mono_on (\x. - (x powr (1 - s) / (1 - s) + z)) {1..}" by blast qed auto lemma zeta_partial_sum_le_pos': assumes "s > 0" "s \ 1" defines "z \ Re (zeta (complex_of_real s))" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c" proof - have "\c>0. \x\1. \sum_upto (\n. n powr -s) x - x powr (1-s) / (1-s)\ \ c * 1" proof (rule sum_upto_asymptotics_lift_nat_real) have "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s) + z)) \ O(\n. real n powr - s)" using zeta_partial_sum_bigo_pos[OF assms(1,2)] unfolding z_def by (simp add: algebra_simps) also have "(\n. real n powr -s) \ O(\n. 1)" using assms by real_asymp finally have "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z) \ O(\n. 1)" by (simp add: algebra_simps) hence "(\n. (\k = 1..n. real k powr - s) - real n powr (1 - s) / (1 - s) - z + z) \ O(\n. 1)" by (rule sum_in_bigo) auto thus "(\n. (\k = 1..n. real k powr - s) - (real n powr (1 - s) / (1 - s))) \ O(\n. 1)" by simp from assms have "s < 1 \ s > 1" by linarith thus "(\n. real n powr (1 - s) / (1 - s) - (real (Suc n) powr (1 - s) / (1 - s))) \ O(\n. 1)" by standard (use \s > 0\ in \real_asymp+\) show "mono_on (\a. 1) {1..} \ mono_on (\x::real. -1 :: real) {1..}" using assms by (intro disjI2) (auto intro!: mono_onI powr_mono2') from assms have "s < 1 \ s > 1" by linarith hence "mono_on (\a. a powr (1 - s) / (1 - s)) {1..}" proof assume "s < 1" thus ?thesis using \s > 0\ by (intro mono_onI powr_mono2 divide_right_mono add_right_mono) auto next assume "s > 1" thus ?thesis by (intro mono_onI le_imp_neg_le add_right_mono divide_right_mono_neg powr_mono2') auto qed thus "mono_on (\a. a powr (1 - s) / (1 - s)) {1..} \ mono_on (\x. - (x powr (1 - s) / (1 - s))) {1..}" by blast qed auto thus ?thesis by simp qed lemma zeta_partial_sum_le_pos'': assumes "s > 0" "s \ 1" shows "\c>0. \x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" proof - from zeta_partial_sum_le_pos'[OF assms] obtain c where c: "c > 0" "\x. x \ 1 \ \sum_upto (\x. real x powr - s) x - x powr (1 - s) / (1 - s)\ \ c" by auto { fix x :: real assume x: "x \ 1" have "\sum_upto (\x. real x powr - s) x\ \ \x powr (1 - s) / (1 - s)\ + c" using c(1) c(2)[OF x] x by linarith also have "\x powr (1 - s) / (1 - s)\ = x powr (1 - s) / \1 - s\" using assms by simp also have "\ \ x powr max 0 (1 - s) / \1 - s\" using x by (intro divide_right_mono powr_mono) auto also have "c = c * x powr 0" using x by simp also have "c * x powr 0 \ c * x powr max 0 (1 - s)" using c(1) x by (intro mult_left_mono powr_mono) auto also have "x powr max 0 (1 - s) / \1 - s\ + c * x powr max 0 (1 - s) = (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by (simp add: algebra_simps) finally have "\sum_upto (\x. real x powr - s) x\ \ (1 / \1 - s\ + c) * x powr max 0 (1 - s)" by simp } moreover have "(1 / \1 - s\ + c) > 0" using c assms by (intro add_pos_pos divide_pos_pos) auto ultimately show ?thesis by blast qed lemma zeta_partial_sum_le_pos_bigo: assumes "s > 0" "s \ 1" shows "(\x. sum_upto (\n. n powr -s) x) \ O(\x. x powr max 0 (1 - s))" proof - from zeta_partial_sum_le_pos''[OF assms] obtain c where "\x\1. \sum_upto (\n. n powr -s) x\ \ c * x powr max 0 (1 - s)" by auto thus ?thesis by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto qed lemma zeta_partial_sum_01_asymp_equiv: assumes "s \ {0<..<1}" shows "sum_upto (\n. n powr -s) \[at_top] (\x. x powr (1 - s) / (1 - s))" proof - from zeta_partial_sum_le_pos'[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)\ \ c" by auto hence "(\x. sum_upto (\x. real x powr -s) x - x powr (1 - s) / (1 - s)) \ O(\_. 1)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\_. 1) \ o(\x. x powr (1 - s) / (1 - s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_gt_1_asymp_equiv: + fixes s :: real assumes "s > 1" defines "\ \ Re (zeta s)" shows "sum_upto (\n. n powr -s) \[at_top] (\x. \)" proof - have [simp]: "\ \ 0" using assms zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) from zeta_partial_sum_le_pos[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr -s) x - (x powr (1 - s) / (1 - s) + \)\ \ c * x powr -s" by (auto simp: \_def) hence "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s)) \ O(\x. x powr -s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr -s) \ o(\_. 1)" using \s > 1\ by real_asymp finally have "(\x. sum_upto (\x. real x powr -s) x - \ - x powr (1 - s) / (1 - s) + x powr (1 - s) / (1 - s)) \ o(\_. 1)" by (rule sum_in_smallo) (use \s > 1\ in real_asymp) thus ?thesis by (simp add: smallo_imp_asymp_equiv) qed lemma zeta_partial_sum_pos_bigtheta: assumes "s > 0" "s \ 1" shows "sum_upto (\n. n powr -s) \ \(\x. x powr max 0 (1 - s))" proof (cases "s > 1") case False thus ?thesis using asymp_equiv_imp_bigtheta[OF zeta_partial_sum_01_asymp_equiv[of s]] assms by (simp add: max_def) next case True have [simp]: "Re (zeta s) \ 0" using True zeta_Re_gt_1_nonzero[of s] zeta_real[of s] by (auto elim!: Reals_cases) show ?thesis using True asymp_equiv_imp_bigtheta[OF zeta_partial_sum_gt_1_asymp_equiv[of s]] by (simp add: max_def) qed lemma zeta_partial_sum_le_neg: assumes "s > 0" shows "\c>0. \x\1. \sum_upto (\n. n powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" proof (rule sum_upto_asymptotics_lift_nat_real) show "(\n. (\k = 1..n. real k powr s) - (real n powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using zeta_partial_sum_bigo_neg[OF assms(1)] by (simp add: algebra_simps) show "(\n. real n powr (1 + s) / (1 + s) - (real (Suc n) powr (1 + s) / (1 + s))) \ O(\n. real n powr s)" using assms by real_asymp show "(\n. real n powr s) \ O(\n. real (Suc n) powr s)" by real_asymp show "mono_on (\a. a powr s) {1..} \ mono_on (\x. - (x powr s)) {1..}" using assms by (intro disjI1) (auto intro!: mono_onI powr_mono2) show "mono_on (\a. a powr (1 + s) / (1 + s)) {1..} \ mono_on (\x. - (x powr (1 + s) / (1 + s))) {1..}" using assms by (intro disjI1 divide_right_mono powr_mono2 mono_onI) auto qed auto lemma zeta_partial_sum_neg_asymp_equiv: assumes "s > 0" shows "sum_upto (\n. n powr s) \[at_top] (\x. x powr (1 + s) / (1 + s))" proof - from zeta_partial_sum_le_neg[of s] assms obtain c where c: "c > 0" "\x\1. \sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)\ \ c * x powr s" by auto hence "(\x. sum_upto (\x. real x powr s) x - x powr (1 + s) / (1 + s)) \ O(\x. x powr s)" by (intro bigoI[of _ c] eventually_mono[OF eventually_ge_at_top[of 1]]) auto also have "(\x. x powr s) \ o(\x. x powr (1 + s) / (1 + s))" using assms by real_asymp finally show ?thesis by (rule smallo_imp_asymp_equiv) qed end \ No newline at end of file diff --git a/thys/Projective_Measurements/Linear_Algebra_Complements.thy b/thys/Projective_Measurements/Linear_Algebra_Complements.thy --- a/thys/Projective_Measurements/Linear_Algebra_Complements.thy +++ b/thys/Projective_Measurements/Linear_Algebra_Complements.thy @@ -1,2269 +1,2269 @@ (* Author: Mnacho Echenim, Université Grenoble Alpes *) theory Linear_Algebra_Complements imports "Isabelle_Marries_Dirac.Tensor" "Isabelle_Marries_Dirac.More_Tensor" "QHLProver.Gates" "HOL-Types_To_Sets.Group_On_With" "HOL-Probability.Probability" begin hide_const(open) S section \Preliminaries\ subsection \Misc\ lemma mult_real_cpx: fixes a::complex fixes b::complex assumes "a\ Reals" shows "a* (Re b) = Re (a * b)" using assms by (metis Reals_cases complex.exhaust complex.sel(1) complex_of_real_mult_Complex of_real_mult) lemma fct_bound: fixes f::"complex\ real" assumes "f(-1) + f 1 = 1" and "0 \ f 1" and "0 \ f (-1)" shows "-1 \ f 1 - f(-1) \ f 1 - f(-1) \ 1" proof have "f 1 - f(-1) = 1 - f(-1) - f(-1)" using assms by simp also have "...\ -1" using assms by simp finally show "-1 \ f 1 - f(-1)" . next have "f(-1) - f 1 = 1 - f 1 - f 1 " using assms by simp also have "... \ -1" using assms by simp finally have "-1 \ f(-1) - f 1" . thus "f 1 - f (-1) \ 1" by simp qed lemma fct_bound': fixes f::"complex\ real" assumes "f(-1) + f 1 = 1" and "0 \ f 1" and "0 \ f (-1)" shows "\f 1 - f(-1)\ \ 1" using assms fct_bound by auto lemma pos_sum_1_le: assumes "finite I" and "\ i \ I. (0::real) \ f i" and "(\i\ I. f i) = 1" and "j\ I" shows "f j \ 1" proof (rule ccontr) assume "\ f j \ 1" hence "1 < f j" by simp hence "1 < (\i\ I. f i)" using assms by (metis \\ f j \ 1\ sum_nonneg_leq_bound) thus False using assms by simp qed lemma last_subset: assumes "A \ {a,b}" and "a\ b" and "A \ {a, b}" and "A\ {}" and "A \ {a}" shows "A = {b}" using assms by blast lemma disjoint_Un: assumes "disjoint_family_on A (insert x F)" and "x\ F" shows "(A x) \ (\ a\ F. A a) = {}" proof - have "(A x) \ (\ a\ F. A a) = (\i\F. (A x) \ A i)" using Int_UN_distrib by simp also have "... = (\i\F. {})" using assms disjoint_family_onD by fastforce also have "... = {}" using SUP_bot_conv(2) by simp finally show ?thesis . qed lemma sum_but_one: assumes "\j < (n::nat). j \i \ f j = (0::'a::ring)" and "i < n" shows "(\ j \ {0 ..< n}. f j * g j) = f i * g i" proof - have "sum (\x. f x * g x) (({0 ..< n} - {i}) \ {i}) = sum (\x. f x * g x) ({0 ..< n} - {i}) + sum (\x. f x * g x) {i}" by (rule sum.union_disjoint, auto) also have "... = sum (\x. f x * g x) {i}" using assms by auto also have "... = f i * g i" by simp finally have "sum (\x. f x * g x) (({0 ..< n} - {i}) \ {i}) = f i * g i" . moreover have "{0 ..< n} = ({0 ..< n} - {i}) \ {i}" using assms by auto ultimately show ?thesis by simp qed lemma sum_2_elems: assumes "I = {a,b}" and "a\ b" shows "(\a\I. f a) = f a + f b" proof - have "(\a\I. f a) = (\a\(insert a {b}). f a)" using assms by simp also have "... = f a + (\a\{b}. f a)" proof (rule sum.insert) show "finite {b}" by simp show "a\ {b}" using assms by simp qed also have "... = f a + f b" by simp finally show ?thesis . qed lemma sum_4_elems: shows "(\i<(4::nat). f i) = f 0 + f 1 + f 2 + f 3" proof - have "(\i<(4::nat). f i) = (\i<(3::nat). f i) + f 3" by (metis Suc_numeral semiring_norm(2) semiring_norm(8) sum.lessThan_Suc) moreover have "(\i<(3::nat). f i) = (\i<(2::nat). f i) + f 2" by (metis Suc_1 add_2_eq_Suc' nat_1_add_1 numeral_code(3) numerals(1) one_plus_numeral_commute sum.lessThan_Suc) moreover have "(\i<(2::nat). f i) = (\i<(1::nat). f i) + f 1" by (metis Suc_1 sum.lessThan_Suc) ultimately show ?thesis by simp qed lemma disj_family_sum: shows "finite I \ disjoint_family_on A I \ (\i. i \ I \ finite (A i)) \ (\ i \ (\n \ I. A n). f i) = (\ n\ I. (\ i \ A n. f i))" proof (induct rule:finite_induct) case empty then show ?case by simp next case (insert x F) hence "disjoint_family_on A F" by (meson disjoint_family_on_mono subset_insertI) have "(\n \ (insert x F). A n) = A x \ (\n \ F. A n)" using insert by simp hence "(\ i \ (\n \ (insert x F). A n). f i) = (\ i \ (A x \ (\n \ F. A n)). f i)" by simp also have "... = (\ i \ A x. f i) + (\ i \ (\n \ F. A n). f i)" by (rule sum.union_disjoint, (simp add: insert disjoint_Un)+) also have "... = (\ i \ A x. f i) + (\n\F. sum f (A n))" using \disjoint_family_on A F\ by (simp add: insert) also have "... = (\n\(insert x F). sum f (A n))" using insert by simp finally show ?case . qed lemma integrable_real_mult_right: fixes c::real assumes "integrable M f" shows "integrable M (\w. c * f w)" proof (cases "c = 0") case True thus ?thesis by simp next case False thus ?thesis using integrable_mult_right[of c] assms by simp qed subsection \Unifying notions between Isabelle Marries Dirac and QHLProver\ lemma mult_conj_cmod_square: fixes z::complex shows "z * conjugate z = (cmod z)\<^sup>2" proof - have "z * conjugate z = (Re z)\<^sup>2 + (Im z)\<^sup>2" using complex_mult_cnj by auto also have "... = (cmod z)\<^sup>2" unfolding cmod_def by simp finally show ?thesis . qed lemma vec_norm_sq_cpx_vec_length_sq: shows "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" proof - have "(vec_norm v)\<^sup>2 = inner_prod v v" unfolding vec_norm_def using power2_csqrt by blast also have "... = (\i2)" unfolding Matrix.scalar_prod_def proof - have "\i. i < dim_vec v \ Matrix.vec_index v i * conjugate (Matrix.vec_index v i) = (cmod (Matrix.vec_index v i))\<^sup>2" using mult_conj_cmod_square by simp thus "(\i = 0..i2)" by (simp add: lessThan_atLeast0) qed finally show "(vec_norm v)\<^sup>2 = (cpx_vec_length v)\<^sup>2" unfolding cpx_vec_length_def by (simp add: sum_nonneg) qed lemma vec_norm_eq_cpx_vec_length: shows "vec_norm v = cpx_vec_length v" using vec_norm_sq_cpx_vec_length_sq by (metis cpx_vec_length_inner_prod inner_prod_csqrt power2_csqrt vec_norm_def) lemma cpx_vec_length_square: shows "\v\\<^sup>2 = (\i = 0..2)" unfolding cpx_vec_length_def by (simp add: lessThan_atLeast0 sum_nonneg) lemma state_qbit_norm_sq: assumes "v\ state_qbit n" shows "(cpx_vec_length v)\<^sup>2 = 1" proof - have "cpx_vec_length v = 1" using assms unfolding state_qbit_def by simp thus ?thesis by simp qed lemma dagger_adjoint: shows "dagger M = Complex_Matrix.adjoint M" unfolding dagger_def Complex_Matrix.adjoint_def by (simp add: cong_mat) subsection \Types to sets lemmata transfers\ context ab_group_add_on_with begin context includes lifting_syntax assumes ltd: "\(Rep::'s \ 'a) (Abs::'a \ 's). type_definition Rep Abs S" begin interpretation local_typedef_ab_group_add_on_with pls z mns um S "TYPE('s)" by unfold_locales fact lemmas lt_sum_union_disjoint = sum.union_disjoint [var_simplified explicit_ab_group_add, unoverload_type 'c, OF type.comm_monoid_add_axioms, untransferred] lemmas lt_disj_family_sum = disj_family_sum [var_simplified explicit_ab_group_add, unoverload_type 'd, OF type.comm_monoid_add_axioms, untransferred] lemmas lt_sum_reindex_cong = sum.reindex_cong [var_simplified explicit_ab_group_add, unoverload_type 'd, OF type.comm_monoid_add_axioms, untransferred] end lemmas sum_with_union_disjoint = lt_sum_union_disjoint [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] lemmas disj_family_sum_with = lt_disj_family_sum [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] lemmas sum_with_reindex_cong = lt_sum_reindex_cong [cancel_type_definition, OF carrier_ne, simplified pred_fun_def, simplified] end lemma (in comm_monoid_add_on_with) sum_with_cong': shows "finite I \ (\i. i\ I \ A i = B i) \ (\i. i\ I \ A i \ S) \ (\i. i\ I \ B i \ S) \ sum_with pls z A I = sum_with pls z B I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "sum_with pls z A (insert x F) = pls (A x) (sum_with pls z A F)" using insert sum_with_insert[of A] by (simp add: image_subset_iff) also have "... = pls (B x) (sum_with pls z B F)" using insert by simp also have "... = sum_with pls z B (insert x F)" using insert sum_with_insert[of B] by (simp add: image_subset_iff) finally show ?case . qed section \Linear algebra complements\ subsection \Additional properties of matrices\ lemma smult_one: shows "(1::'a::monoid_mult) \\<^sub>m A = A" by (simp add:eq_matI) lemma times_diag_index: fixes A::"'a::comm_ring Matrix.mat" assumes "A \ carrier_mat n n" and "B\ carrier_mat n n" and "diagonal_mat B" and "j < n" and "i < n" shows "Matrix.vec_index (Matrix.row (A*B) j) i = diag_mat B ! i *A $$ (j, i)" proof - have "Matrix.vec_index (Matrix.row (A*B) j) i = (A*B) $$ (j,i)" using Matrix.row_def[of "A*B" ] assms by simp also have "... = Matrix.scalar_prod (Matrix.row A j) (Matrix.col B i)" using assms times_mat_def[of A] by simp also have "... = Matrix.scalar_prod (Matrix.col B i) (Matrix.row A j)" using comm_scalar_prod[of "Matrix.row A j" n] assms by auto also have "... = (Matrix.vec_index (Matrix.col B i) i) * (Matrix.vec_index (Matrix.row A j) i)" unfolding Matrix.scalar_prod_def proof (rule sum_but_one) show "i < dim_vec (Matrix.row A j)" using assms by simp show "\ia i \ Matrix.vec_index (Matrix.col B i) ia = 0" using assms by (metis carrier_matD(1) carrier_matD(2) diagonal_mat_def index_col index_row(2)) qed also have "... = B $$(i,i) * (Matrix.vec_index (Matrix.row A j) i)" using assms by auto also have "... = diag_mat B ! i * (Matrix.vec_index (Matrix.row A j) i)" unfolding diag_mat_def using assms by simp also have "... = diag_mat B ! i * A $$ (j, i)" using assms by simp finally show ?thesis . qed lemma inner_prod_adjoint_comp: assumes "(U::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "(V::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "i < n" and "j < n" shows "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" proof - have "Complex_Matrix.inner_prod (Matrix.col V i) (Matrix.col U j) = Matrix.scalar_prod (Matrix.col U j) (Matrix.row (Complex_Matrix.adjoint V) i)" using adjoint_row[of i V] assms by simp also have "... = Matrix.scalar_prod (Matrix.row (Complex_Matrix.adjoint V) i) (Matrix.col U j)" by (metis adjoint_row assms(1) assms(2) assms(3) carrier_matD(1) carrier_matD(2) Matrix.col_dim conjugate_vec_sprod_comm) also have "... = ((Complex_Matrix.adjoint V) * U) $$ (i, j)" using assms by (simp add:times_mat_def) finally show ?thesis . qed lemma mat_unit_vec_col: assumes "(A::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "i < n" shows "A *\<^sub>v (unit_vec n i) = Matrix.col A i" proof show "dim_vec (A *\<^sub>v unit_vec n i) = dim_vec (Matrix.col A i)" using assms by simp show "\j. j < dim_vec (Matrix.col A i) \ Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.vec_index (Matrix.col A i) j" proof - fix j assume "j < dim_vec (Matrix.col A i)" hence "Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.scalar_prod (Matrix.row A j) (unit_vec n i)" unfolding mult_mat_vec_def by simp also have "... = Matrix.scalar_prod (unit_vec n i) (Matrix.row A j)" using comm_scalar_prod assms by auto also have "... = (Matrix.vec_index (unit_vec n i) i) * (Matrix.vec_index (Matrix.row A j) i)" unfolding Matrix.scalar_prod_def proof (rule sum_but_one) show "i < dim_vec (Matrix.row A j)" using assms by auto show "\ia i \ Matrix.vec_index (unit_vec n i) ia = 0" using assms unfolding unit_vec_def by auto qed also have "... = (Matrix.vec_index (Matrix.row A j) i)" using assms by simp also have "... = A $$ (j, i)" using assms \j < dim_vec (Matrix.col A i)\ by simp also have "... = Matrix.vec_index (Matrix.col A i) j" using assms \j < dim_vec (Matrix.col A i)\ by simp finally show "Matrix.vec_index (A *\<^sub>v unit_vec n i) j = Matrix.vec_index (Matrix.col A i) j" . qed qed lemma mat_prod_unit_vec_cong: assumes "(A::'a::conjugatable_field Matrix.mat) \ carrier_mat n n" and "B\ carrier_mat n n" and "\i. i < n \ A *\<^sub>v (unit_vec n i) = B *\<^sub>v (unit_vec n i)" shows "A = B" proof show "dim_row A = dim_row B" using assms by simp show "dim_col A = dim_col B" using assms by simp show "\i j. i < dim_row B \ j < dim_col B \ A $$ (i, j) = B $$ (i, j)" proof - fix i j assume ij: "i < dim_row B" "j < dim_col B" hence "A $$ (i,j) = Matrix.vec_index (Matrix.col A j) i" using assms by simp also have "... = Matrix.vec_index (A *\<^sub>v (unit_vec n j)) i" using mat_unit_vec_col[of A] ij assms by simp also have "... = Matrix.vec_index (B *\<^sub>v (unit_vec n j)) i" using assms ij by simp also have "... = Matrix.vec_index (Matrix.col B j) i" using mat_unit_vec_col ij assms by simp also have "... = B $$ (i,j)" using assms ij by simp finally show "A $$ (i, j) = B $$ (i, j)" . qed qed lemma smult_smult_times: fixes a::"'a::semigroup_mult" shows "a\\<^sub>m (k \\<^sub>m A) = (a * k)\\<^sub>m A" proof show r:"dim_row (a \\<^sub>m (k \\<^sub>m A)) = dim_row (a * k \\<^sub>m A)" by simp show c:"dim_col (a \\<^sub>m (k \\<^sub>m A)) = dim_col (a * k \\<^sub>m A)" by simp show "\i j. i < dim_row (a * k \\<^sub>m A) \ j < dim_col (a * k \\<^sub>m A) \ (a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = (a * k \\<^sub>m A) $$ (i, j)" proof - fix i j assume "i < dim_row (a * k \\<^sub>m A)" and "j < dim_col (a * k \\<^sub>m A)" note ij = this hence "(a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = a * (k \\<^sub>m A) $$ (i, j)" by simp also have "... = a * (k * A $$ (i,j))" using ij by simp also have "... = (a * k) * A $$ (i,j)" by (simp add: semigroup_mult_class.mult.assoc) also have "... = (a * k \\<^sub>m A) $$ (i, j)" using r c ij by simp finally show "(a \\<^sub>m (k \\<^sub>m A)) $$ (i, j) = (a * k \\<^sub>m A) $$ (i, j)" . qed qed lemma mat_minus_minus: fixes A :: "'a :: ab_group_add Matrix.mat" assumes "A \ carrier_mat n m" and "B\ carrier_mat n m" and "C\ carrier_mat n m" shows "A - (B - C) = A - B + C" proof show "dim_row (A - (B - C)) = dim_row (A - B + C)" using assms by simp show "dim_col (A - (B - C)) = dim_col (A - B + C)" using assms by simp show "\i j. i < dim_row (A - B + C) \ j < dim_col (A - B + C) \ (A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" proof - fix i j assume "i < dim_row (A - B + C)" and "j < dim_col (A - B + C)" note ij = this have "(A - (B - C)) $$ (i, j) = (A $$ (i,j) - B $$ (i,j) + C $$ (i,j))" using ij assms by simp also have "... = (A - B + C) $$ (i, j)" using assms ij by simp finally show "(A - (B - C)) $$ (i, j) = (A - B + C) $$ (i, j)" . qed qed subsection \Complements on complex matrices\ lemma hermitian_square: assumes "hermitian M" shows "M \ carrier_mat (dim_row M) (dim_row M)" proof - have "dim_col M = dim_row M" using assms unfolding hermitian_def adjoint_def by (metis adjoint_dim_col) thus ?thesis by auto qed lemma hermitian_add: assumes "A\ carrier_mat n n" and "B\ carrier_mat n n" and "hermitian A" and "hermitian B" shows "hermitian (A + B)" unfolding hermitian_def by (metis adjoint_add assms hermitian_def) lemma hermitian_minus: assumes "A\ carrier_mat n n" and "B\ carrier_mat n n" and "hermitian A" and "hermitian B" shows "hermitian (A - B)" unfolding hermitian_def by (metis adjoint_minus assms hermitian_def) lemma hermitian_smult: fixes a::real fixes A::"complex Matrix.mat" assumes "A \ carrier_mat n n" and "hermitian A" shows "hermitian (a \\<^sub>m A)" proof - have dim: "Complex_Matrix.adjoint A \ carrier_mat n n" using assms by (simp add: adjoint_dim) { fix i j assume "i < n" and "j < n" hence "Complex_Matrix.adjoint (a \\<^sub>m A) $$ (i,j) = a * (Complex_Matrix.adjoint A $$ (i,j))" using adjoint_scale[of a A] assms by simp also have "... = a * (A $$ (i,j))" using assms unfolding hermitian_def by simp also have "... = (a \\<^sub>m A) $$ (i,j)" using \i < n\ \j < n\ assms by simp finally have "Complex_Matrix.adjoint (a \\<^sub>m A) $$ (i,j) = (a \\<^sub>m A) $$ (i,j)" . } thus ?thesis using dim assms unfolding hermitian_def by auto qed lemma unitary_eigenvalues_norm_square: fixes U::"complex Matrix.mat" assumes "unitary U" and "U \ carrier_mat n n" and "eigenvalue U k" shows "conjugate k * k = 1" proof - have "\v. eigenvector U v k" using assms unfolding eigenvalue_def by simp from this obtain v where "eigenvector U v k" by auto define vn where "vn = vec_normalize v" have "eigenvector U vn k" using normalize_keep_eigenvector \eigenvector U v k\ using assms(2) eigenvector_def vn_def by blast have "vn \ carrier_vec n" using \eigenvector U v k\ assms(2) eigenvector_def normalized_vec_dim vn_def by blast have "Complex_Matrix.inner_prod vn vn = 1" using \vn = vec_normalize v\ \eigenvector U v k\ eigenvector_def normalized_vec_norm by blast hence "conjugate k * k = conjugate k * k * Complex_Matrix.inner_prod vn vn" by simp also have "... = conjugate k * Complex_Matrix.inner_prod vn (k \\<^sub>v vn)" proof - have "k * Complex_Matrix.inner_prod vn vn = Complex_Matrix.inner_prod vn (k \\<^sub>v vn)" using inner_prod_smult_left[of vn n vn k] \vn \ carrier_vec n\ by simp thus ?thesis by simp qed also have "... = Complex_Matrix.inner_prod (k \\<^sub>v vn) (k \\<^sub>v vn)" using inner_prod_smult_right[of vn n _ k] by (simp add: \vn \ carrier_vec n\) also have "... = Complex_Matrix.inner_prod (U *\<^sub>v vn) (U *\<^sub>v vn)" using \eigenvector U vn k\ unfolding eigenvector_def by simp also have "... = Complex_Matrix.inner_prod (Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn)) vn" using adjoint_def_alter[of "U *\<^sub>v vn" n vn n U] assms by (metis \eigenvector U vn k\ carrier_matD(1) carrier_vec_dim_vec dim_mult_mat_vec eigenvector_def) also have "... = Complex_Matrix.inner_prod vn vn" proof - have "Complex_Matrix.adjoint U *\<^sub>v (U *\<^sub>v vn) = (Complex_Matrix.adjoint U * U) *\<^sub>v vn" using assms by (metis \eigenvector U vn k\ adjoint_dim assoc_mult_mat_vec carrier_matD(1) eigenvector_def) also have "... = vn" using assms unfolding unitary_def inverts_mat_def by (metis \eigenvector U vn k\ assms(1) eigenvector_def one_mult_mat_vec unitary_simps(1)) finally show ?thesis by simp qed also have "... = 1" using \vn = vec_normalize v\ \eigenvector U v k\ eigenvector_def normalized_vec_norm by blast finally show ?thesis . qed lemma outer_prod_smult_left: fixes v::"complex Matrix.vec" shows "outer_prod (a \\<^sub>v v) w = a \\<^sub>m outer_prod v w" proof - define paw where "paw = outer_prod (a \\<^sub>v v) w" define apw where "apw = a \\<^sub>m outer_prod v w" have "paw = apw" proof have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_vec(2)) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim using carrier_vec_dim_vec by blast also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = a * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_smult_right: fixes v::"complex Matrix.vec" shows "outer_prod v (a \\<^sub>v w) = (conjugate a) \\<^sub>m outer_prod v w" proof - define paw where "paw = outer_prod v (a \\<^sub>v w)" define apw where "apw = (conjugate a) \\<^sub>m outer_prod v w" have "paw = apw" proof have "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec index_smult_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim using carrier_vec_dim_vec by (metis carrier_matD(2) index_smult_vec(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec smult_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (conjugate a) * (Matrix.vec_index v i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_add_left: fixes v::"complex Matrix.vec" assumes "dim_vec v = dim_vec x" shows "outer_prod (v + x) w = outer_prod v w + (outer_prod x w)" proof - define paw where "paw = outer_prod (v+x) w" define apw where "apw = outer_prod v w + (outer_prod x w)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec add_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (Matrix.vec_index v i + Matrix.vec_index x i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) + Matrix.vec_index x i * cnj (Matrix.vec_index w j)" by (simp add: ring_class.ring_distribs(2)) also have "... = (outer_prod v w) $$ (i,j) + (outer_prod x w) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_add_right: fixes v::"complex Matrix.vec" assumes "dim_vec w = dim_vec x" shows "outer_prod v (w + x) = outer_prod v w + (outer_prod v x)" proof - define paw where "paw = outer_prod v (w+x)" define apw where "apw = outer_prod v w + (outer_prod v x)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_add_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2) index_add_vec(2) paw_def) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis assms carrier_matD(2) carrier_vec_dim_vec index_add_mat(3)) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = Matrix.vec_index v i * (cnj (Matrix.vec_index w j + (Matrix.vec_index x j)))" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) + Matrix.vec_index v i * cnj (Matrix.vec_index x j)" by (simp add: ring_class.ring_distribs(1)) also have "... = (outer_prod v w) $$ (i,j) + (outer_prod v x) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_minus_left: fixes v::"complex Matrix.vec" assumes "dim_vec v = dim_vec x" shows "outer_prod (v - x) w = outer_prod v w - (outer_prod x w)" proof - define paw where "paw = outer_prod (v-x) w" define apw where "apw = outer_prod v w - (outer_prod x w)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_vec(2) paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2)) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis apw_def carrier_matD(2) carrier_vec_dim_vec minus_carrier_mat) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = (Matrix.vec_index v i - Matrix.vec_index x i) * cnj (Matrix.vec_index w j)" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) - Matrix.vec_index x i * cnj (Matrix.vec_index w j)" by (simp add: ring_class.ring_distribs) also have "... = (outer_prod v w) $$ (i,j) - (outer_prod x w) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_prod_minus_right: fixes v::"complex Matrix.vec" assumes "dim_vec w = dim_vec x" shows "outer_prod v (w - x) = outer_prod v w - (outer_prod v x)" proof - define paw where "paw = outer_prod v (w-x)" define apw where "apw = outer_prod v w - (outer_prod v x)" have "paw = apw" proof have rv: "dim_row paw = dim_vec v" unfolding paw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec paw_def) also have "... = dim_row apw" unfolding apw_def using outer_prod_dim assms by (metis carrier_matD(1) carrier_vec_dim_vec index_minus_mat(2)) finally show dr: "dim_row paw = dim_row apw" . have cw: "dim_col paw = dim_vec w" unfolding paw_def using outer_prod_dim assms using carrier_vec_dim_vec by (metis carrier_matD(2) index_minus_vec(2) paw_def) also have "... = dim_col apw" unfolding apw_def using outer_prod_dim by (metis assms carrier_matD(2) carrier_vec_dim_vec index_minus_mat(3)) finally show dc: "dim_col paw = dim_col apw" . show "\i j. i < dim_row apw \ j < dim_col apw \ paw $$ (i, j) = apw $$ (i, j)" proof - fix i j assume "i < dim_row apw" and "j < dim_col apw" note ij = this hence "paw $$ (i,j) = Matrix.vec_index v i * (cnj (Matrix.vec_index w j - (Matrix.vec_index x j)))" using dr dc unfolding paw_def outer_prod_def by simp also have "... = Matrix.vec_index v i * cnj (Matrix.vec_index w j) - Matrix.vec_index v i * cnj (Matrix.vec_index x j)" by (simp add: ring_class.ring_distribs) also have "... = (outer_prod v w) $$ (i,j) - (outer_prod v x) $$ (i,j)" using rv cw dr dc ij assms unfolding outer_prod_def by auto also have "... = apw $$ (i,j)" using dr dc ij unfolding apw_def outer_prod_def by simp finally show "paw $$ (i, j) = apw $$ (i, j)" . qed qed thus ?thesis unfolding paw_def apw_def by simp qed lemma outer_minus_minus: fixes a::"complex Matrix.vec" assumes "dim_vec a = dim_vec b" and "dim_vec u = dim_vec v" shows "outer_prod (a - b) (u - v) = outer_prod a u - outer_prod a v - outer_prod b u + outer_prod b v" proof - have "outer_prod (a - b) (u - v) = outer_prod a (u - v) - outer_prod b (u - v)" using outer_prod_minus_left assms by simp also have "... = outer_prod a u - outer_prod a v - outer_prod b (u - v)" using assms outer_prod_minus_right by simp also have "... = outer_prod a u - outer_prod a v - (outer_prod b u - outer_prod b v)" using assms outer_prod_minus_right by simp also have "... = outer_prod a u - outer_prod a v - outer_prod b u + outer_prod b v" proof (rule mat_minus_minus) show "outer_prod b u \ carrier_mat (dim_vec b) (dim_vec u)" by simp show "outer_prod b v \ carrier_mat (dim_vec b) (dim_vec u)" using assms by simp show "outer_prod a u - outer_prod a v \ carrier_mat (dim_vec b) (dim_vec u)" using assms by (metis carrier_vecI minus_carrier_mat outer_prod_dim) qed finally show ?thesis . qed lemma outer_trace_inner: assumes "A \ carrier_mat n n" and "dim_vec u = n" and "dim_vec v = n" shows "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.inner_prod v (A *\<^sub>v u)" proof - have "Complex_Matrix.trace (outer_prod u v * A) = Complex_Matrix.trace (A * outer_prod u v)" proof (rule trace_comm) show "A \ carrier_mat n n" using assms by simp show "outer_prod u v \ carrier_mat n n" using assms by (metis carrier_vec_dim_vec outer_prod_dim) qed also have "... = Complex_Matrix.inner_prod v (A *\<^sub>v u)" using trace_outer_prod_right[of A n u v] assms carrier_vec_dim_vec by metis finally show ?thesis . qed lemma zero_hermitian: shows "hermitian (0\<^sub>m n n)" unfolding hermitian_def by (metis adjoint_minus hermitian_def hermitian_one minus_r_inv_mat one_carrier_mat) lemma trace_1: shows "Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat) =(n::complex)" using one_mat_def by (simp add: Complex_Matrix.trace_def Matrix.mat_def) lemma trace_add: assumes "square_mat A" and "square_mat B" and "dim_row A = dim_row B" shows "Complex_Matrix.trace (A + B) = Complex_Matrix.trace A + Complex_Matrix.trace B" using assms by (simp add: Complex_Matrix.trace_def sum.distrib) lemma bra_vec_carrier: shows "bra_vec v \ carrier_mat 1 (dim_vec v)" proof - have "dim_row (ket_vec v) = dim_vec v" unfolding ket_vec_def by simp thus ?thesis using bra_bra_vec[of v] bra_def[of "ket_vec v"] by simp qed lemma mat_mult_ket_carrier: assumes "A\ carrier_mat n m" shows "A * |v\ \ carrier_mat n 1" using assms by (metis bra_bra_vec bra_vec_carrier carrier_matD(1) carrier_matI dagger_of_ket_is_bra dim_row_of_dagger index_mult_mat(2) index_mult_mat(3)) lemma mat_mult_ket: assumes "A \ carrier_mat n m" and "dim_vec v = m" shows "A * |v\ = |A *\<^sub>v v\" proof - have rn: "dim_row (A * |v\) = n" unfolding times_mat_def using assms by simp have co: "dim_col |A *\<^sub>v v\ = 1" using assms unfolding ket_vec_def by simp have cov: "dim_col |v\ = 1" using assms unfolding ket_vec_def by simp have er: "dim_row (A * |v\) = dim_row |A *\<^sub>v v\" using assms by (metis bra_bra_vec bra_vec_carrier carrier_matD(2) dagger_of_ket_is_bra dim_col_of_dagger dim_mult_mat_vec index_mult_mat(2)) have ec: "dim_col (A * |v\) = dim_col |A *\<^sub>v v\" using assms by (metis carrier_matD(2) index_mult_mat(3) mat_mult_ket_carrier) { fix i::nat fix j::nat assume "i < n" and "j < 1" hence "j = 0" by simp have "(A * |v\) $$ (i,0) = Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\ 0)" using times_mat_def[of A] \i < n\ rn cov by simp also have "... = Matrix.scalar_prod (Matrix.row A i) v" using ket_vec_col by simp also have "... = |A *\<^sub>v v\ $$ (i,j)" unfolding mult_mat_vec_def using \i < n\ \j = 0\ assms(1) by auto } note idx = this have "A * |v\ = Matrix.mat n 1 (\(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col |v\ j))" using assms unfolding times_mat_def ket_vec_def by simp also have "... = |A *\<^sub>v v\" using er ec idx rn co by auto finally show ?thesis . qed lemma unitary_density: assumes "density_operator R" and "unitary U" and "R\ carrier_mat n n" and "U\ carrier_mat n n" shows "density_operator (U * R * (Complex_Matrix.adjoint U))" unfolding density_operator_def proof (intro conjI) show "Complex_Matrix.positive (U * R * Complex_Matrix.adjoint U)" proof (rule positive_close_under_left_right_mult_adjoint) show "U \ carrier_mat n n" using assms by simp show "R \ carrier_mat n n" using assms by simp show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp qed have "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = Complex_Matrix.trace (Complex_Matrix.adjoint U * U * R)" using trace_comm[of "U * R" n "Complex_Matrix.adjoint U"] assms by (metis adjoint_dim mat_assoc_test(10)) also have "... = Complex_Matrix.trace R" using assms by simp also have "... = 1" using assms unfolding density_operator_def by simp finally show "Complex_Matrix.trace (U * R * Complex_Matrix.adjoint U) = 1" . qed subsection \Tensor product complements\ lemma tensor_vec_dim[simp]: shows "dim_vec (tensor_vec u v) = dim_vec u * (dim_vec v)" proof - have "length (mult.vec_vec_Tensor (*) (list_of_vec u) (list_of_vec v)) = length (list_of_vec u) * length (list_of_vec v)" using mult.vec_vec_Tensor_length[of "1::real" "(*)" "list_of_vec u" "list_of_vec v"] by (simp add: Matrix_Tensor.mult_def) thus ?thesis unfolding tensor_vec_def by simp qed lemma index_tensor_vec[simp]: assumes "0 < dim_vec v" and "i < dim_vec u * dim_vec v" shows "vec_index (tensor_vec u v) i = vec_index u (i div (dim_vec v)) * vec_index v (i mod dim_vec v)" proof - have m: "Matrix_Tensor.mult (1::complex) (*)" by (simp add: Matrix_Tensor.mult_def) have "length (list_of_vec v) = dim_vec v" using assms by simp hence "vec_index (tensor_vec u v) i = (*) (vec_index u (i div dim_vec v)) (vec_index v (i mod dim_vec v))" unfolding tensor_vec_def using mult.vec_vec_Tensor_elements assms m by (metis (mono_tags, lifting) length_greater_0_conv length_list_of_vec list_of_vec_index mult.vec_vec_Tensor_elements vec_of_list_index) thus ?thesis by simp qed lemma outer_prod_tensor_comm: fixes a::"complex Matrix.vec" fixes u::"complex Matrix.vec" assumes "0 < dim_vec a" and "0 < dim_vec b" shows "outer_prod (tensor_vec u v) (tensor_vec a b) = tensor_mat (outer_prod u a) (outer_prod v b)" proof - define ot where "ot = outer_prod (tensor_vec u v) (tensor_vec a b)" define to where "to = tensor_mat (outer_prod u a) (outer_prod v b)" define dv where "dv = dim_vec v" define db where "db = dim_vec b" have "ot = to" proof have ro: "dim_row ot = dim_vec u * dim_vec v" unfolding ot_def outer_prod_def by simp have "dim_row to = dim_row (outer_prod u a) * dim_row (outer_prod v b)" unfolding to_def by simp also have "... = dim_vec u * dim_vec v" using outer_prod_dim by (metis carrier_matD(1) carrier_vec_dim_vec) finally have rt: "dim_row to = dim_vec u * dim_vec v" . show "dim_row ot = dim_row to" using ro rt by simp have co: "dim_col ot = dim_vec a * dim_vec b" unfolding ot_def outer_prod_def by simp have "dim_col to = dim_col (outer_prod u a) * dim_col (outer_prod v b)" unfolding to_def by simp also have "... = dim_vec a * dim_vec b" using outer_prod_dim by (metis carrier_matD(2) carrier_vec_dim_vec) finally have ct: "dim_col to = dim_vec a * dim_vec b" . show "dim_col ot = dim_col to" using co ct by simp show "\i j. i < dim_row to \ j < dim_col to \ ot $$ (i, j) = to $$ (i, j)" proof - fix i j assume "i < dim_row to" and "j < dim_col to" note ij = this have "ot $$ (i,j) = Matrix.vec_index (tensor_vec u v) i * (conjugate (Matrix.vec_index (tensor_vec a b) j))" unfolding ot_def outer_prod_def using ij rt ct by simp also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (Matrix.vec_index (tensor_vec a b) j))" using ij rt assms unfolding dv_def by (metis index_tensor_vec less_nat_zero_code nat_0_less_mult_iff neq0_conv) also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (vec_index a (j div db) * vec_index b (j mod db)))" using ij ct assms unfolding db_def by simp also have "... = vec_index u (i div dv) * vec_index v (i mod dv) * (conjugate (vec_index a (j div db))) * (conjugate (vec_index b (j mod db)))" by simp also have "... = vec_index u (i div dv) * (conjugate (vec_index a (j div db))) * vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" by simp also have "... = (outer_prod u a) $$ (i div dv, j div db) * vec_index v (i mod dv) * (conjugate (vec_index b (j mod db)))" proof - have "i div dv < dim_vec u" using ij rt unfolding dv_def by (simp add: less_mult_imp_div_less) moreover have "j div db < dim_vec a" using ij ct assms unfolding db_def by (simp add: less_mult_imp_div_less) ultimately have "vec_index u (i div dv) * (conjugate (vec_index a (j div db))) = (outer_prod u a) $$ (i div dv, j div db)" unfolding outer_prod_def by simp thus ?thesis by simp qed also have "... = (outer_prod u a) $$ (i div dv, j div db) * (outer_prod v b) $$ (i mod dv, j mod db)" proof - have "i mod dv < dim_vec v" using ij rt unfolding dv_def using assms mod_less_divisor by (metis less_nat_zero_code mult.commute neq0_conv times_nat.simps(1)) moreover have "j mod db < dim_vec b" using ij ct assms unfolding db_def by (simp add: less_mult_imp_div_less) ultimately have "vec_index v (i mod dv) * (conjugate (vec_index b (j mod db))) = (outer_prod v b) $$ (i mod dv, j mod db)" unfolding outer_prod_def by simp thus ?thesis by simp qed also have "... = tensor_mat (outer_prod u a) (outer_prod v b) $$ (i, j)" proof (rule index_tensor_mat[symmetric]) show "dim_row (outer_prod u a) = dim_vec u" unfolding outer_prod_def by simp show "dim_row (outer_prod v b) = dv" unfolding outer_prod_def dv_def by simp show "dim_col (outer_prod v b) = db" unfolding db_def outer_prod_def by simp show "i < dim_vec u * dv" unfolding dv_def using ij rt by simp show "dim_col (outer_prod u a) = dim_vec a" unfolding outer_prod_def by simp show "j < dim_vec a * db" unfolding db_def using ij ct by simp show "0 < dim_vec a" using assms by simp show "0 < db" unfolding db_def using assms by simp qed finally show "ot $$ (i, j) = to $$ (i, j)" unfolding to_def . qed qed thus ?thesis unfolding ot_def to_def by simp qed lemma tensor_mat_adjoint: assumes "m1 \ carrier_mat r1 c1" and "m2 \ carrier_mat r2 c2" and "0 < c1" and "0 < c2" and "0 < r1" and "0 < r2" shows "Complex_Matrix.adjoint (tensor_mat m1 m2) = tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)" apply (rule eq_matI, auto) proof - fix i j assume "i < dim_col m1 * dim_col m2" and "j < dim_row m1 * dim_row m2" note ij = this have c1: "dim_col m1 = c1" using assms by simp have r1: "dim_row m1 = r1" using assms by simp have c2: "dim_col m2 = c2" using assms by simp have r2: "dim_row m2 = r2" using assms by simp have "Complex_Matrix.adjoint (m1 \ m2) $$ (i, j) = conjugate ((m1 \ m2) $$ (j, i))" using ij by (simp add: adjoint_eval) also have "... = conjugate (m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2))" proof - have "(m1 \ m2) $$ (j, i) = m1 $$ (j div r2, i div c2) * m2 $$ (j mod r2, i mod c2)" proof (rule index_tensor_mat[of m1 r1 c1 m2 r2 c2 j i], (auto simp add: assms ij c1 c2 r1 r2)) show "j < r1 * r2" using ij r1 r2 by simp show "i < c1 * c2" using ij c1 c2 by simp qed thus ?thesis by simp qed also have "... = conjugate (m1 $$ (j div r2, i div c2)) * conjugate ( m2 $$ (j mod r2, i mod c2))" by simp also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) * conjugate ( m2 $$ (j mod r2, i mod c2))" by (metis adjoint_eval c2 ij less_mult_imp_div_less r2) also have "... = (Complex_Matrix.adjoint m1) $$ (i div c2, j div r2) * (Complex_Matrix.adjoint m2) $$ (i mod c2, j mod r2)" by (metis Euclidean_Division.div_eq_0_iff adjoint_eval assms(4) bits_mod_div_trivial c2 gr_implies_not_zero ij(2) mult_not_zero r2) also have "... = (tensor_mat (Complex_Matrix.adjoint m1) (Complex_Matrix.adjoint m2)) $$ (i,j)" proof (rule index_tensor_mat[symmetric], (simp add: ij c1 c2 r1 r2) +) show "i < c1 * c2" using ij c1 c2 by simp show "j < r1 * r2" using ij r1 r2 by simp show "0 < r1" using assms by simp show "0 < r2" using assms by simp qed finally show "Complex_Matrix.adjoint (m1 \ m2) $$ (i, j) = (Complex_Matrix.adjoint m1 \ Complex_Matrix.adjoint m2) $$ (i, j)" . qed lemma index_tensor_mat': assumes "0 < dim_col A" and "0 < dim_col B" and "i < dim_row A * dim_row B" and "j < dim_col A * dim_col B" shows "(A \ B) $$ (i, j) = A $$ (i div (dim_row B), j div (dim_col B)) * B $$ (i mod (dim_row B), j mod (dim_col B))" by (rule index_tensor_mat, (simp add: assms)+) lemma tensor_mat_carrier: shows "tensor_mat U V \ carrier_mat (dim_row U * dim_row V) (dim_col U * dim_col V)" by auto lemma tensor_mat_id: assumes "0 < d1" and "0 < d2" shows "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) = 1\<^sub>m (d1 * d2)" proof (rule eq_matI, auto) show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, i) = 1" if "i < (d1 * d2)" for i using that index_tensor_mat'[of "1\<^sub>m d1" "1\<^sub>m d2"] by (simp add: assms less_mult_imp_div_less) next show "tensor_mat (1\<^sub>m d1) (1\<^sub>m d2) $$ (i, j) = 0" if "i < d1 * d2" "j < d1 * d2" "i \ j" for i j using that index_tensor_mat[of "1\<^sub>m d1" d1 d1 "1\<^sub>m d2" d2 d2 i j] by (metis assms(1) assms(2) index_one_mat(1) index_one_mat(2) index_one_mat(3) less_mult_imp_div_less mod_less_divisor mult_div_mod_eq mult_not_zero) qed lemma tensor_mat_hermitian: assumes "A \ carrier_mat n n" and "B \ carrier_mat n' n'" and "0 < n" and "0 < n'" and "hermitian A" and "hermitian B" shows "hermitian (tensor_mat A B)" using assms by (metis hermitian_def tensor_mat_adjoint) lemma tensor_mat_unitary: assumes "Complex_Matrix.unitary U" and "Complex_Matrix.unitary V" and "0 < dim_row U" and "0 < dim_row V" shows "Complex_Matrix.unitary (tensor_mat U V)" proof - define UI where "UI = tensor_mat U V" have "Complex_Matrix.adjoint UI = tensor_mat (Complex_Matrix.adjoint U) (Complex_Matrix.adjoint V)" unfolding UI_def proof (rule tensor_mat_adjoint) show "U \ carrier_mat (dim_row U) (dim_row U)" using assms unfolding Complex_Matrix.unitary_def by simp show "V \ carrier_mat (dim_row V) (dim_row V)" using assms unfolding Complex_Matrix.unitary_def by simp show "0 < dim_row V" using assms by simp show "0 < dim_row U" using assms by simp show "0 < dim_row V" using assms by simp show "0 < dim_row U" using assms by simp qed hence "UI * (Complex_Matrix.adjoint UI) = tensor_mat (U * Complex_Matrix.adjoint U) (V * Complex_Matrix.adjoint V)" using mult_distr_tensor[of U "Complex_Matrix.adjoint U" "V" "Complex_Matrix.adjoint V"] unfolding UI_def by (metis (no_types, lifting) Complex_Matrix.unitary_def adjoint_dim_col adjoint_dim_row assms carrier_matD(2) ) also have "... = tensor_mat (1\<^sub>m (dim_row U)) (1\<^sub>m (dim_row V))" using assms unitary_simps(2) by (metis Complex_Matrix.unitary_def) also have "... = (1\<^sub>m (dim_row U * dim_row V))" using tensor_mat_id assms by simp finally have "UI * (Complex_Matrix.adjoint UI) = (1\<^sub>m (dim_row U * dim_row V))" . hence "inverts_mat UI (Complex_Matrix.adjoint UI)" unfolding inverts_mat_def UI_def by simp thus ?thesis using assms unfolding Complex_Matrix.unitary_def UI_def by (metis carrier_matD(2) carrier_matI dim_col_tensor_mat dim_row_tensor_mat) qed subsection \Fixed carrier matrices locale\ text \We define a locale for matrices with a fixed number of rows and columns, and define a finite sum operation on this locale. The \verb+Type_To_Sets+ transfer tools then permits to obtain lemmata on the locale for free. \ locale fixed_carrier_mat = fixes fc_mats::"'a::field Matrix.mat set" fixes dimR dimC assumes fc_mats_carrier: "fc_mats = carrier_mat dimR dimC" begin sublocale semigroup_add_on_with fc_mats "(+)" proof (unfold_locales) show "\a b. a \ fc_mats \ b \ fc_mats \ a + b \ fc_mats" using fc_mats_carrier by simp show "\a b c. a \ fc_mats \ b \ fc_mats \ c \ fc_mats \ a + b + c = a + (b + c)" using fc_mats_carrier by simp qed sublocale ab_semigroup_add_on_with fc_mats "(+)" proof (unfold_locales) show "\a b. a \ fc_mats \ b \ fc_mats \ a + b = b + a" using fc_mats_carrier by (simp add: comm_add_mat) qed sublocale comm_monoid_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" proof (unfold_locales) show "0\<^sub>m dimR dimC \ fc_mats" using fc_mats_carrier by simp show "\a. a \ fc_mats \ 0\<^sub>m dimR dimC + a = a" using fc_mats_carrier by simp qed sublocale ab_group_add_on_with fc_mats "(+)" "0\<^sub>m dimR dimC" "(-)" "uminus" proof (unfold_locales) show "\a. a \ fc_mats \ - a + a = 0\<^sub>m dimR dimC" using fc_mats_carrier by simp show "\a b. a \ fc_mats \ b \ fc_mats \ a - b = a + - b" using fc_mats_carrier by (simp add: add_uminus_minus_mat) show "\a. a \ fc_mats \ - a \ fc_mats" using fc_mats_carrier by simp qed end lemma (in fixed_carrier_mat) smult_mem: assumes "A \ fc_mats" shows "a \\<^sub>m A \ fc_mats" using fc_mats_carrier assms by auto definition (in fixed_carrier_mat) sum_mat where "sum_mat A I = sum_with (+) (0\<^sub>m dimR dimC) A I" lemma (in fixed_carrier_mat) sum_mat_empty[simp]: shows "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp lemma (in fixed_carrier_mat) sum_mat_carrier: shows "(\i. i \ I \ (A i)\ fc_mats) \ sum_mat A I \ carrier_mat dimR dimC" unfolding sum_mat_def using sum_with_mem[of A I] fc_mats_carrier by auto lemma (in fixed_carrier_mat) sum_mat_insert: assumes "A x \ fc_mats" "A ` I \ fc_mats" and A: "finite I" and x: "x \ I" shows "sum_mat A (insert x I) = A x + sum_mat A I" unfolding sum_mat_def using assms sum_with_insert[of A x I] by simp subsection \A locale for square matrices\ locale cpx_sq_mat = fixed_carrier_mat "fc_mats::complex Matrix.mat set" for fc_mats + assumes dim_eq: "dimR = dimC" and npos: "0 < dimR" lemma (in cpx_sq_mat) one_mem: shows "1\<^sub>m dimR \ fc_mats" using fc_mats_carrier dim_eq by simp lemma (in cpx_sq_mat) square_mats: assumes "A \ fc_mats" shows "square_mat A" using fc_mats_carrier dim_eq assms by simp lemma (in cpx_sq_mat) cpx_sq_mat_mult: assumes "A \ fc_mats" and "B \ fc_mats" shows "A * B \ fc_mats" proof - have "dim_row (A * B) = dimR" using assms fc_mats_carrier by simp moreover have "dim_col (A * B) = dimR" using assms fc_mats_carrier dim_eq by simp ultimately show ?thesis using fc_mats_carrier carrier_mat_def dim_eq by auto qed lemma (in cpx_sq_mat) sum_mat_distrib_left: shows "finite I \ R\ fc_mats \ (\i. i \ I \ (A i)\ fc_mats) \ sum_mat (\i. R * (A i)) I = R * (sum_mat A I)" proof (induct rule: finite_induct) case empty hence a: "sum_mat (\i. R * (A i)) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp hence "R * (sum_mat A {}) = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R dimR dimC dimC] empty dim_eq by simp thus ?case using a by simp next case (insert x F) hence "sum_mat (\i. R * A i) (insert x F) = R * (A x) + sum_mat (\i. R * A i) F" using sum_mat_insert[of "\i. R * A i" x F] by (simp add: image_subsetI fc_mats_carrier dim_eq) also have "... = R * (A x) + R * (sum_mat A F)" using insert by simp also have "... = R * (A x + (sum_mat A F))" by (metis dim_eq fc_mats_carrier insert.prems(1) insert.prems(2) insertCI mult_add_distrib_mat sum_mat_carrier) also have "... = R * sum_mat A (insert x F)" proof - have "A x + (sum_mat A F) = sum_mat A (insert x F)" by (rule sum_mat_insert[symmetric], (auto simp add: insert)) thus ?thesis by simp qed finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_distrib_right: shows "finite I \ R\ fc_mats \ (\i. i \ I \ (A i)\ fc_mats) \ sum_mat (\i. (A i) * R) I = (sum_mat A I) * R" proof (induct rule: finite_induct) case empty hence a: "sum_mat (\i. (A i) * R) {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp have "sum_mat A {} = 0\<^sub>m dimR dimC" unfolding sum_mat_def by simp hence "(sum_mat A {}) * R = 0\<^sub>m dimR dimC" using fc_mats_carrier right_mult_zero_mat[of R ] dim_eq empty by simp thus ?case using a by simp next case (insert x F) have a: "(\i. A i * R) ` F \ fc_mats" using insert cpx_sq_mat_mult by (simp add: image_subsetI) have "A x * R \ fc_mats" using insert by (metis insertI1 local.fc_mats_carrier mult_carrier_mat dim_eq) hence "sum_mat (\i. A i * R) (insert x F) = (A x) * R + sum_mat (\i. A i * R) F" using insert a using sum_mat_insert[of "\i. A i * R" x F] by (simp add: image_subsetI local.fc_mats_carrier) also have "... = (A x) * R + (sum_mat A F) * R" using insert by simp also have "... = (A x + (sum_mat A F)) * R" proof (rule add_mult_distrib_mat[symmetric]) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier by blast show "R \ carrier_mat dimC dimC" using insert fc_mats_carrier dim_eq by simp qed also have "... = sum_mat A (insert x F) * R" proof - have "A x + (sum_mat A F) = sum_mat A (insert x F)" by (rule sum_mat_insert[symmetric], (auto simp add: insert)) thus ?thesis by simp qed finally show ?case . qed lemma (in cpx_sq_mat) trace_sum_mat: fixes A::"'b \ complex Matrix.mat" shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ Complex_Matrix.trace (sum_mat A I) = (\ i\ I. Complex_Matrix.trace (A i))" unfolding sum_mat_def proof (induct rule: finite_induct) case empty then show ?case using trace_zero dim_eq by simp next case (insert x F) have "Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A (insert x F)) = Complex_Matrix.trace (A x + sum_with (+) (0\<^sub>m dimR dimC) A F)" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff dim_eq) also have "... = Complex_Matrix.trace (A x) + Complex_Matrix.trace (sum_with (+) (0\<^sub>m dimR dimC) A F)" using trace_add square_mats insert by (metis carrier_matD(1) fc_mats_carrier image_subset_iff insert_iff sum_with_mem) also have "... = Complex_Matrix.trace (A x) + (\ i\ F. Complex_Matrix.trace (A i))" using insert by simp also have "... = (\ i\ (insert x F). Complex_Matrix.trace (A i))" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff) finally show ?case . qed lemma (in cpx_sq_mat) cpx_sq_mat_smult: assumes "A \ fc_mats" shows "x \\<^sub>m A \ fc_mats" using assms fc_mats_carrier by auto lemma (in cpx_sq_mat) mult_add_distrib_right: assumes "A\ fc_mats" "B\ fc_mats" "C\ fc_mats" shows "A * (B + C) = A * B + A * C" using assms fc_mats_carrier mult_add_distrib_mat dim_eq by simp lemma (in cpx_sq_mat) mult_add_distrib_left: assumes "A\ fc_mats" "B\ fc_mats" "C\ fc_mats" shows "(B + C) * A = B * A + C * A" using assms fc_mats_carrier add_mult_distrib_mat dim_eq by simp lemma (in cpx_sq_mat) mult_sum_mat_distrib_left: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ B \ fc_mats \ (sum_mat (\i. B * (A i)) I) = B * (sum_mat A I)" proof (induct rule: finite_induct) case empty hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp hence "B * (sum_mat A {}) = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq) moreover have "sum_mat (\i. B * (A i)) {} = 0\<^sub>m dimR dimC" using sum_mat_empty[of "\i. B * (A i)"] by simp ultimately show ?case by simp next case (insert x F) have "sum_mat (\i. B * (A i)) (insert x F) = B * (A x) + sum_mat (\i. B * (A i)) F" using sum_with_insert[of "\i. B * (A i)" x F] insert by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult) also have "... = B * (A x) + B * (sum_mat A F)" using insert by simp also have "... = B * (A x + (sum_mat A F))" proof (rule mult_add_distrib_right[symmetric]) show "B\ fc_mats" using insert by simp show "A x \ fc_mats" using insert by simp show "sum_mat A F \ fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier) qed also have "... = B * (sum_mat A (insert x F))" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff sum_mat_def) finally show ?case . qed lemma (in cpx_sq_mat) mult_sum_mat_distrib_right: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ B \ fc_mats \ (sum_mat (\i. (A i) * B) I) = (sum_mat A I) * B" proof (induct rule: finite_induct) case empty hence "sum_mat A {} = 0\<^sub>m dimR dimC" using sum_mat_empty by simp hence "(sum_mat A {}) * B = 0\<^sub>m dimR dimC" using empty by (simp add: fc_mats_carrier dim_eq) moreover have "sum_mat (\i. (A i) * B) {} = 0\<^sub>m dimR dimC" by simp ultimately show ?case by simp next case (insert x F) have "sum_mat (\i. (A i) * B) (insert x F) = (A x) * B + sum_mat (\i. (A i) * B) F" using sum_with_insert[of "\i. (A i) * B" x F] insert by (simp add: image_subset_iff local.sum_mat_def cpx_sq_mat_mult) also have "... = (A x) * B + (sum_mat A F) * B" using insert by simp also have "... = (A x + (sum_mat A F)) * B" proof (rule mult_add_distrib_left[symmetric]) show "B\ fc_mats" using insert by simp show "A x \ fc_mats" using insert by simp show "sum_mat A F \ fc_mats" using insert by (simp add: fc_mats_carrier sum_mat_carrier) qed also have "... = (sum_mat A (insert x F)) * B" using sum_with_insert[of A x F] insert by (simp add: image_subset_iff sum_mat_def) finally show ?case . qed lemma (in cpx_sq_mat) trace_sum_mat_mat_distrib: assumes "finite I" and "\i. i\ I \ B i \ fc_mats" and "A\ fc_mats" and "C \ fc_mats" shows "(\ i\ I. Complex_Matrix.trace(A * (B i) * C)) = Complex_Matrix.trace (A * (sum_mat B I) * C)" proof - have seq: "sum_mat (\i. A * (B i) * C) I = A * (sum_mat B I) * C" proof - have "sum_mat (\i. A * (B i) * C) I = (sum_mat (\i. A * (B i)) I) * C" proof (rule mult_sum_mat_distrib_right) show "finite I" using assms by simp show "C\ fc_mats" using assms by simp show "\i. i \ I \ A * B i \ fc_mats" using assms cpx_sq_mat_mult by simp qed moreover have "sum_mat (\i. A * (B i)) I = A * (sum_mat B I)" by (rule mult_sum_mat_distrib_left, (auto simp add: assms)) ultimately show "sum_mat (\i. A * (B i) * C) I = A * (sum_mat B I) * C" by simp qed have "(\ i\ I. Complex_Matrix.trace(A * (B i) * C)) = Complex_Matrix.trace (sum_mat (\i. A * (B i) * C) I)" proof (rule trace_sum_mat[symmetric]) show "finite I" using assms by simp fix i assume "i\ I" thus "A * B i * C \ fc_mats" using assms by (simp add: cpx_sq_mat_mult) qed also have "... = Complex_Matrix.trace (A * (sum_mat B I) * C)" using seq by simp finally show ?thesis . qed definition (in cpx_sq_mat) zero_col where "zero_col U = (\i. if i < dimR then Matrix.col U i else 0\<^sub>v dimR)" lemma (in cpx_sq_mat) zero_col_dim: assumes "U \ fc_mats" shows "dim_vec (zero_col U i) = dimR" using assms fc_mats_carrier unfolding zero_col_def by simp lemma (in cpx_sq_mat) zero_col_col: assumes "i < dimR" shows "zero_col U i = Matrix.col U i" using assms unfolding zero_col_def by simp lemma (in cpx_sq_mat) sum_mat_index: shows "finite I \ (\i. i \ I \ (A i)\ fc_mats) \ i < dimR \ j < dimC \ (sum_mat (\k. (A k)) I) $$ (i,j) = (\ k\I. (A k) $$ (i,j))" proof (induct rule: finite_induct) case empty thus ?case unfolding sum_mat_def by simp next case (insert x F) hence "(sum_mat (\k. (A k)) (insert x F)) $$ (i,j) = (A x + (sum_mat (\k. (A k)) F)) $$ (i,j)" using insert sum_mat_insert[of A] by (simp add: image_subsetI local.fc_mats_carrier) also have "... = (A x) $$ (i,j) + (sum_mat (\k. (A k)) F) $$ (i,j)" using insert sum_mat_carrier[of F A] fc_mats_carrier by simp also have "... = (A x) $$ (i,j) + (\ k\F. (A k) $$ (i,j))" using insert by simp also have "... = (\ k\(insert x F). (A k) $$ (i,j))" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_cong: shows "finite I \ (\i. i\ I \ A i = B i) \ (\i. i\ I \ A i \ fc_mats) \ (\i. i\ I \ B i \ fc_mats) \ sum_mat A I = sum_mat B I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "sum_mat A (insert x F) = A x + sum_mat A F" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = B x + sum_mat B F" using insert by simp also have "... = sum_mat B (insert x F)" using insert sum_mat_insert[of B] by (simp add: image_subset_iff) finally show ?case . qed lemma (in cpx_sq_mat) smult_sum_mat: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ a \\<^sub>m sum_mat A I = sum_mat (\i. a \\<^sub>m (A i)) I" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) have "a \\<^sub>m sum_mat A (insert x F) = a \\<^sub>m (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = a \\<^sub>m A x + a \\<^sub>m (sum_mat A F)" using insert by (metis add_smult_distrib_left_mat fc_mats_carrier insert_iff sum_mat_carrier) also have "... = a \\<^sub>m A x + sum_mat (\i. a \\<^sub>m (A i)) F" using insert by simp also have "... = sum_mat (\i. a \\<^sub>m (A i)) (insert x F)" using insert sum_mat_insert[of "(\i. a \\<^sub>m (A i))"] by (simp add: image_subset_iff cpx_sq_mat_smult) finally show ?case . qed lemma (in cpx_sq_mat) zero_sum_mat: shows "finite I \ sum_mat (\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) I = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" proof (induct rule:finite_induct) case empty then show ?case using dim_eq sum_mat_empty by auto next case (insert x F) have "sum_mat (\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)) (insert x F) = 0\<^sub>m dimR dimR + sum_mat (\i. 0\<^sub>m dimR dimR) F" using insert dim_eq zero_mem sum_mat_insert[of "\i. ((0\<^sub>m dimR dimR)::complex Matrix.mat)"] by fastforce also have "... = ((0\<^sub>m dimR dimR)::complex Matrix.mat)" using insert by auto finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_adjoint: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\i. Complex_Matrix.adjoint (A i)) I" proof (induct rule: finite_induct) case empty then show ?case using zero_hermitian[of dimR] by (metis (no_types) dim_eq hermitian_def sum_mat_empty) next case (insert x F) have "Complex_Matrix.adjoint (sum_mat A (insert x F)) = Complex_Matrix.adjoint (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = Complex_Matrix.adjoint (A x) + Complex_Matrix.adjoint (sum_mat A F)" proof (rule adjoint_add) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert fc_mats_carrier sum_mat_carrier[of F] by simp qed also have "... = Complex_Matrix.adjoint (A x) + sum_mat (\i. Complex_Matrix.adjoint (A i)) F" using insert by simp also have "... = sum_mat (\i. Complex_Matrix.adjoint (A i)) (insert x F)" proof (rule sum_mat_insert[symmetric], (auto simp add: insert)) show "Complex_Matrix.adjoint (A x) \ fc_mats" using insert fc_mats_carrier dim_eq by (simp add: adjoint_dim) show "\i. i \ F \ Complex_Matrix.adjoint (A i) \ fc_mats" using insert fc_mats_carrier dim_eq by (simp add: adjoint_dim) qed finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_hermitian: assumes "finite I" and "\i\ I. hermitian (A i)" and "\i\ I. A i\ fc_mats" shows "hermitian (sum_mat A I)" proof - have "Complex_Matrix.adjoint (sum_mat A I) = sum_mat (\i. Complex_Matrix.adjoint (A i)) I" using assms sum_mat_adjoint[of I] by simp also have "... = sum_mat A I" proof (rule sum_mat_cong, (auto simp add: assms)) show "\i. i \ I \ Complex_Matrix.adjoint (A i) = A i" using assms unfolding hermitian_def by simp show "\i. i \ I \ Complex_Matrix.adjoint (A i) \ fc_mats" using assms fc_mats_carrier dim_eq by (simp add: adjoint_dim) qed finally show ?thesis unfolding hermitian_def . qed lemma (in cpx_sq_mat) sum_mat_positive: shows "finite I \ (\i. i\ I \ Complex_Matrix.positive (A i)) \ (\i. i \ I \ A i \ fc_mats) \ Complex_Matrix.positive (sum_mat A I)" proof (induct rule: finite_induct) case empty then show ?case using positive_zero[of dimR] by (metis (no_types) dim_eq sum_mat_empty) next case (insert x F) hence "sum_mat A (insert x F) = A x + (sum_mat A F)" using sum_mat_insert[of A] by (simp add: image_subset_iff) moreover have "Complex_Matrix.positive (A x + (sum_mat A F))" proof (rule positive_add, (auto simp add: insert)) show "A x \ carrier_mat dimR dimR" using insert fc_mats_carrier dim_eq by simp show "sum_mat A F \ carrier_mat dimR dimR" using insert sum_mat_carrier dim_eq by (metis insertCI) qed ultimately show "Complex_Matrix.positive (sum_mat A (insert x F))" by simp qed lemma (in cpx_sq_mat) sum_mat_left_ortho_zero: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ (B \ fc_mats) \ (\ i. i\ I \ A i * B = (0\<^sub>m dimR dimR)) \ (sum_mat A I) * B = 0\<^sub>m dimR dimR" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_right) next case (insert x F) have "(sum_mat A (insert x F)) * B = (A x + sum_mat A F) * B" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = A x * B + sum_mat A F * B" proof (rule add_mult_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) show "B \ carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp qed also have "... = A x * B + (0\<^sub>m dimR dimR)" using insert by simp also have "... = 0\<^sub>m dimR dimR" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_right_ortho_zero: shows "finite I \ (\i. i\ I \ A i \ fc_mats) \ (B \ fc_mats) \ (\ i. i\ I \ B * A i = (0\<^sub>m dimR dimR)) \ B * (sum_mat A I) = 0\<^sub>m dimR dimR" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis finite.intros(1) sum_mat_empty mult_sum_mat_distrib_left) next case (insert x F) have "B * (sum_mat A (insert x F)) = B * (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: image_subset_iff) also have "... = B * A x + B * sum_mat A F" proof (rule mult_add_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) show "B \ carrier_mat dimC dimR" using insert fc_mats_carrier dim_eq by simp qed also have "... = B * A x + (0\<^sub>m dimR dimR)" using insert by simp also have "... = 0\<^sub>m dimR dimR" using insert by simp finally show ?case . qed lemma (in cpx_sq_mat) sum_mat_ortho_square: shows "finite I \ (\i. i\ I \ ((A i)::complex Matrix.mat) * (A i) = A i) \ (\i. i\ I \ A i \ fc_mats) \ (\ i j. i\ I \ j\ I \ i\ j \ A i * (A j) = (0\<^sub>m dimR dimR)) \ (sum_mat A I) * (sum_mat A I) = (sum_mat A I)" proof (induct rule:finite_induct) case empty then show ?case using dim_eq by (metis fc_mats_carrier right_mult_zero_mat sum_mat_empty zero_mem) next case (insert x F) have "(sum_mat A (insert x F)) * (sum_mat A (insert x F)) = (A x + sum_mat A F) * (A x + sum_mat A F)" using insert sum_mat_insert[of A] by (simp add: \\i. i \ insert x F \ A i * A i = A i\ image_subset_iff) also have "... = A x * (A x + sum_mat A F) + sum_mat A F * (A x + sum_mat A F)" proof (rule add_mult_distrib_mat) show "A x \ carrier_mat dimR dimC" using insert fc_mats_carrier by simp show "sum_mat A F \ carrier_mat dimR dimC" using insert by (metis insert_iff local.fc_mats_carrier sum_mat_carrier) thus "A x + sum_mat A F \ carrier_mat dimC dimC" using insert dim_eq by simp qed also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * (A x + sum_mat A F)" proof - have "A x * (A x + sum_mat A F) = A x * A x + A x * (sum_mat A F)" using dim_eq insert.prems(2) mult_add_distrib_right sum_mat_carrier by (metis fc_mats_carrier insertI1 subsetD subset_insertI) thus ?thesis by simp qed also have "... = A x * A x + A x * (sum_mat A F) + sum_mat A F * A x + sum_mat A F * (sum_mat A F)" proof - have "sum_mat A F * (A x + local.sum_mat A F) = sum_mat A F * A x + local.sum_mat A F * local.sum_mat A F" using insert dim_eq add_assoc add_mem mult_add_distrib_right cpx_sq_mat_mult sum_mat_carrier by (metis fc_mats_carrier insertI1 subsetD subset_insertI) hence "A x * A x + A x * sum_mat A F + sum_mat A F * (A x + sum_mat A F) = A x * A x + A x * sum_mat A F + (sum_mat A F * A x + sum_mat A F * sum_mat A F)" by simp also have "... = A x * A x + A x * sum_mat A F + sum_mat A F * A x + sum_mat A F * sum_mat A F" proof (rule assoc_add_mat[symmetric]) show "A x * A x + A x * sum_mat A F \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis add_mem cpx_sq_mat_mult insertCI) show "sum_mat A F * A x \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI) show "sum_mat A F * sum_mat A F \ carrier_mat dimR dimR" using sum_mat_carrier insert dim_eq fc_mats_carrier by (metis cpx_sq_mat_mult insertCI) qed finally show ?thesis . qed also have "... = A x + sum_mat A F" proof - have "A x * A x = A x" using insert by simp moreover have "sum_mat A F * sum_mat A F = sum_mat A F" using insert by simp moreover have "A x * (sum_mat A F) = 0\<^sub>m dimR dimR" proof - have "A x * (sum_mat A F) = sum_mat (\i. A x * (A i)) F" by (rule sum_mat_distrib_left[symmetric], (simp add: insert)+) also have "... = sum_mat (\i. 0\<^sub>m dimR dimR) F" proof (rule sum_mat_cong, (auto simp add: insert zero_mem)) show "\i. i \ F \ A x * A i = 0\<^sub>m dimR dimR" using insert by auto show "\i. i \ F \ A x * A i \ fc_mats" using insert cpx_sq_mat_mult by auto show "\i. i \ F \ 0\<^sub>m dimR dimR \ fc_mats" using zero_mem dim_eq by simp qed also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp finally show ?thesis . qed moreover have "sum_mat A F * A x = 0\<^sub>m dimR dimR" proof - have "sum_mat A F * A x = sum_mat (\i. A i * (A x)) F" by (rule sum_mat_distrib_right[symmetric], (simp add: insert)+) also have "... = sum_mat (\i. 0\<^sub>m dimR dimR) F" proof (rule sum_mat_cong, (auto simp add: insert zero_mem)) show "\i. i \ F \ A i * A x = 0\<^sub>m dimR dimR" using insert by auto show "\i. i \ F \ A i * A x \ fc_mats" using insert cpx_sq_mat_mult by auto show "\i. i \ F \ 0\<^sub>m dimR dimR \ fc_mats" using zero_mem dim_eq by simp qed also have "... = 0\<^sub>m dimR dimR" using zero_sum_mat insert by simp finally show ?thesis . qed ultimately show ?thesis using add_commute add_zero insert.prems(2) zero_mem dim_eq by auto qed also have "... = sum_mat A (insert x F)" using insert sum_mat_insert[of A x F] by (simp add: \\i. i \ insert x F \ A i * A i = A i\ image_subsetI) finally show ?case . qed lemma diagonal_unit_vec: assumes "B \ carrier_mat n n" and "diagonal_mat (B::complex Matrix.mat)" shows "B *\<^sub>v (unit_vec n i) = B $$ (i,i) \\<^sub>v (unit_vec n i)" proof - define v::"complex Matrix.vec" where "v = unit_vec n i" have "B *\<^sub>v v = Matrix.vec n (\ i. Matrix.scalar_prod (Matrix.row B i) v)" using assms unfolding mult_mat_vec_def by simp also have "... = Matrix.vec n (\ i. B $$(i,i) * Matrix.vec_index v i)" proof - have "\i < n. (Matrix.scalar_prod (Matrix.row B i) v = B $$(i,i) * Matrix.vec_index v i)" proof (intro allI impI) fix i assume "i < n" have "(Matrix.scalar_prod (Matrix.row B i) v) = (\ j \ {0 ..< n}. Matrix.vec_index (Matrix.row B i) j * Matrix.vec_index v j)" using assms unfolding Matrix.scalar_prod_def v_def by simp also have "... = Matrix.vec_index (Matrix.row B i) i * Matrix.vec_index v i" proof (rule sum_but_one) show "\j < n. j \ i \ Matrix.vec_index (Matrix.row B i) j = 0" proof (intro allI impI) fix j assume "j < n" and "j \ i" hence "Matrix.vec_index (Matrix.row B i) j = B $$ (i,j)" using \i < n\ \j < n\ assms by auto also have "... = 0" using assms \i < n\ \j < n\ \j\ i\ unfolding diagonal_mat_def by simp finally show "Matrix.vec_index (Matrix.row B i) j = 0" . qed show "i < n" using \i < n\ . qed also have "... = B $$(i,i) * Matrix.vec_index v i" using assms \i < n\ by auto finally show "(Matrix.scalar_prod (Matrix.row B i) v) = B $$(i,i) * Matrix.vec_index v i" . qed thus ?thesis by auto qed also have "... = B $$ (i,i) \\<^sub>v v" unfolding v_def unit_vec_def by auto finally have "B *\<^sub>v v = B $$ (i,i) \\<^sub>v v" . thus ?thesis unfolding v_def by simp qed lemma mat_vec_mult_assoc: assumes "A \ carrier_mat n p" and "B\ carrier_mat p q" and "dim_vec v = q" shows "A *\<^sub>v (B *\<^sub>v v) = (A * B) *\<^sub>v v" using assms by auto lemma (in cpx_sq_mat) similar_eigenvectors: assumes "A\ fc_mats" and "B\ fc_mats" and "P\ fc_mats" and "similar_mat_wit A B P (Complex_Matrix.adjoint P)" and "diagonal_mat B" and "i < n" shows "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = B $$ (i,i) \\<^sub>v (P *\<^sub>v (unit_vec dimR i))" proof - have "A *\<^sub>v (P *\<^sub>v (unit_vec dimR i)) = (P * B * (Complex_Matrix.adjoint P)) *\<^sub>v (P *\<^sub>v (unit_vec dimR i))" using assms unfolding similar_mat_wit_def by metis also have "... = P * B * (Complex_Matrix.adjoint P) * P *\<^sub>v (unit_vec dimR i)" proof (rule mat_vec_mult_assoc[of _ dimR dimR], (auto simp add: assms fc_mats_carrier)) show "P \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "P * B * Complex_Matrix.adjoint P \ carrier_mat dimR dimR" using assms fc_mats_carrier by auto qed also have "... = P * B * ((Complex_Matrix.adjoint P) * P) *\<^sub>v (unit_vec dimR i)" using assms dim_eq by (smt fc_mats_carrier mat_assoc_test(1) similar_mat_witD2(6) similar_mat_wit_sym) also have "... = P * B *\<^sub>v (unit_vec dimR i)" proof - have "(Complex_Matrix.adjoint P) * P = 1\<^sub>m dimR" using assms dim_eq unfolding similar_mat_wit_def by (simp add: fc_mats_carrier) thus ?thesis using assms(2) local.fc_mats_carrier dim_eq by auto qed also have "... = P *\<^sub>v (B *\<^sub>v (unit_vec dimR i))" using mat_vec_mult_assoc assms fc_mats_carrier dim_eq by simp also have "... = P *\<^sub>v (B $$ (i,i) \\<^sub>v (unit_vec dimR i))" using assms diagonal_unit_vec fc_mats_carrier dim_eq by simp also have "... = B $$ (i,i) \\<^sub>v (P *\<^sub>v (unit_vec dimR i))" proof (rule mult_mat_vec) show "P \ carrier_mat dimR dimC" using assms fc_mats_carrier by simp show "unit_vec dimR i \ carrier_vec dimC" using dim_eq by simp qed finally show ?thesis . qed subsection \Projectors\ definition projector where "projector M \ (hermitian M \ M * M = M)" lemma projector_hermitian: assumes "projector M" shows "hermitian M" using assms unfolding projector_def by simp lemma zero_projector[simp]: shows "projector (0\<^sub>m n n)" unfolding projector_def proof show "hermitian (0\<^sub>m n n)" using zero_hermitian[of n] by simp show "0\<^sub>m n n * 0\<^sub>m n n = 0\<^sub>m n n" by simp qed lemma projector_square_eq: assumes "projector M" shows "M * M = M" using assms unfolding projector_def by simp lemma projector_positive: assumes "projector M" shows "Complex_Matrix.positive M" proof (rule positive_if_decomp) show "M \ carrier_mat (dim_row M) (dim_row M)" using assms projector_hermitian hermitian_square by auto next have "M = Complex_Matrix.adjoint M" using assms projector_hermitian[of M] unfolding hermitian_def by simp hence "M * Complex_Matrix.adjoint M = M * M" by simp also have "... = M" using assms projector_square_eq by auto finally have "M * Complex_Matrix.adjoint M = M" . thus "\Ma. Ma * Complex_Matrix.adjoint Ma = M" by auto qed lemma projector_collapse_trace: assumes "projector (P::complex Matrix.mat)" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Complex_Matrix.trace (P * R * P) = Complex_Matrix.trace (R * P)" proof - have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R)" using trace_comm assms by auto also have "... = Complex_Matrix.trace ((P * P) * R)" using assms projector_square_eq[of P] by simp also have "... = Complex_Matrix.trace (P * (P * R))" using assms by auto also have "... = Complex_Matrix.trace (P * R * P)" using trace_comm[of P n "P * R"] assms by auto finally have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace (P * R * P)" . thus ?thesis by simp qed lemma positive_proj_trace: assumes "projector (P::complex Matrix.mat)" and "Complex_Matrix.positive R" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Complex_Matrix.trace (R * P) \ 0" proof - have "Complex_Matrix.trace (R * P) = Complex_Matrix.trace ((P * R) * P)" using assms projector_collapse_trace by auto also have "... = Complex_Matrix.trace ((P * R) * (Complex_Matrix.adjoint P))" using assms projector_hermitian[of P] unfolding hermitian_def by simp also have "... \ 0" proof (rule positive_trace) show " P * R * Complex_Matrix.adjoint P \ carrier_mat n n" using assms by auto show "Complex_Matrix.positive (P * R * Complex_Matrix.adjoint P)" by (rule positive_close_under_left_right_mult_adjoint[of _ n], (auto simp add: assms)) qed finally show ?thesis . qed lemma trace_proj_pos_real: assumes "projector (P::complex Matrix.mat)" and "Complex_Matrix.positive R" and "P \ carrier_mat n n" and "R\ carrier_mat n n" shows "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)" proof - have "Complex_Matrix.trace (R * P) \ 0" using assms positive_proj_trace by simp - thus ?thesis by (simp add: complex_eqI) + thus ?thesis by (simp add: complex_eqI less_eq_complex_def) qed lemma (in cpx_sq_mat) trace_sum_mat_proj_pos_real: fixes f::"'a \ real" assumes "finite I" and "\ i\ I. projector (P i)" and "Complex_Matrix.positive R" and "\i\ I. P i \ fc_mats" and "R \ fc_mats" shows "Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)) = Re (Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)))" proof - have sm: "\x. x \ I \ Complex_Matrix.trace (f x \\<^sub>m (R * P x)) = f x * Complex_Matrix.trace (R * P x)" proof - fix i assume "i\ I" show "Complex_Matrix.trace (f i \\<^sub>m (R * P i)) = f i * Complex_Matrix.trace (R * P i)" proof (rule trace_smult) show "R * P i \ carrier_mat dimR dimR" using assms cpx_sq_mat_mult fc_mats_carrier \i\ I\ dim_eq by simp qed qed have sw: "\x. x \ I \ R * (f x \\<^sub>m P x) = f x \\<^sub>m (R * P x)" proof - fix i assume "i \ I" show "R * (f i \\<^sub>m P i) = f i \\<^sub>m (R * P i)" proof (rule mult_smult_distrib) show "R\ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "P i \ carrier_mat dimR dimR" using assms \i\ I\ fc_mats_carrier dim_eq by simp qed qed have dr: "Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i)) I)) = Complex_Matrix.trace (sum_mat (\i. (R * (f i \\<^sub>m (P i)))) I)" using sum_mat_distrib_left[of I] assms by (simp add: cpx_sq_mat_smult) also have trs: "... = (\ i\ I. Complex_Matrix.trace (R * (f i \\<^sub>m (P i))))" proof (rule trace_sum_mat, (simp add: assms)) show "\i. i \ I \ R * (f i \\<^sub>m P i) \ fc_mats" using assms by (simp add: cpx_sq_mat_smult cpx_sq_mat_mult) qed also have "... = (\ i\ I. Complex_Matrix.trace (f i \\<^sub>m (R * (P i))))" by (rule sum.cong, (simp add: sw)+) also have "... = (\ i\ I. f i * Complex_Matrix.trace (R * (P i)))" by (rule sum.cong, (simp add: sm)+) also have "... = (\ i\ I. complex_of_real (f i * Re (Complex_Matrix.trace (R * (P i)))))" proof (rule sum.cong, simp) show "\x. x \ I \ complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" proof - fix x assume "x\ I" have "complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x) * complex_of_real (Re (Complex_Matrix.trace (R * P x)))" using assms sum.cong[of I I] fc_mats_carrier trace_proj_pos_real \x \ I\ dim_eq by auto also have "... = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" by simp finally show "complex_of_real (f x) * Complex_Matrix.trace (R * P x) = complex_of_real (f x * Re (Complex_Matrix.trace (R * P x)))" . qed qed also have "... = (\ i\ I. f i * Re (Complex_Matrix.trace (R * (P i))))" by simp also have "... = (\ i\ I. Re (Complex_Matrix.trace (f i \\<^sub>m (R * (P i)))))" proof - have "(\ i\ I. f i * Re (Complex_Matrix.trace (R * (P i)))) = (\ i\ I. Re (Complex_Matrix.trace (f i \\<^sub>m (R * (P i)))))" by (rule sum.cong, (simp add: sm)+) thus ?thesis by simp qed also have "... = (\ i\ I. Re (Complex_Matrix.trace (R * (f i \\<^sub>m (P i)))))" proof - have "\i. i \ I \ f i \\<^sub>m (R * (P i)) = R * (f i \\<^sub>m (P i))" using sw by simp thus ?thesis by simp qed also have "... = Re (\ i\ I. (Complex_Matrix.trace (R * (f i \\<^sub>m (P i)))))" by simp also have "... = Re (Complex_Matrix.trace (sum_mat (\i. R * (f i \\<^sub>m (P i))) I))" using trs by simp also have "... = Re (Complex_Matrix.trace (R * (sum_mat (\i. f i \\<^sub>m (P i))) I))" using dr by simp finally show ?thesis . qed subsection \Rank 1 projection\ definition rank_1_proj where "rank_1_proj v = outer_prod v v" lemma rank_1_proj_square_mat: shows "square_mat (rank_1_proj v)" using outer_prod_dim unfolding rank_1_proj_def by (metis carrier_matD(1) carrier_matD(2) carrier_vec_dim_vec square_mat.simps) lemma rank_1_proj_dim[simp]: shows "dim_row (rank_1_proj v) = dim_vec v" using outer_prod_dim unfolding rank_1_proj_def using carrier_vecI by blast lemma rank_1_proj_carrier[simp]: shows "rank_1_proj v \ carrier_mat (dim_vec v) (dim_vec v)" using outer_prod_dim unfolding rank_1_proj_def using carrier_vecI by blast lemma rank_1_proj_coord: assumes "i < dim_vec v" and "j < dim_vec v" shows "(rank_1_proj v) $$ (i, j) = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" using assms unfolding rank_1_proj_def outer_prod_def by auto lemma rank_1_proj_adjoint: shows "Complex_Matrix.adjoint (rank_1_proj (v::complex Matrix.vec)) = rank_1_proj v" proof show "dim_row (Complex_Matrix.adjoint (rank_1_proj v)) = dim_row (rank_1_proj v)" using rank_1_proj_square_mat by auto thus "dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)" by auto fix i j assume "i < dim_row (rank_1_proj v)" and "j < dim_col (rank_1_proj v)" note ij = this have "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = conjugate ((rank_1_proj v) $$ (j, i))" using ij \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_eval by fastforce also have "... = conjugate (Matrix.vec_index v j * (cnj (Matrix.vec_index v i)))" using rank_1_proj_coord ij by (metis \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_dim_col rank_1_proj_dim) also have "... = Matrix.vec_index v i * (cnj (Matrix.vec_index v j))" by simp also have "... = rank_1_proj v $$ (i, j)" using ij rank_1_proj_coord by (metis \dim_col (Complex_Matrix.adjoint (rank_1_proj v)) = dim_col (rank_1_proj v)\ adjoint_dim_col rank_1_proj_dim) finally show "Complex_Matrix.adjoint (rank_1_proj v) $$ (i, j) = rank_1_proj v $$ (i, j)". qed lemma rank_1_proj_hermitian: shows "hermitian (rank_1_proj (v::complex Matrix.vec))" using rank_1_proj_adjoint unfolding hermitian_def by simp lemma rank_1_proj_trace: assumes "\v\ = 1" shows "Complex_Matrix.trace (rank_1_proj v) = 1" proof - have "Complex_Matrix.trace (rank_1_proj v) = inner_prod v v" using trace_outer_prod unfolding rank_1_proj_def using carrier_vecI by blast also have "... = (vec_norm v)\<^sup>2" unfolding vec_norm_def using power2_csqrt by presburger also have "... = \v\\<^sup>2" using vec_norm_sq_cpx_vec_length_sq by simp also have "... = 1" using assms by simp finally show ?thesis . qed lemma rank_1_proj_mat_col: assumes "A \ carrier_mat n n" and "i < n" and "j < n" and "k < n" shows "(rank_1_proj (Matrix.col A i)) $$ (j, k) = A $$ (j, i) * conjugate (A $$ (k,i))" proof - have "(rank_1_proj (Matrix.col A i)) $$ (j, k) = Matrix.vec_index (Matrix.col A i) j * conjugate (Matrix.vec_index (Matrix.col A i) k)" using index_outer_prod[of "Matrix.col A i" n "Matrix.col A i" n] assms unfolding rank_1_proj_def by simp also have "... = A $$ (j, i) * conjugate (A $$ (k,i))" using assms by simp finally show ?thesis . qed lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary_index: assumes "A \ fc_mats" and "B \ fc_mats" and "diagonal_mat B" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (A * B * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (\ i\ {..< dimR}. ((diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) $$ (j,k))" proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms)) show "\i. i < dimR \ (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" using rank_1_proj_carrier assms fc_mats_carrier dim_eq by (metis smult_carrier_mat zero_col_col zero_col_dim) show "k < dimC" using assms dim_eq by simp qed also have "... = (\ i\ {..< dimR}. (diag_mat B)!i* A $$ (j, i) * conjugate (A $$ (k,i)))" proof (rule sum.cong, simp) have idx: "\x. x \ {.. (rank_1_proj (Matrix.col A x)) $$ (j, k) = A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms fc_mats_carrier dim_eq by blast show "\x. x \ {.. ((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" proof - fix x assume "x\ {..< dimR}" have "((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * (rank_1_proj (Matrix.col A x)) $$ (j, k)" proof (rule index_smult_mat) show "j < dim_row (rank_1_proj (Matrix.col A x))" using \x\ {..< dimR}\ assms fc_mats_carrier by simp show "k < dim_col (rank_1_proj (Matrix.col A x))" using \x\ {..< dimR}\ assms fc_mats_carrier rank_1_proj_carrier[of "Matrix.col A x"] by simp qed thus "((diag_mat B)!x \\<^sub>m rank_1_proj (Matrix.col A x)) $$ (j, k) = (diag_mat B)!x * A $$ (j, x) * conjugate (A $$ (k, x))" using idx \x\ {..< dimR}\ by simp qed qed also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j) " unfolding Matrix.scalar_prod_def proof (rule sum.cong) show "{..x. x \ {0.. diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row (A*B) j) x" proof - fix x assume "x\ {0.. {0.. carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "B \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "x < dimR" using \x\{0..< dimR}\ by simp qed moreover have "conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \x\ {0.. assms fc_mats_carrier dim_eq by (simp add: adjoint_col) ultimately show "diag_mat B ! x * A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row (A*B) j) x" by simp qed qed also have "... = (A*B * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "Matrix.mat (dim_row (A*B)) (dim_col (Complex_Matrix.adjoint A)) (\(i, j). Matrix.scalar_prod (Matrix.row (A*B) i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$ (j, k) = Matrix.scalar_prod (Matrix.row (A*B) j) (Matrix.col (Complex_Matrix.adjoint A) k)" using assms fc_mats_carrier by simp also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row (A*B) j)" using assms comm_scalar_prod[of "Matrix.row (A*B) j" dimR] fc_mats_carrier dim_eq by (metis adjoint_dim Matrix.col_carrier_vec row_carrier_vec cpx_sq_mat_mult) thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp qed finally show ?thesis . qed lemma (in cpx_sq_mat) weighted_sum_rank_1_proj_unitary: assumes "A \ fc_mats" and "B \ fc_mats" and "diagonal_mat B" and "Complex_Matrix.unitary A" shows "(sum_mat (\i. (diag_mat B)!i \\<^sub>m rank_1_proj (Matrix.col A i)) {..< dimR}) = (A * B * (Complex_Matrix.adjoint A))" proof show "dim_row (sum_mat (\i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..i j. i < dim_row (A * B * Complex_Matrix.adjoint A) \ j < dim_col (A * B * Complex_Matrix.adjoint A) \ local.sum_mat (\i. diag_mat B ! i \\<^sub>m rank_1_proj (Matrix.col A i)) {..v\ = 1" shows "projector (rank_1_proj v)" proof - have "Complex_Matrix.adjoint (rank_1_proj v) = rank_1_proj v" using rank_1_proj_adjoint by simp hence a: "hermitian (rank_1_proj v)" unfolding hermitian_def by simp have "rank_1_proj v * rank_1_proj v = inner_prod v v \\<^sub>m outer_prod v v" unfolding rank_1_proj_def using outer_prod_mult_outer_prod assms using carrier_vec_dim_vec by blast also have "... = (vec_norm v)\<^sup>2 \\<^sub>m outer_prod v v" unfolding vec_norm_def using power2_csqrt by presburger also have "... = (cpx_vec_length v)\<^sup>2 \\<^sub>m outer_prod v v " using vec_norm_sq_cpx_vec_length_sq by simp also have "... = outer_prod v v" using assms state_qbit_norm_sq smult_one[of "outer_prod v v"] by simp also have "... = rank_1_proj v" unfolding rank_1_proj_def by simp finally show ?thesis using a unfolding projector_def by simp qed lemma rank_1_proj_positive: assumes "\v\ = 1" shows "Complex_Matrix.positive (rank_1_proj v)" proof - have "projector (rank_1_proj v)" using assms rank_1_proj_projector by simp thus ?thesis using projector_positive by simp qed lemma rank_1_proj_density: assumes "\v\ = 1" shows "density_operator (rank_1_proj v)" unfolding density_operator_def using assms by (simp add: rank_1_proj_positive rank_1_proj_trace) lemma (in cpx_sq_mat) sum_rank_1_proj_unitary_index: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (1\<^sub>m dimR) $$ (j,k)" proof - have "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR}) $$ (j,k) = (\ i\ {..< dimR}. (rank_1_proj (Matrix.col A i)) $$ (j,k))" proof (rule sum_mat_index, (auto simp add: fc_mats_carrier assms)) show "\i. i < dimR \ rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" proof - fix i assume "i < dimR" hence "dim_vec (Matrix.col A i) = dimR" using assms fc_mats_carrier by simp thus "rank_1_proj (Matrix.col A i) \ carrier_mat dimR dimC" using rank_1_proj_carrier assms fc_mats_carrier dim_eq fc_mats_carrier by (metis dim_col fc_mats_carrier) qed show "k < dimC" using assms dim_eq by simp qed also have "... = (\ i\ {..< dimR}. A $$ (j, i) * conjugate (A $$ (k,i)))" proof (rule sum.cong, simp) show "\x. x \ {.. rank_1_proj (Matrix.col A x) $$ (j, k) = A $$ (j, x) * conjugate (A $$ (k, x))" using rank_1_proj_mat_col assms using local.fc_mats_carrier dim_eq by blast qed also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j) " unfolding Matrix.scalar_prod_def proof (rule sum.cong) show "{..x. x \ {0.. A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x" proof - fix x assume "x\ {0.. {0..x\ {0.. assms fc_mats_carrier dim_eq by simp moreover have "conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x" using \x\ {0.. assms fc_mats_carrier dim_eq by (simp add: adjoint_col) ultimately show "A $$ (j, x) * conjugate (A $$ (k, x)) = Matrix.vec_index ((Matrix.col (Complex_Matrix.adjoint A) k)) x * Matrix.vec_index (Matrix.row A j) x" by simp qed qed also have "... = (A * (Complex_Matrix.adjoint A)) $$ (j,k)" proof - have "Matrix.mat (dim_row A) (dim_col (Complex_Matrix.adjoint A)) (\(i, j). Matrix.scalar_prod (Matrix.row A i) (Matrix.col (Complex_Matrix.adjoint A) j)) $$ (j, k) = Matrix.scalar_prod (Matrix.row A j) (Matrix.col (Complex_Matrix.adjoint A) k)" using assms fc_mats_carrier by simp also have "... = Matrix.scalar_prod (Matrix.col (Complex_Matrix.adjoint A) k) (Matrix.row A j)" using assms comm_scalar_prod[of "Matrix.row A j" dimR] fc_mats_carrier by (simp add: adjoint_dim dim_eq) thus ?thesis using assms fc_mats_carrier unfolding times_mat_def by simp qed also have "... = (1\<^sub>m dimR) $$ (j,k)" using assms dim_eq by (simp add: fc_mats_carrier) finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_sum_density: assumes "finite I" and "\i\ I. \u i\ = 1" and "\i\ I. dim_vec (u i) = dimR" and "\i\ I. 0 \ p i" and "(\i\ I. p i) = 1" shows "density_operator (sum_mat (\i. p i \\<^sub>m (rank_1_proj (u i))) I)" unfolding density_operator_def proof have "Complex_Matrix.trace (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I) = (\i\ I. Complex_Matrix.trace (p i \\<^sub>m rank_1_proj (u i)))" proof (rule trace_sum_mat, (simp add: assms)) show "\i. i \ I \ p i \\<^sub>m rank_1_proj (u i) \ fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto qed also have "... = (\i\ I. p i * Complex_Matrix.trace (rank_1_proj (u i)))" proof - { fix i assume "i \ I" hence "Complex_Matrix.trace (p i \\<^sub>m rank_1_proj (u i)) = p i * Complex_Matrix.trace (rank_1_proj (u i))" using trace_smult[of "rank_1_proj (u i)" dimR] assms fc_mats_carrier dim_eq by auto } thus ?thesis by simp qed also have "... = 1" using assms rank_1_proj_trace assms by auto finally show "Complex_Matrix.trace (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I) = 1" . next show "Complex_Matrix.positive (sum_mat (\i. p i \\<^sub>m rank_1_proj (u i)) I)" proof (rule sum_mat_positive, (simp add: assms)) fix i assume "i\ I" thus "p i \\<^sub>m rank_1_proj (u i) \ fc_mats" using assms smult_mem fc_mats_carrier dim_eq by auto show "Complex_Matrix.positive (p i \\<^sub>m rank_1_proj (u i))" using assms \i\ I\ rank_1_proj_positive positive_smult[of "rank_1_proj (u i)" dimR] dim_eq by auto qed qed lemma (in cpx_sq_mat) sum_rank_1_proj_unitary: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" shows "(sum_mat (\i. rank_1_proj (Matrix.col A i)) {..< dimR})= (1\<^sub>m dimR)" proof show "dim_row (sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR)" using assms fc_mats_carrier by (metis carrier_matD(1) dim_col dim_eq index_one_mat(2) rank_1_proj_carrier sum_mat_carrier) show "dim_col (sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR)" using assms fc_mats_carrier rank_1_proj_carrier sum_mat_carrier by (metis carrier_matD(2) dim_col dim_eq index_one_mat(3) square_mat.simps square_mats) show "\i j. i < dim_row (1\<^sub>m dimR) \ j < dim_col (1\<^sub>m dimR) \ sum_mat (\i. rank_1_proj (Matrix.col A i)) {..m dimR $$ (i, j)" using assms sum_rank_1_proj_unitary_index by simp qed lemma (in cpx_sq_mat) rank_1_proj_unitary: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (1\<^sub>m dimR) $$ (j,k) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" proof - have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = Complex_Matrix.inner_prod (Matrix.col A j) (Matrix.col A k) \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)" using outer_prod_mult_outer_prod assms Matrix.col_dim local.fc_mats_carrier unfolding rank_1_proj_def by blast also have "... = (Complex_Matrix.adjoint A * A) $$ (j, k)\\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using inner_prod_adjoint_comp[of A dimR A] assms fc_mats_carrier dim_eq by simp also have "... = (1\<^sub>m dimR) $$ (j,k) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms unfolding Complex_Matrix.unitary_def by (metis add_commute assms(2) index_add_mat(2) index_one_mat(2) one_mem unitary_simps(1)) finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_unitary_ne: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" and "k < dimR" and "j\ k" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0\<^sub>m dimR dimR)" proof - have "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (outer_prod (Matrix.col A j) (Matrix.col A k))" by simp also have "... = dimR" using assms fc_mats_carrier dim_eq by auto finally have rw: "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" . have "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (outer_prod (Matrix.col A j) (Matrix.col A k))" by simp also have "... = dimR" using assms fc_mats_carrier dim_eq by auto finally have cl: "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dimR" . have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A k)) = (0::complex) \\<^sub>m (outer_prod (Matrix.col A j) (Matrix.col A k))" using assms rank_1_proj_unitary by simp also have "... = (0\<^sub>m dimR dimR)" proof show "dim_row (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_row (0\<^sub>m dimR dimR)" using rw by simp next show "dim_col (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) = dim_col (0\<^sub>m dimR dimR)" using cl by simp next show "\i p. i < dim_row (0\<^sub>m dimR dimR) \ p < dim_col (0\<^sub>m dimR dimR) \ (0 \\<^sub>m outer_prod (Matrix.col A j) (Matrix.col A k)) $$ (i, p) = 0\<^sub>m dimR dimR $$ (i, p)" using rw cl by auto qed finally show ?thesis . qed lemma (in cpx_sq_mat) rank_1_proj_unitary_eq: assumes "A \ fc_mats" and "Complex_Matrix.unitary A" and "j < dimR" shows "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = rank_1_proj (Matrix.col A j)" proof - have "rank_1_proj (Matrix.col A j) * (rank_1_proj (Matrix.col A j)) = (1::complex) \\<^sub>m (rank_1_proj (Matrix.col A j))" using assms rank_1_proj_unitary unfolding rank_1_proj_def by simp also have "... = (rank_1_proj (Matrix.col A j))" by (simp add: smult_one) finally show ?thesis . qed end \ No newline at end of file diff --git a/thys/Projective_Measurements/Projective_Measurements.thy b/thys/Projective_Measurements/Projective_Measurements.thy --- a/thys/Projective_Measurements/Projective_Measurements.thy +++ b/thys/Projective_Measurements/Projective_Measurements.thy @@ -1,1670 +1,1670 @@ (* Author: Mnacho Echenim, Université Grenoble Alpes *) theory Projective_Measurements imports Linear_Algebra_Complements begin section \Projective measurements\ text \In this part we define projective measurements, also refered to as von Neumann measurements. The latter are characterized by a set of orthogonal projectors, which are used to compute the probabilities of measure outcomes and to represent the state of the system after the measurement.\ text \The state of the system (a density operator in this case) after a measurement is represented by the \verb+density_collapse+ function.\ subsection \First definitions\ text \We begin by defining a type synonym for couples of measurement values (reals) and the associated projectors (complex matrices).\ type_synonym measure_outcome = "real \ complex Matrix.mat" text \The corresponding values and projectors are retrieved thanks to \verb+meas_outcome_val+ and \verb+meas_outcome_prj+.\ definition meas_outcome_val::"measure_outcome \ real" where "meas_outcome_val Mi = fst Mi" definition meas_outcome_prj::"measure_outcome \ complex Matrix.mat" where "meas_outcome_prj Mi = snd Mi" text \We define a predicate for projective measurements. A projective measurement is characterized by the number $p$ of possible measure outcomes, and a function $M$ mapping outcome $i$ to the corresponding \verb+measure_outcome+.\ definition (in cpx_sq_mat) proj_measurement::"nat \ (nat \ measure_outcome) \ bool" where "proj_measurement n M \ (inj_on (\i. meas_outcome_val (M i)) {..< n}) \ (\j < n. meas_outcome_prj (M j) \ fc_mats \ projector (meas_outcome_prj (M j))) \ (\ i < n. \ j < n. i \ j \ meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR) \ sum_mat (\j. meas_outcome_prj (M j)) {..< n} = 1\<^sub>m dimR" lemma (in cpx_sq_mat) proj_measurement_inj: assumes "proj_measurement p M" shows "inj_on (\i. meas_outcome_val (M i)) {..< p}" using assms unfolding proj_measurement_def by simp lemma (in cpx_sq_mat) proj_measurement_carrier: assumes "proj_measurement p M" and "i < p" shows "meas_outcome_prj (M i) \ fc_mats" using assms unfolding proj_measurement_def by simp lemma (in cpx_sq_mat) proj_measurement_ortho: assumes "proj_measurement p M" and "i < p" and "j < p" and "i\ j" shows "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR" using assms unfolding proj_measurement_def by simp lemma (in cpx_sq_mat) proj_measurement_id: assumes "proj_measurement p M" shows "sum_mat (\j. meas_outcome_prj (M j)) {..< p} = 1\<^sub>m dimR" using assms unfolding proj_measurement_def by simp lemma (in cpx_sq_mat) proj_measurement_square: assumes "proj_measurement p M" and "i < p" shows "meas_outcome_prj (M i) \ fc_mats" using assms unfolding proj_measurement_def by simp lemma (in cpx_sq_mat) proj_measurement_proj: assumes "proj_measurement p M" and "i < p" shows "projector (meas_outcome_prj (M i))" using assms unfolding proj_measurement_def by simp text \We define the probability of obtaining a measurement outcome: this is a positive number and the sum over all the measurement outcomes is 1.\ definition (in cpx_sq_mat) meas_outcome_prob :: "complex Matrix.mat \ (nat \ real \ complex Matrix.mat) \ nat \ complex" where "meas_outcome_prob R M i = Complex_Matrix.trace (R* (meas_outcome_prj (M i)))" lemma (in cpx_sq_mat) meas_outcome_prob_real: assumes "R\ fc_mats" and "density_operator R" and "proj_measurement n M" and "i < n" shows "meas_outcome_prob R M i \ \" proof - have "complex_of_real (Re (Complex_Matrix.trace (R * meas_outcome_prj (M i)))) = Complex_Matrix.trace (R * meas_outcome_prj (M i))" proof (rule trace_proj_pos_real) show "R \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp show "meas_outcome_prj (M i) \ carrier_mat dimR dimR" using assms proj_measurement_carrier fc_mats_carrier dim_eq by simp qed thus ?thesis unfolding meas_outcome_prob_def by (metis Reals_of_real) qed lemma (in cpx_sq_mat) meas_outcome_prob_pos: assumes "R\ fc_mats" and "density_operator R" and "proj_measurement n M" and "i < n" shows "0 \ meas_outcome_prob R M i" unfolding meas_outcome_prob_def proof (rule positive_proj_trace) show "R \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq by simp show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp show "projector (meas_outcome_prj (M i))" using assms proj_measurement_proj by simp show "meas_outcome_prj (M i) \ carrier_mat dimR dimR" using assms proj_measurement_carrier fc_mats_carrier dim_eq by simp qed lemma (in cpx_sq_mat) meas_outcome_prob_sum: assumes "density_operator R" and "R\ fc_mats" and" proj_measurement n M" shows "(\ j \ {..< n}. meas_outcome_prob R M j) = 1" proof - have "(\ j \ {..< n}. Complex_Matrix.trace (R* (meas_outcome_prj (M j)))) = Complex_Matrix.trace (sum_mat (\j. R* (meas_outcome_prj (M j))) {..< n})" proof (rule trace_sum_mat[symmetric], auto) fix j assume "j < n" thus "R * meas_outcome_prj (M j) \ fc_mats" using cpx_sq_mat_mult assms unfolding proj_measurement_def by simp qed also have "... = Complex_Matrix.trace (R * (sum_mat (\j. (meas_outcome_prj (M j))) {..< n}))" proof - have "sum_mat (\j. R* (meas_outcome_prj (M j))) {..< n} = R * (sum_mat (\j. (meas_outcome_prj (M j))) {..< n})" proof (rule mult_sum_mat_distrib_left, (auto simp add: assms)) fix j assume "j < n" thus "meas_outcome_prj (M j) \ fc_mats" using assms unfolding proj_measurement_def by simp qed thus ?thesis by simp qed also have "... = Complex_Matrix.trace (R * 1\<^sub>m dimR)" using assms unfolding proj_measurement_def by simp also have "... = Complex_Matrix.trace R" using assms by (simp add: fc_mats_carrier dim_eq) also have "... = 1" using assms unfolding density_operator_def by simp finally show ?thesis unfolding meas_outcome_prob_def . qed text \We introduce the maximally mixed density operator. Intuitively, this density operator corresponds to a uniform distribution of the states of an orthonormal basis. This operator will be used to define the density operator after a measurement for the measure outcome probabilities equal to zero.\ definition max_mix_density :: "nat \ complex Matrix.mat" where "max_mix_density n = ((1::real)/ n) \\<^sub>m (1\<^sub>m n)" lemma max_mix_density_carrier: shows "max_mix_density n \ carrier_mat n n" unfolding max_mix_density_def by simp lemma max_mix_is_density: assumes "0 < n" shows "density_operator (max_mix_density n)" unfolding density_operator_def max_mix_density_def proof have "Complex_Matrix.trace (complex_of_real ((1::real)/n) \\<^sub>m 1\<^sub>m n) = (complex_of_real ((1::real)/n)) * (Complex_Matrix.trace ((1\<^sub>m n)::complex Matrix.mat))" using one_carrier_mat trace_smult[of "(1\<^sub>m n)::complex Matrix.mat"] by blast also have "... = (complex_of_real ((1::real)/n)) * (real n)" using trace_1[of n] by simp also have "... = 1" using assms by simp finally show "Complex_Matrix.trace (complex_of_real ((1::real)/n) \\<^sub>m 1\<^sub>m n) = 1" . next show "Complex_Matrix.positive (complex_of_real (1 / real n) \\<^sub>m 1\<^sub>m n)" - by (rule positive_smult, (auto simp add: positive_one)) + by (rule positive_smult, (auto simp add: positive_one less_eq_complex_def)) qed lemma (in cpx_sq_mat) max_mix_density_square: shows "max_mix_density dimR \ fc_mats" unfolding max_mix_density_def using fc_mats_carrier dim_eq by simp text \Given a measurement outcome, \verb+density_collapse+ represents the resulting density operator. In practice only the measure outcomes with nonzero probabilities are of interest; we (arbitrarily) collapse the density operator for zero-probability outcomes to the maximally mixed density operator.\ definition density_collapse ::"complex Matrix.mat \ complex Matrix.mat \ complex Matrix.mat" where "density_collapse R P = (if ((Complex_Matrix.trace (R * P)) = 0) then (max_mix_density (dim_row R)) else ((1::real)/ ((Complex_Matrix.trace (R * P)))) \\<^sub>m (P * R * P))" lemma density_collapse_carrier: assumes "0 < dim_row R" and "P \ carrier_mat n n" and "R \ carrier_mat n n" shows "(density_collapse R P) \ carrier_mat n n" proof (cases "(Complex_Matrix.trace (R * P)) = 0") case True hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp then show ?thesis using max_mix_is_density assms max_mix_density_carrier by auto next case False hence "density_collapse R P = complex_of_real 1 / Complex_Matrix.trace (R * P) \\<^sub>m (P * R * P)" unfolding density_collapse_def by simp thus ?thesis using assms by auto qed lemma density_collapse_operator: assumes "projector P" and "density_operator R" and "0 < dim_row R" and "P \ carrier_mat n n" and "R \ carrier_mat n n" shows "density_operator (density_collapse R P)" proof (cases "(Complex_Matrix.trace (R * P)) = 0") case True hence "density_collapse R P = max_mix_density (dim_row R)" unfolding density_collapse_def by simp then show ?thesis using max_mix_is_density assms by simp next case False show ?thesis unfolding density_operator_def proof (intro conjI) have "Complex_Matrix.positive ((1 / (Complex_Matrix.trace (R * P))) \\<^sub>m (P * R * P))" proof (rule positive_smult) show "P * R * P \ carrier_mat n n" using assms by simp have "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp hence "0 \ (Complex_Matrix.trace (R * P))" using positive_proj_trace[of P R n] assms False by auto - hence "0 \ Re (Complex_Matrix.trace (R * P))" by simp + hence "0 \ Re (Complex_Matrix.trace (R * P))" by (simp add: less_eq_complex_def) hence "0 \ 1/(Re (Complex_Matrix.trace (R * P)))" by simp have "Re (Complex_Matrix.trace (R * P)) = Complex_Matrix.trace (R * P)" using assms \Complex_Matrix.positive R\ trace_proj_pos_real by simp hence inv: "1/ (Complex_Matrix.trace (R * P)) = 1/(Re (Complex_Matrix.trace (R * P)))" by simp thus "0 \ 1/ (Complex_Matrix.trace (R * P))" - using \0 \ 1/(Re (Complex_Matrix.trace (R * P)))\ by (simp add: inv) + using \0 \ 1/(Re (Complex_Matrix.trace (R * P)))\ by (simp add: inv less_eq_complex_def) show "Complex_Matrix.positive (P * R * P)" using assms positive_close_under_left_right_mult_adjoint[of P n R] by (simp add: \Complex_Matrix.positive R\ hermitian_def projector_def) qed thus "Complex_Matrix.positive (density_collapse R P)" using False unfolding density_collapse_def by simp next have "Complex_Matrix.trace (density_collapse R P) = Complex_Matrix.trace ((1/ (Complex_Matrix.trace (R * P))) \\<^sub>m (P * R * P))" using False unfolding density_collapse_def by simp also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (P * R * P)" using trace_smult[of "P * R * P" n] assms by simp also have "... = 1/ (Complex_Matrix.trace (R * P)) * Complex_Matrix.trace (R * P)" using projector_collapse_trace assms by simp also have "... = 1" using False by simp finally show "Complex_Matrix.trace (density_collapse R P) = 1" . qed qed subsection \Measurements with observables\ text \It is standard in quantum mechanics to represent projective measurements with so-called \emph{observables}. These are Hermitian matrices which encode projective measurements as follows: the eigenvalues of an observable represent the possible projective measurement outcomes, and the associated projectors are the projectors onto the corresponding eigenspaces. The results in this part are based on the spectral theorem, which states that any Hermitian matrix admits an orthonormal basis consisting of eigenvectors of the matrix.\ subsubsection \On the diagonal elements of a matrix\ text \We begin by introducing definitions that will be used on the diagonalized version of a Hermitian matrix.\ definition diag_elems where "diag_elems B = {B$$(i,i) |i. i < dim_row B}" text \Relationship between \verb+diag_elems+ and the list \verb+diag_mat+\ lemma diag_elems_set_diag_mat: shows "diag_elems B = set (diag_mat B)" unfolding diag_mat_def diag_elems_def proof show "{B $$ (i, i) |i. i < dim_row B} \ set (map (\i. B $$ (i, i)) [0.. {B $$ (i, i) |i. i < dim_row B}" hence "\i < dim_row B. x = B $$(i,i)" by auto from this obtain i where "i < dim_row B" and "x = B $$(i,i)" by auto thus "x \ set (map (\i. B $$ (i, i)) [0..i. B $$ (i, i)) [0.. {B $$ (i, i) |i. i < dim_row B}" proof fix x assume "x \ set (map (\i. B $$ (i, i)) [0.. {B $$ (i, i) |i. i < dim_row B}" by auto qed qed lemma diag_elems_finite[simp]: shows "finite (diag_elems B)" unfolding diag_elems_def by simp lemma diag_elems_mem[simp]: assumes "i < dim_row B" shows "B $$(i,i) \ diag_elems B" using assms unfolding diag_elems_def by auto text \When $x$ is a diagonal element of $B$, \verb+diag_elem_indices+ returns the set of diagonal indices of $B$ with value $x$.\ definition diag_elem_indices where "diag_elem_indices x B = {i|i. i < dim_row B \ B $$ (i,i) = x}" lemma diag_elem_indices_elem: assumes "a \ diag_elem_indices x B" shows "a < dim_row B \ B$$(a,a) = x" using assms unfolding diag_elem_indices_def by simp lemma diag_elem_indices_itself: assumes "i < dim_row B" shows "i \ diag_elem_indices (B $$(i,i)) B" using assms unfolding diag_elem_indices_def by simp lemma diag_elem_indices_finite: shows "finite (diag_elem_indices x B)" unfolding diag_elem_indices_def by simp text \We can therefore partition the diagonal indices of a matrix $B$ depending on the value of the diagonal elements. If $B$ admits $p$ elements on its diagonal, then we define bijections between its set of diagonal elements and the initial segment $[0..p-1]$.\ definition dist_el_card where "dist_el_card B = card (diag_elems B)" definition diag_idx_to_el where "diag_idx_to_el B = (SOME h. bij_betw h {..< dist_el_card B} (diag_elems B))" definition diag_el_to_idx where "diag_el_to_idx B = inv_into {..< dist_el_card B} (diag_idx_to_el B)" lemma diag_idx_to_el_bij: shows "bij_betw (diag_idx_to_el B) {..< dist_el_card B} (diag_elems B)" proof - let ?V = "SOME h. bij_betw h {..< dist_el_card B} (diag_elems B)" have vprop: "bij_betw ?V {..< dist_el_card B} (diag_elems B)" using someI_ex[of "\h. bij_betw h {..< dist_el_card B} (diag_elems B)"] diag_elems_finite unfolding dist_el_card_def using bij_betw_from_nat_into_finite by blast show ?thesis by (simp add:diag_idx_to_el_def vprop) qed lemma diag_el_to_idx_bij: shows "bij_betw (diag_el_to_idx B) (diag_elems B) {..< dist_el_card B}" unfolding diag_el_to_idx_def proof (rule bij_betw_inv_into_subset[of _ _ "diag_elems B"], (simp add: diag_idx_to_el_bij)+) show "diag_idx_to_el B ` {.. {..< dist_el_card B}" using assms by simp moreover have "j\ {..< dist_el_card B}" using assms by simp moreover have "inj_on (diag_idx_to_el B) {.. diag_elems B" shows "\k \ {..< dist_el_card B}. x = diag_idx_to_el B k" proof - have "diag_idx_to_el B ` {.. diag_elems B" proof - have "diag_idx_to_el B ` {..i < dim_row B. B$$(i,i) \ Reals" shows "diag_elems B \ Reals" proof fix x assume "x\ diag_elems B" hence "\i < dim_row B. x = B $$(i,i)" using assms unfolding diag_elems_def by auto from this obtain i where "i < dim_row B" "x = B $$ (i,i)" by auto thus "x \ Reals" using assms by simp qed lemma diag_elems_Re: fixes B::"complex Matrix.mat" assumes "\i < (dim_row B). B$$(i,i) \ Reals" shows "{Re x|x. x\ diag_elems B} = diag_elems B" proof show "{complex_of_real (Re x) |x. x \ diag_elems B} \ diag_elems B" proof fix x assume "x \ {complex_of_real (Re x) |x. x \ diag_elems B}" hence "\y \ diag_elems B. x = Re y" by auto from this obtain y where "y\ diag_elems B" and "x = Re y" by auto hence "y = x" using assms diag_elems_real[of B] by auto thus "x\ diag_elems B" using \y\ diag_elems B\ by simp qed show "diag_elems B \ {complex_of_real (Re x) |x. x \ diag_elems B}" proof fix x assume "x \ diag_elems B" hence "Re x = x" using assms diag_elems_real[of B] by auto thus "x\ {complex_of_real (Re x) |x. x \ diag_elems B}" using \x\ diag_elems B\ by force qed qed lemma diag_idx_to_el_real: fixes B::"complex Matrix.mat" assumes "\i < dim_row B. B$$(i,i) \ Reals" and "i < dist_el_card B" shows "Re (diag_idx_to_el B i) = diag_idx_to_el B i" proof - have "diag_idx_to_el B i \ diag_elems B" using diag_idx_to_el_img[of i B] assms by simp hence "diag_idx_to_el B i \ Reals" using diag_elems_real[of B] assms by auto thus ?thesis by simp qed lemma diag_elem_indices_empty: assumes "B \ carrier_mat dimR dimC" and "i < (dist_el_card B)" and "j < (dist_el_card B)" and "i\ j" shows "diag_elem_indices (diag_idx_to_el B i) B \ (diag_elem_indices (diag_idx_to_el B j) B) = {}" proof (rule ccontr) assume "diag_elem_indices (diag_idx_to_el B i) B \ diag_elem_indices (diag_idx_to_el B j) B \ {}" hence "\x. x\ diag_elem_indices (diag_idx_to_el B i) B \ diag_elem_indices (diag_idx_to_el B j) B" by auto from this obtain x where xprop: "x\ diag_elem_indices (diag_idx_to_el B i) B \ diag_elem_indices (diag_idx_to_el B j) B" by auto hence "B $$ (x,x) = (diag_idx_to_el B i)" using diag_elem_indices_elem[of x "diag_idx_to_el B i"] by simp moreover have "B $$ (x,x) = (diag_idx_to_el B j)" using diag_elem_indices_elem[of x "diag_idx_to_el B j"] xprop by simp ultimately have "diag_idx_to_el B i = diag_idx_to_el B j" by simp hence "i = j" using diag_idx_to_el_less_inj assms by auto thus False using assms by simp qed lemma (in cpx_sq_mat) diag_elem_indices_disjoint: assumes "B\ carrier_mat dimR dimC" shows "disjoint_family_on (\n. diag_elem_indices (diag_idx_to_el B n) B) {.. {..< dist_el_card B}" and "p\ {..< dist_el_card B}" and "m\ p" thus "diag_elem_indices (diag_idx_to_el B m) B \ diag_elem_indices (diag_idx_to_el B p) B = {}" using diag_elem_indices_empty assms fc_mats_carrier by simp qed lemma diag_elem_indices_union: assumes "B\ carrier_mat dimR dimC" shows "(\ i \ {..< dist_el_card B}. diag_elem_indices (diag_idx_to_el B i) B) = {..< dimR}" proof show "(\i {.. (\ii < dist_el_card B. x\ diag_elem_indices (diag_idx_to_el B i) B" by auto from this obtain i where "i < dist_el_card B" "x\ diag_elem_indices (diag_idx_to_el B i) B" by auto hence "x < dimR" using diag_elem_indices_elem[of x _ B] assms by simp thus "x \ {..< dimR}" by auto qed next show "{.. (\i {..< dimR}" hence "j < dim_row B" using assms by simp hence "B$$(j,j) \ diag_elems B" by simp hence "\k \ {..< dist_el_card B}. B$$(j,j) = diag_idx_to_el B k" using diag_idx_to_el_less_surj[of "B $$(j,j)"] by simp from this obtain k where kprop: "k \ {..< dist_el_card B}" "B$$(j,j) = diag_idx_to_el B k" by auto hence "j \ diag_elem_indices (diag_idx_to_el B k) B" using \j < dim_row B\ diag_elem_indices_itself by fastforce thus "j \ (\iConstruction of measurement outcomes\ text \The construction of a projective measurement for a hermitian matrix $A$ is based on the Schur decomposition $A = U*B*U^\dagger$, where $B$ is diagonal and $U$ is unitary. The columns of $U$ are normalized and pairwise orthogonal; they are used to construct the projectors associated with a measurement value\ definition (in cpx_sq_mat) project_vecs where "project_vecs (f::nat \ complex Matrix.vec) N = sum_mat (\i. rank_1_proj (f i)) N" lemma (in cpx_sq_mat) project_vecs_dim: assumes "\i \ N. dim_vec (f i) = dimR" shows "project_vecs f N \ fc_mats" proof - have "project_vecs f N \ carrier_mat dimR dimC" unfolding project_vecs_def proof (rule sum_mat_carrier) show "\i. i \ N \ rank_1_proj (f i) \ fc_mats" using assms fc_mats_carrier rank_1_proj_dim dim_eq rank_1_proj_carrier by fastforce qed thus ?thesis using fc_mats_carrier by simp qed definition (in cpx_sq_mat) mk_meas_outcome where "mk_meas_outcome B U = (\i. (Re (diag_idx_to_el B i), project_vecs (\i. zero_col U i) (diag_elem_indices (diag_idx_to_el B i) B)))" lemma (in cpx_sq_mat) mk_meas_outcome_carrier: assumes "Complex_Matrix.unitary U" and "U \ fc_mats" and "B\ fc_mats" shows "meas_outcome_prj ((mk_meas_outcome B U) j) \ fc_mats" proof - have "project_vecs (\i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B) \ fc_mats" using project_vecs_dim by (simp add: assms(2) zero_col_dim) thus ?thesis unfolding mk_meas_outcome_def meas_outcome_prj_def by simp qed lemma (in cpx_sq_mat) mk_meas_outcome_sum_id: assumes "Complex_Matrix.unitary U" and "U \ fc_mats" and "B\ fc_mats" shows "sum_mat (\j. meas_outcome_prj ((mk_meas_outcome B U) j)) {..<(dist_el_card B)} = 1\<^sub>m dimR" proof - have "sum_mat (\j. meas_outcome_prj ((mk_meas_outcome B U) j)) {..<(dist_el_card B)} = sum_mat (\j. project_vecs (\i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B)) {..<(dist_el_card B)}" unfolding mk_meas_outcome_def meas_outcome_prj_def by simp also have "... = sum_mat (\i. rank_1_proj (zero_col U i)) (\jj. rank_1_proj (zero_col U j) \ fc_mats" using zero_col_dim fc_mats_carrier dim_eq by (metis assms(2) rank_1_proj_carrier) show "finite {..i. i \ {.. finite (diag_elem_indices (diag_idx_to_el B i) B)" using diag_elem_indices_finite by simp show "disjoint_family_on (\n. diag_elem_indices (diag_idx_to_el B n) B) {.. {..< dist_el_card B}" and "p\ {..< dist_el_card B}" and "m\ p" thus "diag_elem_indices (diag_idx_to_el B m) B \ diag_elem_indices (diag_idx_to_el B p) B = {}" using diag_elem_indices_empty assms fc_mats_carrier by simp qed qed also have "... = sum_mat (\i. rank_1_proj (zero_col U i)) {..< dimR}" using diag_elem_indices_union[of B] assms fc_mats_carrier by simp also have "... = sum_mat (\i. rank_1_proj (Matrix.col U i)) {..< dimR}" proof (rule sum_mat_cong, simp) show "\i. i \ {.. rank_1_proj (zero_col U i) \ fc_mats" using dim_eq by (metis assms(2) local.fc_mats_carrier rank_1_proj_carrier zero_col_dim) thus "\i. i \ {.. rank_1_proj (Matrix.col U i) \ fc_mats" using dim_eq by (metis lessThan_iff zero_col_col) show "\i. i \ {.. rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)" by (simp add: zero_col_col) qed also have "... = 1\<^sub>m dimR" using sum_rank_1_proj_unitary assms by simp finally show ?thesis . qed lemma (in cpx_sq_mat) make_meas_outcome_prj_ortho: assumes "Complex_Matrix.unitary U" and "U \ fc_mats" and "B\ fc_mats" and "i < dist_el_card B" and "j < dist_el_card B" and "i \ j" shows "meas_outcome_prj ((mk_meas_outcome B U) i) * meas_outcome_prj ((mk_meas_outcome B U) j) = 0\<^sub>m dimR dimR" proof - define Pi where "Pi = sum_mat (\k. rank_1_proj (zero_col U k)) (diag_elem_indices (diag_idx_to_el B i) B)" have sneqi: "meas_outcome_prj (mk_meas_outcome B U i) = Pi" unfolding mk_meas_outcome_def project_vecs_def Pi_def meas_outcome_prj_def by simp define Pj where "Pj = sum_mat (\k. rank_1_proj (zero_col U k)) (diag_elem_indices (diag_idx_to_el B j) B)" have sneqj: "meas_outcome_prj (mk_meas_outcome B U j) = Pj" unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp have "Pi * Pj = 0\<^sub>m dimR dimR" unfolding Pi_def proof (rule sum_mat_left_ortho_zero) show "finite (diag_elem_indices (diag_idx_to_el B i) B)" using diag_elem_indices_finite[of _ B] by simp show km: "\k. k \ diag_elem_indices (diag_idx_to_el B i) B \ rank_1_proj (zero_col U k) \ fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq by (metis rank_1_proj_carrier) show "Pj \ fc_mats" using sneqj assms mk_meas_outcome_carrier by auto show "\k. k \ diag_elem_indices (diag_idx_to_el B i) B \ rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR" proof - fix k assume "k \ diag_elem_indices (diag_idx_to_el B i) B" show "rank_1_proj (zero_col U k) * Pj = 0\<^sub>m dimR dimR" unfolding Pj_def proof (rule sum_mat_right_ortho_zero) show "finite (diag_elem_indices (diag_idx_to_el B j) B)" using diag_elem_indices_finite[of _ B] by simp show "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ rank_1_proj (zero_col U i) \ fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq by (metis rank_1_proj_carrier) show "rank_1_proj (zero_col U k) \ fc_mats" by (simp add: km \k \ diag_elem_indices (diag_idx_to_el B i) B\) show "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ rank_1_proj (zero_col U k) * rank_1_proj (zero_col U i) = 0\<^sub>m dimR dimR" proof - fix m assume "m \ diag_elem_indices (diag_idx_to_el B j) B" hence "m \ k" using \k \ diag_elem_indices (diag_idx_to_el B i) B\ diag_elem_indices_disjoint[of B] fc_mats_carrier assms unfolding disjoint_family_on_def by auto have "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ i < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1)) hence "m < dimR" using \m \ diag_elem_indices (diag_idx_to_el B j) B\ by simp have "\k. k \ diag_elem_indices (diag_idx_to_el B i) B \ k < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1)) hence "k < dimR" using \k \ diag_elem_indices (diag_idx_to_el B i) B\ by simp show "rank_1_proj (zero_col U k) * rank_1_proj (zero_col U m) = 0\<^sub>m dimR dimR" using rank_1_proj_unitary_ne[of U k m] assms \m < dimR\ \k < dimR\ by (metis \m \ k\ zero_col_col) qed qed qed qed thus ?thesis using sneqi sneqj by simp qed lemma (in cpx_sq_mat) make_meas_outcome_prjectors: assumes "Complex_Matrix.unitary U" and "U \ fc_mats" and "B\ fc_mats" and "j < dist_el_card B" shows "projector (meas_outcome_prj ((mk_meas_outcome B U) j))" unfolding projector_def proof define Pj where "Pj = sum_mat (\i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B)" have sneq: "meas_outcome_prj (mk_meas_outcome B U j) = Pj" unfolding mk_meas_outcome_def project_vecs_def Pj_def meas_outcome_prj_def by simp moreover have "hermitian Pj" unfolding Pj_def proof (rule sum_mat_hermitian) show "finite (diag_elem_indices (diag_idx_to_el B j) B)" using diag_elem_indices_finite[of _ B] by simp show "\i\diag_elem_indices (diag_idx_to_el B j) B. hermitian (rank_1_proj (zero_col U i))" using rank_1_proj_hermitian by simp show "\i\diag_elem_indices (diag_idx_to_el B j) B. rank_1_proj (zero_col U i) \ fc_mats" using zero_col_dim fc_mats_carrier dim_eq by (metis assms(2) rank_1_proj_carrier) qed ultimately show "hermitian (meas_outcome_prj (mk_meas_outcome B U j))" by simp have "Pj * Pj = Pj" unfolding Pj_def proof (rule sum_mat_ortho_square) show "finite (diag_elem_indices (diag_idx_to_el B j) B)" using diag_elem_indices_finite[of _ B] by simp show "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) = rank_1_proj (zero_col U i)" proof - fix i assume imem: "i\ diag_elem_indices (diag_idx_to_el B j) B" hence "i < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1)) hence "zero_col U i = Matrix.col U i" using zero_col_col[of i] by simp hence "rank_1_proj (zero_col U i) = rank_1_proj (Matrix.col U i)" by simp moreover have "rank_1_proj (Matrix.col U i) * rank_1_proj (Matrix.col U i) = rank_1_proj (Matrix.col U i)" by (rule rank_1_proj_unitary_eq, (auto simp add: assms \i < dimR\)) ultimately show "rank_1_proj (zero_col U i) * rank_1_proj (zero_col U i) = rank_1_proj (zero_col U i)" by simp qed show "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ rank_1_proj (zero_col U i) \ fc_mats" using zero_col_dim assms fc_mats_carrier dim_eq by (metis rank_1_proj_carrier) have "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ i < dimR" using diag_elem_indices_elem fc_mats_carrier assms dim_eq by (metis carrier_matD(1)) thus "\i ja. i \ diag_elem_indices (diag_idx_to_el B j) B \ ja \ diag_elem_indices (diag_idx_to_el B j) B \ i \ ja \ rank_1_proj (zero_col U i) * rank_1_proj (zero_col U ja) = 0\<^sub>m dimR dimR" using rank_1_proj_unitary_ne by (simp add: assms(1) assms(2) zero_col_col) qed thus "meas_outcome_prj (mk_meas_outcome B U j) * meas_outcome_prj (mk_meas_outcome B U j) = meas_outcome_prj (mk_meas_outcome B U j)" using sneq by simp qed lemma (in cpx_sq_mat) mk_meas_outcome_fst_inj: assumes "\i < (dim_row B). B$$(i,i) \ Reals" shows "inj_on (\i. meas_outcome_val ((mk_meas_outcome B U) i)) {.. {.. {..i < (dim_row B). B$$(i,i) \ Reals" shows "bij_betw (\i. meas_outcome_val ((mk_meas_outcome B U) i)) {..< dist_el_card B} {Re x|x. x\ diag_elems B}" unfolding bij_betw_def proof have "inj_on (\x. (meas_outcome_val (mk_meas_outcome B U x))) {.. diag_elems B} = diag_elems B" using diag_elems_Re[of B] assms by simp ultimately show "inj_on (\x. meas_outcome_val (mk_meas_outcome B U x)) {..i. meas_outcome_val (mk_meas_outcome B U i)) ` {.. diag_elems B}" unfolding meas_outcome_val_def mk_meas_outcome_def proof show "(\i. fst (Re (diag_idx_to_el B i), project_vecs (zero_col U) (diag_elem_indices (diag_idx_to_el B i) B))) ` {.. {Re x |x. x \ diag_elems B}" using diag_idx_to_el_bij[of B] bij_betw_apply by fastforce show "{Re x |x. x \ diag_elems B} \ (\i. fst (Re (diag_idx_to_el B i), project_vecs (zero_col U) (diag_elem_indices (diag_idx_to_el B i) B))) ` {..Projective measurement associated with an observable\ definition eigvals where "eigvals M = (SOME as. char_poly M = (\a\as. [:- a, 1:]) \ length as = dim_row M)" lemma eigvals_poly_length: assumes "(M::complex Matrix.mat) \ carrier_mat n n" shows "char_poly M = (\a\(eigvals M). [:- a, 1:]) \ length (eigvals M) = dim_row M" proof - let ?V = "SOME as. char_poly M = (\a\as. [:- a, 1:]) \ length as = dim_row M" have vprop: "char_poly M = (\a\?V. [:- a, 1:]) \ length ?V = dim_row M" using someI_ex[of "\as. char_poly M = (\a\as. [:- a, 1:]) \ length as = dim_row M"] complex_mat_char_poly_factorizable assms by blast show ?thesis by (metis (no_types) eigvals_def vprop) qed text \We define the spectrum of a matrix $A$: this is its set of eigenvalues; its elements are roots of the characteristic polynomial of $A$.\ definition spectrum where "spectrum M = set (eigvals M)" lemma spectrum_finite: shows "finite (spectrum M)" unfolding spectrum_def by simp lemma spectrum_char_poly_root: fixes A::"complex Matrix.mat" assumes "A\ carrier_mat n n" and "k \ spectrum A" shows "poly (char_poly A) k = 0" using eigvals_poly_length[of A n] assms unfolding spectrum_def eigenvalue_root_char_poly by (simp add: linear_poly_root) lemma spectrum_eigenvalues: fixes A::"complex Matrix.mat" assumes "A\ carrier_mat n n" and "k\ spectrum A" shows "eigenvalue A k" using eigenvalue_root_char_poly[of A n k] spectrum_char_poly_root[of A n k] assms by simp text \The main result that is used to construct a projective measurement for a Hermitian matrix is that it is always possible to decompose it as $A = U*B*U^\dagger$, where $B$ is diagonal and only contains real elements, and $U$ is unitary.\ definition hermitian_decomp where "hermitian_decomp A B U \ similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dim_row B. B$$(i, i) \ Reals)" lemma hermitian_decomp_sim: assumes "hermitian_decomp A B U" shows "similar_mat_wit A B U (Complex_Matrix.adjoint U)" using assms unfolding hermitian_decomp_def by simp lemma hermitian_decomp_diag_mat: assumes "hermitian_decomp A B U" shows "diagonal_mat B" using assms unfolding hermitian_decomp_def by simp lemma hermitian_decomp_eigenvalues: assumes "hermitian_decomp A B U" shows "diag_mat B = (eigvals A)" using assms unfolding hermitian_decomp_def by simp lemma hermitian_decomp_unitary: assumes "hermitian_decomp A B U" shows "unitary U" using assms unfolding hermitian_decomp_def by simp lemma hermitian_decomp_real_eigvals: assumes "hermitian_decomp A B U" shows "\i < dim_row B. B$$(i, i) \ Reals" using assms unfolding hermitian_decomp_def by simp lemma hermitian_decomp_dim_carrier: assumes "hermitian_decomp A B U" shows "B \ carrier_mat (dim_row A) (dim_col A)" using assms unfolding hermitian_decomp_def similar_mat_wit_def by (metis (full_types) index_mult_mat(3) index_one_mat(3) insert_subset) lemma similar_mat_wit_dim_row: assumes "similar_mat_wit A B Q R" shows "dim_row B = dim_row A" using assms Let_def unfolding similar_mat_wit_def by (meson carrier_matD(1) insert_subset) lemma (in cpx_sq_mat) hermitian_schur_decomp: assumes "hermitian A" and "A\ fc_mats" obtains B U where "hermitian_decomp A B U" proof - { have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)") hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A] pr by auto ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp hence "\ B U. hermitian_decomp A B U" by auto } thus ?thesis using that by auto qed lemma (in cpx_sq_mat) hermitian_spectrum_real: assumes "A \ fc_mats" and "hermitian A" and "a \ spectrum A" shows "a \ Reals" proof - obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto hence dimB: "B \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq hermitian_decomp_dim_carrier[of A] by simp hence Bii: "\i. i < dimR \ B$$(i, i) \ Reals" using hermitian_decomp_real_eigvals[of A B U] bu assms fc_mats_carrier by simp have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp hence "a \ set (diag_mat B)" using assms unfolding spectrum_def by simp hence "\i < length (diag_mat B). a = diag_mat B ! i" by (metis in_set_conv_nth) from this obtain i where "i < length (diag_mat B)" and "a = diag_mat B ! i" by auto hence "a = B $$ (i,i)" unfolding diag_mat_def by simp moreover have "i < dimR" using dimB \i < length (diag_mat B)\ unfolding diag_mat_def by simp ultimately show ?thesis using Bii by simp qed lemma (in cpx_sq_mat) spectrum_ne: assumes "A\ fc_mats" and "hermitian A" shows "spectrum A \ {}" proof - obtain B U where bu: "hermitian_decomp A B U" using assms hermitian_schur_decomp by auto hence dimB: "B \ carrier_mat dimR dimR" using assms fc_mats_carrier dim_eq hermitian_decomp_dim_carrier[of A] by simp have "diag_mat B = (eigvals A)" using bu hermitian_decomp_eigenvalues[of A B] by simp have "B$$(0,0) \ set (diag_mat B)" using dimB npos unfolding diag_mat_def by simp hence "set (diag_mat B) \ {}" by auto thus ?thesis unfolding spectrum_def using \diag_mat B = eigvals A\ by auto qed lemma unitary_hermitian_eigenvalues: fixes U::"complex Matrix.mat" assumes "unitary U" and "hermitian U" and "U \ carrier_mat n n" and "0 < n" and "k \ spectrum U" shows "k \ {-1, 1}" proof - have "cpx_sq_mat n n (carrier_mat n n)" by (unfold_locales, (simp add: assms)+) have "eigenvalue U k" using spectrum_eigenvalues assms by simp have "k \ Reals" using assms \cpx_sq_mat n n (carrier_mat n n)\ cpx_sq_mat.hermitian_spectrum_real by simp hence "conjugate k = k" by (simp add: Reals_cnj_iff) hence "k\<^sup>2 = 1" using unitary_eigenvalues_norm_square[of U n k] assms by (simp add: \eigenvalue U k\ power2_eq_square) thus ?thesis using power2_eq_1_iff by auto qed lemma unitary_hermitian_Re_spectrum: fixes U::"complex Matrix.mat" assumes "unitary U" and "hermitian U" and "U \ carrier_mat n n" and "0 < n" shows "{Re x|x. x\ spectrum U} \ {-1,1}" proof fix y assume "y\ {Re x|x. x\ spectrum U}" hence "\x \ spectrum U. y = Re x" by auto from this obtain x where "x\ spectrum U" and "y = Re x" by auto hence "x\ {-1,1}" using unitary_hermitian_eigenvalues assms by simp thus "y \ {-1, 1}" using \y = Re x\ by auto qed text \The projective measurement associated with matrix $M$ is obtained from its Schur decomposition, by considering the number of distinct elements on the resulting diagonal matrix and constructing the projectors associated with each of them.\ type_synonym proj_meas_rep = "nat \ (nat \ measure_outcome)" definition proj_meas_size::"proj_meas_rep \ nat" where "proj_meas_size P = fst P" definition proj_meas_outcomes::"proj_meas_rep \ (nat \ measure_outcome)" where "proj_meas_outcomes P = snd P" definition (in cpx_sq_mat) make_pm::"complex Matrix.mat \ proj_meas_rep" where "make_pm A = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in (dist_el_card B, mk_meas_outcome B U))" lemma (in cpx_sq_mat) make_pm_decomp: shows "make_pm A = (proj_meas_size (make_pm A), proj_meas_outcomes (make_pm A))" unfolding proj_meas_size_def proj_meas_outcomes_def by simp lemma (in cpx_sq_mat) make_pm_proj_measurement: assumes "A \ fc_mats" and "hermitian A" and "make_pm A = (n, M)" shows "proj_measurement n M" proof - have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)", auto) then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B" and Bii: "\i. i < dimR \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat dimR dimR" and dimP: "U \ carrier_mat dimR dimR" and dimaP: "Complex_Matrix.adjoint U \ carrier_mat dimR dimR" and unit: "unitary U" unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def unfolding make_pm_def by simp hence "n = dist_el_card B" using assms by simp have "M = mk_meas_outcome B U" using assms mpeq by simp show ?thesis unfolding proj_measurement_def proof (intro conjI) show "inj_on (\i. meas_outcome_val (M i)) {..n = dist_el_card B\ \M = mk_meas_outcome B U\ Bii fc_mats_carrier dimB by auto show "\j fc_mats \ projector (meas_outcome_prj (M j))" proof (intro allI impI conjI) fix j assume "j < n" show "meas_outcome_prj (M j) \ fc_mats" using \M = mk_meas_outcome B U\ assms fc_mats_carrier \j < n\ \n = dist_el_card B\ dim_eq mk_meas_outcome_carrier dimB dimP unit by blast show "projector (meas_outcome_prj (M j))" using make_meas_outcome_prjectors \M = mk_meas_outcome B U\ fc_mats_carrier \n = dist_el_card B\ unit \j < n\ dimB dimP dim_eq by blast qed show "\ij j \ meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR" proof (intro allI impI) fix i fix j assume "i < n" and "j < n" and "i\ j" thus "meas_outcome_prj (M i) * meas_outcome_prj (M j) = 0\<^sub>m dimR dimR" using make_meas_outcome_prj_ortho[of U B i j] assms dimB dimP fc_mats_carrier dim_eq by (simp add: \M = mk_meas_outcome B U\ \n = dist_el_card B\ unit) qed show "sum_mat (\j. meas_outcome_prj (M j)) {..m dimR" using mk_meas_outcome_sum_id \M = mk_meas_outcome B U\ fc_mats_carrier dim_eq \n = dist_el_card B\ unit dimB dimP by blast qed qed lemma (in cpx_sq_mat) make_pm_proj_measurement': assumes "A\ fc_mats" and "hermitian A" shows "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))" unfolding proj_meas_size_def proj_meas_outcomes_def by (rule make_pm_proj_measurement[of A], (simp add: assms)+) lemma (in cpx_sq_mat) make_pm_projectors: assumes "A\ fc_mats" and "hermitian A" and "i < proj_meas_size (make_pm A)" shows "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) i))" proof - have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))" using assms make_pm_proj_measurement' by simp thus ?thesis using proj_measurement_proj assms by simp qed lemma (in cpx_sq_mat) make_pm_square: assumes "A\ fc_mats" and "hermitian A" and "i < proj_meas_size (make_pm A)" shows "meas_outcome_prj (proj_meas_outcomes (make_pm A) i) \ fc_mats" proof - have "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))" using assms make_pm_proj_measurement' by simp thus ?thesis using proj_measurement_square assms by simp qed subsubsection \Properties on the spectrum of a Hermitian matrix\ lemma (in cpx_sq_mat) hermitian_schur_decomp': assumes "hermitian A" and "A\ fc_mats" obtains B U where "hermitian_decomp A B U \ make_pm A = (dist_el_card B, mk_meas_outcome B U)" proof - { have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)") hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto moreover have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A] pr by auto ultimately have "hermitian_decomp A B U" unfolding hermitian_decomp_def by simp moreover have "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def unfolding make_pm_def hermitian_decomp_def by simp ultimately have "\ B U. hermitian_decomp A B U \ make_pm A = (dist_el_card B, mk_meas_outcome B U)" by auto } thus ?thesis using that by auto qed lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq: assumes "hermitian A" and "A \ fc_mats" and "make_pm A = (p, M)" shows "spectrum A = (\i. meas_outcome_val (M i)) `{..< p}" proof - obtain B U where bu: "hermitian_decomp A B U \ make_pm A = (dist_el_card B, mk_meas_outcome B U)" using assms hermitian_schur_decomp'[OF assms(1)] by auto hence "p = dist_el_card B" using assms by simp have dimB: "B \ carrier_mat dimR dimR" using hermitian_decomp_dim_carrier[of A] dim_eq bu assms fc_mats_carrier by auto have Bvals: "diag_mat B = eigvals A" using bu hermitian_decomp_eigenvalues[of A B U] by simp have Bii: "\i. i < dimR \ B$$(i, i) \ Reals" using bu hermitian_decomp_real_eigvals[of A B U] dimB by simp have "diag_elems B = set (eigvals A)" using dimB Bvals diag_elems_set_diag_mat[of B] by simp have "M = mk_meas_outcome B U" using assms bu by simp { fix i assume "i < p" have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)" using \M = mk_meas_outcome B U\ unfolding meas_outcome_val_def mk_meas_outcome_def by simp also have "... = diag_idx_to_el B i" using diag_idx_to_el_real dimB \i < p\ \p = dist_el_card B\ Bii by simp finally have "meas_outcome_val (M i) = diag_idx_to_el B i" . } note eq = this have "bij_betw (diag_idx_to_el B) {..p = dist_el_card B\ unfolding bij_betw_def by simp thus ?thesis using eq \diag_elems B = set (eigvals A)\ unfolding spectrum_def by auto qed lemma (in cpx_sq_mat) spectrum_meas_outcome_val_eq': assumes "hermitian A" and "A \ fc_mats" and "make_pm A = (p, M)" shows "{Re x|x. x\ spectrum A} = (\i. meas_outcome_val (M i)) `{..< p}" proof show "{Re x |x. x \ spectrum A} \ (\i. meas_outcome_val (M i)) ` {..i. meas_outcome_val (M i)) ` {.. {Re x |x. x \ spectrum A}" using spectrum_meas_outcome_val_eq assms by force qed lemma (in cpx_sq_mat) make_pm_eigenvalues: assumes "A\ fc_mats" and "hermitian A" and "i < proj_meas_size (make_pm A)" shows "meas_outcome_val (proj_meas_outcomes (make_pm A) i) \ spectrum A" using spectrum_meas_outcome_val_eq[of A] assms make_pm_decomp by auto lemma (in cpx_sq_mat) make_pm_spectrum: assumes "A\ fc_mats" and "hermitian A" and "make_pm A = (p,M)" shows "(\i. meas_outcome_val (proj_meas_outcomes (make_pm A) i)) ` {..< p} = spectrum A" proof show "(\x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {.. spectrum A" using assms make_pm_eigenvalues make_pm_decomp by (metis (mono_tags, lifting) Pair_inject image_subsetI lessThan_iff) show "spectrum A \ (\x. complex_of_real (meas_outcome_val (proj_meas_outcomes (make_pm A) x))) ` {.. fc_mats" and "make_pm A = (p, M)" shows "p = card (spectrum A)" proof - obtain B U where bu: "hermitian_decomp A B U \ make_pm A = (dist_el_card B, mk_meas_outcome B U)" using assms hermitian_schur_decomp'[OF assms(1)] by auto hence "p = dist_el_card B" using assms by simp have "spectrum A = set (diag_mat B)" using bu hermitian_decomp_eigenvalues[of A B U] unfolding spectrum_def by simp also have "... = diag_elems B" using diag_elems_set_diag_mat[of B] by simp finally have "spectrum A = diag_elems B" . thus ?thesis using \p = dist_el_card B\ unfolding dist_el_card_def by simp qed lemma (in cpx_sq_mat) spectrum_size': assumes "hermitian A" and "A\ fc_mats" shows "proj_meas_size (make_pm A) = card (spectrum A)" using spectrum_size unfolding proj_meas_size_def by (metis assms fst_conv surj_pair) lemma (in cpx_sq_mat) make_pm_projectors': assumes "hermitian A" and "A \ fc_mats" and "a fc_mats" and "a fc_mats" proof (rule proj_measurement_square) show "proj_measurement (proj_meas_size (make_pm A)) (proj_meas_outcomes (make_pm A))" using make_pm_proj_measurement' assms by simp show "a < proj_meas_size (make_pm A)" using assms spectrum_size[of _ "proj_meas_size (make_pm A)"] make_pm_decomp[of A] by simp qed lemma (in cpx_sq_mat) meas_outcome_prj_trace_real: assumes "hermitian A" and "density_operator R" and "R \ fc_mats" and "A\ fc_mats" and "a carrier_mat dimR dimR" using fc_mats_carrier assms dim_eq by simp show "Complex_Matrix.positive R" using assms unfolding density_operator_def by simp show "projector (meas_outcome_prj (proj_meas_outcomes (make_pm A) a))" using assms make_pm_projectors' by simp show "meas_outcome_prj (proj_meas_outcomes (make_pm A) a) \ carrier_mat dimR dimR" using meas_outcome_prj_carrier assms dim_eq fc_mats_carrier by simp qed lemma (in cpx_sq_mat) sum_over_spectrum: assumes "A\ fc_mats" and "hermitian A" and "make_pm A = (p, M)" shows "(\ y \ spectrum A. f y) = (\ i < p. f (meas_outcome_val (M i)))" proof (rule sum.reindex_cong) show "spectrum A =(\x. (meas_outcome_val (M x)))` {..< p}" using spectrum_meas_outcome_val_eq assms by simp show "inj_on (\x. complex_of_real (meas_outcome_val (M x))) {..x. (meas_outcome_val (M x))) {.. fc_mats" and "hermitian A" and "make_pm A = (p, M)" shows "(\ y \ {Re x|x. x \ spectrum A}. f y) = (\ i < p. f (meas_outcome_val (M i)))" proof (rule sum.reindex_cong) show "{Re x|x. x \ spectrum A} =(\x. (meas_outcome_val (M x)))` {..< p}" using spectrum_meas_outcome_val_eq' assms by simp show "inj_on (\x. (meas_outcome_val (M x))) {..When a matrix $A$ is decomposed into a projective measurement $\{(\lambda_a, \pi_a)\}$, it can be recovered by the formula $A = \sum \lambda_a \pi_a$.\ lemma (in cpx_sq_mat) make_pm_sum: assumes "A \ fc_mats" and "hermitian A" and "make_pm A = (p, M)" shows "sum_mat (\i. (meas_outcome_val (M i)) \\<^sub>m meas_outcome_prj (M i)) {..< p} = A" proof - have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)", auto) then have "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto hence A: "A = U * B * (Complex_Matrix.adjoint U)" and dB: "diagonal_mat B" and Bii: "\i. i < dimR \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat dimR dimR" and dimP: "U \ carrier_mat dimR dimR" and dimaP: "Complex_Matrix.adjoint U \ carrier_mat dimR dimR" and unit: "unitary U" unfolding similar_mat_wit_def Let_def using assms fc_mats_carrier by auto hence mpeq: "make_pm A = (dist_el_card B, mk_meas_outcome B U)" using us Let_def unfolding make_pm_def by simp hence "p = dist_el_card B" using assms by simp have "M = mk_meas_outcome B U" using assms mpeq by simp have "sum_mat (\i. meas_outcome_val (M i) \\<^sub>m meas_outcome_prj (M i)) {..< p} = sum_mat (\j. Re (diag_idx_to_el B j)\\<^sub>m project_vecs (\i. zero_col U i) (diag_elem_indices (diag_idx_to_el B j) B)) {..<(dist_el_card B)}" using \p = dist_el_card B\ \M = mk_meas_outcome B U\ unfolding meas_outcome_val_def meas_outcome_prj_def mk_meas_outcome_def by simp also have "... = sum_mat (\j. sum_mat (\i. (Re (diag_idx_to_el B j)) \\<^sub>m rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B)) {..i. i \ {.. complex_of_real (Re (diag_idx_to_el B i)) \\<^sub>m sum_mat (\i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B i) B) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by auto show "\i. i \ {.. sum_mat (\ia. complex_of_real (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier) show "\j. j \ {.. (Re (diag_idx_to_el B j)) \\<^sub>m sum_mat (\i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B) = sum_mat (\i. complex_of_real (Re (diag_idx_to_el B j)) \\<^sub>m rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B)" proof - fix j assume "j \ {..\<^sub>m sum_mat (\i. rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B) = sum_mat (\i. (Re (diag_idx_to_el B j)) \\<^sub>m rank_1_proj (zero_col U i)) (diag_elem_indices (diag_idx_to_el B j) B)" proof (rule smult_sum_mat) show "finite (diag_elem_indices (diag_idx_to_el B j) B)" using diag_elem_indices_finite[of _ B] by simp show "\i. i \ diag_elem_indices (diag_idx_to_el B j) B \ rank_1_proj (zero_col U i) \ fc_mats" using dim_eq by (metis dimP local.fc_mats_carrier rank_1_proj_carrier zero_col_dim) qed qed qed also have "... = sum_mat (\j. sum_mat (\ia. (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B j) B)) {..i. i \ {.. sum_mat (\ia. complex_of_real (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier) show "\i. i \ {.. local.sum_mat (\ia. (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult sum_mat_carrier) show "\i. i \ {.. sum_mat (\ia. (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B) = sum_mat (\ia. (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B)" proof - fix i assume "i\ {..< dist_el_card B}" show "sum_mat (\ia. (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B) = sum_mat (\ia. (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)) (diag_elem_indices (diag_idx_to_el B i) B)" proof (rule sum_mat_cong) show "finite (diag_elem_indices (diag_idx_to_el B i) B)" using diag_elem_indices_finite[of _ B] by simp show "\ia. ia \ diag_elem_indices (diag_idx_to_el B i) B \ (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult) show "\ia. ia \ diag_elem_indices (diag_idx_to_el B i) B \ (diag_mat B !ia) \\<^sub>m rank_1_proj (zero_col U ia) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult) show "\ia. ia \ diag_elem_indices (diag_idx_to_el B i) B \ (Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia) = (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)" proof - fix ia assume ia: "ia \ diag_elem_indices (diag_idx_to_el B i) B" hence "ia < dim_row B" using diag_elem_indices_elem[of ia _ B] by simp have "Re (diag_idx_to_el B i) = Re (B $$ (ia, ia))" using diag_elem_indices_elem[of ia _ B] ia by simp also have "... = B $$ (ia, ia)" using Bii using \ia < dim_row B\ dimB of_real_Re by blast also have "... = diag_mat B ! ia" using \ia < dim_row B\ unfolding diag_mat_def by simp finally have "Re (diag_idx_to_el B i) = diag_mat B ! ia" . thus "(Re (diag_idx_to_el B i)) \\<^sub>m rank_1_proj (zero_col U ia) = (diag_mat B ! ia) \\<^sub>m rank_1_proj (zero_col U ia)" by simp qed qed qed qed also have "... = sum_mat (\i. (diag_mat B ! i) \\<^sub>m rank_1_proj (zero_col U i)) (\jj. (diag_mat B ! j) \\<^sub>m rank_1_proj (zero_col U j) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult) show "\i. i \ {.. finite (diag_elem_indices (diag_idx_to_el B i) B)" by (simp add: diag_elem_indices_finite) show "disjoint_family_on (\n. diag_elem_indices (diag_idx_to_el B n) B) {..i. (diag_mat B ! i) \\<^sub>m rank_1_proj (zero_col U i)) {..< dimR}" using diag_elem_indices_union[of B] dimB dim_eq by simp also have "... = sum_mat (\i. (diag_mat B ! i) \\<^sub>mrank_1_proj (Matrix.col U i)) {..< dimR}" proof (rule sum_mat_cong, simp) show res: "\i. i \ {.. (diag_mat B ! i) \\<^sub>m rank_1_proj (zero_col U i) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim dim_eq by (metis (no_types, lifting) rank_1_proj_carrier cpx_sq_mat_smult) show "\i. i \ {.. (diag_mat B ! i) \\<^sub>m rank_1_proj (Matrix.col U i) \ fc_mats" using assms fc_mats_carrier dimP project_vecs_def project_vecs_dim zero_col_dim by (metis res lessThan_iff zero_col_col) show "\i. i \ {.. (diag_mat B ! i) \\<^sub>m rank_1_proj (zero_col U i) = (diag_mat B ! i) \\<^sub>m rank_1_proj (Matrix.col U i)" by (simp add: zero_col_col) qed also have "... = U * B * Complex_Matrix.adjoint U" proof (rule weighted_sum_rank_1_proj_unitary) show "diagonal_mat B" using dB . show "Complex_Matrix.unitary U" using unit . show "U \ fc_mats" using fc_mats_carrier dim_eq dimP by simp show "B\ fc_mats" using fc_mats_carrier dim_eq dimB by simp qed also have "... = A" using A by simp finally show ?thesis . qed lemma (in cpx_sq_mat) trace_hermitian_pos_real: fixes f::"'a \ real" assumes "hermitian A" and "Complex_Matrix.positive R" and "A \ fc_mats" and "R \ fc_mats" shows "Complex_Matrix.trace (R * A) = Re (Complex_Matrix.trace (R * A))" proof - define prj_mems where "prj_mems = make_pm A" define p where "p = proj_meas_size prj_mems" define meas where "meas = proj_meas_outcomes prj_mems" have tre: "Complex_Matrix.trace (R * A) = Complex_Matrix.trace (R * sum_mat (\i. (meas_outcome_val (meas i)) \\<^sub>m meas_outcome_prj (meas i)) {..< p})" using make_pm_sum assms meas_def p_def proj_meas_size_def proj_meas_outcomes_def prj_mems_def meas_outcome_val_def meas_outcome_prj_def by auto also have "... = Re (Complex_Matrix.trace (R * sum_mat (\i. (meas_outcome_val (meas i)) \\<^sub>m meas_outcome_prj (meas i)) {..< p}))" proof (rule trace_sum_mat_proj_pos_real, (auto simp add: assms)) fix i assume "i < p" thus "projector (meas_outcome_prj (meas i))" using make_pm_projectors assms unfolding p_def meas_def prj_mems_def by simp show "meas_outcome_prj (meas i) \ fc_mats" using make_pm_square assms \i < p\ unfolding p_def meas_def prj_mems_def by simp qed also have "... = Re (Complex_Matrix.trace (R * A))" using tre by simp finally show ?thesis . qed lemma (in cpx_sq_mat) hermitian_Re_spectrum: assumes "hermitian A" and "A\ fc_mats" and "{Re x |x. x \ spectrum A} = {a,b}" shows "spectrum A = {complex_of_real a, complex_of_real b}" proof have ar: "\(a::complex). a \ spectrum A \ Re a = a" using hermitian_spectrum_real assms by simp show "spectrum A \ {complex_of_real a, complex_of_real b}" proof fix x assume "x\ spectrum A" hence "Re x = x" using ar by simp have "Re x \ {a,b}" using \x\ spectrum A\ assms by blast thus "x \ {complex_of_real a, complex_of_real b}" using \Re x = x\ by auto qed show "{complex_of_real a, complex_of_real b} \ spectrum A" proof fix x assume "x \ {complex_of_real a, complex_of_real b}" hence "x \ {complex_of_real (Re x) |x. x \ spectrum A}" using assms proof - have "\r. r \ {a, b} \ (\c. r = Re c \ c \ spectrum A)" using \{Re x |x. x \ spectrum A} = {a, b}\ by blast then show ?thesis using \x \ {complex_of_real a, complex_of_real b}\ by blast qed thus "x\ spectrum A" using ar by auto qed qed subsubsection \Similar properties for eigenvalues rather than spectrum indices\ text \In this part we go the other way round: given an eigenvalue of $A$, \verb+spectrum_to_pm_idx+ permits to retrieve its index in the associated projective measurement\ definition (in cpx_sq_mat) spectrum_to_pm_idx where "spectrum_to_pm_idx A x = (let (B,U,_) = unitary_schur_decomposition A (eigvals A) in diag_el_to_idx B x)" lemma (in cpx_sq_mat) spectrum_to_pm_idx_bij: assumes "hermitian A" and "A\ fc_mats" shows "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}" proof - define p where "p = proj_meas_size (make_pm A)" define M where "M = proj_meas_outcomes (make_pm A)" have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)") hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diag_mat B = (eigvals A)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp hence "p = dist_el_card B" using assms us unfolding make_pm_def by simp have dimB: "B \ carrier_mat dimR dimR" using dim_eq pr assms fc_mats_carrier unfolding similar_mat_wit_def by auto have Bvals: "diag_mat B = eigvals A" using pr hermitian_decomp_eigenvalues[of A B U] by simp have "diag_elems B = spectrum A" unfolding spectrum_def using dimB Bvals diag_elems_set_diag_mat[of B] by simp moreover have "dist_el_card B = card (spectrum A)" using spectrum_size[of A p M] assms \(p,M) = make_pm A\ \p = dist_el_card B\ by simp ultimately show "bij_betw (spectrum_to_pm_idx A) (spectrum A) {..< card (spectrum A)}" using diag_el_to_idx_bij us unfolding spectrum_to_pm_idx_def Let_def by (metis (mono_tags, lifting) bij_betw_cong case_prod_conv) qed lemma (in cpx_sq_mat) spectrum_to_pm_idx_lt_card: assumes "A\ fc_mats" and "hermitian A" and "a\ spectrum A" shows "spectrum_to_pm_idx A a < card (spectrum A)" using spectrum_to_pm_idx_bij[of A] assms by (meson bij_betw_apply lessThan_iff) lemma (in cpx_sq_mat) spectrum_to_pm_idx_inj: assumes "hermitian A" and "A\ fc_mats" shows "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_bij unfolding bij_betw_def by simp lemma (in cpx_sq_mat) spectrum_meas_outcome_val_inv: assumes "A\ fc_mats" and "hermitian A" and "make_pm A = (p,M)" and "i < p" shows "spectrum_to_pm_idx A (meas_outcome_val (M i)) = i" proof - have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)") hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diag_mat B = (eigvals A) \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A] pr by auto hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us unfolding make_pm_def by simp hence "M = mk_meas_outcome B U" by simp have "meas_outcome_val (M i) = Re (diag_idx_to_el B i)" using \M = mk_meas_outcome B U\ unfolding mk_meas_outcome_def meas_outcome_val_def by simp also have "... = diag_idx_to_el B i" using pr \(p, M) = (dist_el_card B, mk_meas_outcome B U)\ \dim_row B = dimR\ assms diag_idx_to_el_real by auto finally have "meas_outcome_val (M i) = diag_idx_to_el B i" . hence "spectrum_to_pm_idx A (meas_outcome_val (M i)) = spectrum_to_pm_idx A (diag_idx_to_el B i)" by simp also have "... = diag_el_to_idx B (diag_idx_to_el B i)" unfolding spectrum_to_pm_idx_def using us by simp also have "... = i" using assms unfolding diag_el_to_idx_def using \(p, M) = (dist_el_card B, mk_meas_outcome B U)\ bij_betw_inv_into_left diag_idx_to_el_bij by fastforce finally show ?thesis . qed lemma (in cpx_sq_mat) meas_outcome_val_spectrum_inv: assumes "A\ fc_mats" and "hermitian A" and "x\ spectrum A" and "make_pm A = (p,M)" shows "meas_outcome_val (M (spectrum_to_pm_idx A x)) = x" proof - have es: "char_poly A = (\ (e :: complex) \ (eigvals A). [:- e, 1:])" using assms fc_mats_carrier eigvals_poly_length dim_eq by auto obtain B U Q where us: "unitary_schur_decomposition A (eigvals A) = (B,U,Q)" by (cases "unitary_schur_decomposition A (eigvals A)") hence pr: "similar_mat_wit A B U (Complex_Matrix.adjoint U) \ diagonal_mat B \ diag_mat B = (eigvals A) \ unitary U \ (\i < dimR. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real assms fc_mats_carrier es dim_eq by auto have "dim_row B = dimR" using assms fc_mats_carrier dim_eq similar_mat_wit_dim_row[of A] pr by auto hence "(p,M) = (dist_el_card B, mk_meas_outcome B U)" using assms us unfolding make_pm_def by simp hence "M = mk_meas_outcome B U" by simp have "diag_mat B = (eigvals A)" using pr by simp hence "x \ set (diag_mat B)" using \diag_mat B = eigvals A\ assms unfolding spectrum_def by simp hence "x \ diag_elems B" using diag_elems_set_diag_mat[of B] by simp hence "diag_idx_to_el B (diag_el_to_idx B x) = x" unfolding diag_el_to_idx_def by (meson bij_betw_inv_into_right diag_idx_to_el_bij) moreover have "spectrum_to_pm_idx A x = diag_el_to_idx B x" unfolding spectrum_to_pm_idx_def using us by simp moreover have "meas_outcome_val (M (spectrum_to_pm_idx A x)) = Re (diag_idx_to_el B (diag_el_to_idx B x))" using \M = mk_meas_outcome B U\ unfolding mk_meas_outcome_def meas_outcome_val_def by (simp add: calculation(2)) ultimately show ?thesis by simp qed definition (in cpx_sq_mat) eigen_projector where "eigen_projector A a = meas_outcome_prj ((proj_meas_outcomes (make_pm A)) (spectrum_to_pm_idx A a))" lemma (in cpx_sq_mat) eigen_projector_carrier: assumes "A\ fc_mats" and "a\ spectrum A" and "hermitian A" shows "eigen_projector A a \ fc_mats" unfolding eigen_projector_def using meas_outcome_prj_carrier[of A "spectrum_to_pm_idx A a"] spectrum_to_pm_idx_lt_card[of A a] assms by simp text \We obtain the following result, which is similar to \verb+make_pm_sum+ but with a sum on the elements of the spectrum of Hermitian matrix $A$, which is a more standard statement of the spectral decomposition theorem.\ lemma (in cpx_sq_mat) make_pm_sum': assumes "A \ fc_mats" and "hermitian A" shows "sum_mat (\a. a \\<^sub>m (eigen_projector A a)) (spectrum A) = A" proof - define p where "p = proj_meas_size (make_pm A)" define M where "M = proj_meas_outcomes (make_pm A)" have "(p,M) = make_pm A" unfolding p_def M_def using make_pm_decomp by simp define g where "g = (\i. (if i < p then complex_of_real (meas_outcome_val (M i)) \\<^sub>m meas_outcome_prj (M i) else (0\<^sub>m dimR dimC)))" have g: "\x. g x \ fc_mats" proof fix x show "g x \ fc_mats" proof (cases "x < p") case True hence "(meas_outcome_val (M x)) \\<^sub>m meas_outcome_prj (M x) \ fc_mats" using meas_outcome_prj_carrier spectrum_size assms cpx_sq_mat_smult make_pm_proj_measurement proj_measurement_square \(p,M) = make_pm A\ by metis then show ?thesis unfolding g_def using True by simp next case False then show ?thesis unfolding g_def using zero_mem by simp qed qed define h where "h = (\a. (if a \ (spectrum A) then a \\<^sub>m eigen_projector A a else (0\<^sub>m dimR dimC)))" have h: "\x. h x \ fc_mats" proof fix x show "h x \ fc_mats" proof (cases "x\ spectrum A") case True then show ?thesis unfolding h_def using eigen_projector_carrier[of A x] assms True by (simp add: cpx_sq_mat_smult) next case False then show ?thesis unfolding h_def using zero_mem by simp qed qed have "inj_on (spectrum_to_pm_idx A) (spectrum A)" using assms spectrum_to_pm_idx_inj by simp moreover have "{..(p,M) = make_pm A\ spectrum_to_pm_idx_bij spectrum_size unfolding bij_betw_def by (metis assms(1) assms(2)) moreover have "\x. x \ spectrum A \ g (spectrum_to_pm_idx A x) = h x" proof - fix a assume "a \ spectrum A" hence "Re a = a" using hermitian_spectrum_real assms by simp have "spectrum_to_pm_idx A a < p" using spectrum_to_pm_idx_lt_card[of A] spectrum_size assms \a \ spectrum A\ \(p,M) = make_pm A\ by metis have "g (spectrum_to_pm_idx A a) = (meas_outcome_val (M (spectrum_to_pm_idx A a))) \\<^sub>m meas_outcome_prj (M (spectrum_to_pm_idx A a))" using \spectrum_to_pm_idx A a < p\ unfolding g_def by simp also have "... = a \\<^sub>m meas_outcome_prj (M (spectrum_to_pm_idx A a))" using meas_outcome_val_spectrum_inv[of A "Re a"] \Re a = a\ assms \a \ spectrum A\ \(p,M) = make_pm A\ by metis also have "... = h a" using \a \ spectrum A\ unfolding eigen_projector_def M_def h_def by simp finally show "g (spectrum_to_pm_idx A a) = h a" . qed ultimately have "sum_mat h (spectrum A) = sum_mat g (spectrum_to_pm_idx A ` spectrum A)" unfolding sum_mat_def using sum_with_reindex_cong[symmetric, of g h "spectrum_to_pm_idx A" "spectrum A" "{..< p}"] g h by simp also have "... = sum_mat g {..< p}" using \{.. by simp also have "... = sum_mat (\i. (meas_outcome_val (M i)) \\<^sub>m meas_outcome_prj (M i)) {.. {..< p}" hence "i < p" by simp show "g i \ fc_mats" using g by simp show "g i = (meas_outcome_val (M i)) \\<^sub>m meas_outcome_prj (M i)" unfolding g_def using \i < p\ by simp show "(meas_outcome_val (M i)) \\<^sub>m meas_outcome_prj (M i) \ fc_mats" using meas_outcome_prj_carrier spectrum_size assms cpx_sq_mat_smult make_pm_proj_measurement proj_measurement_square \i < p\ \(p,M) = make_pm A\ by metis qed also have "... = A" using make_pm_sum assms \(p,M) = make_pm A\ unfolding g_def by simp finally have "sum_mat h (spectrum A) = A" . moreover have "sum_mat h (spectrum A) = sum_mat (\a. a \\<^sub>m (eigen_projector A a)) (spectrum A)" proof (rule sum_mat_cong) show "finite (spectrum A)" using spectrum_finite[of A] by simp fix i assume "i\ spectrum A" thus "h i = i \\<^sub>m eigen_projector A i" unfolding h_def by simp show "h i \ fc_mats" using h by simp show "i \\<^sub>m eigen_projector A i \ fc_mats" using eigen_projector_carrier[of A i] assms \i\ spectrum A\ by (metis \h i = i \\<^sub>m eigen_projector A i\ h) qed ultimately show ?thesis by simp qed end \ No newline at end of file diff --git a/thys/QHLProver/Complex_Matrix.thy b/thys/QHLProver/Complex_Matrix.thy --- a/thys/QHLProver/Complex_Matrix.thy +++ b/thys/QHLProver/Complex_Matrix.thy @@ -1,2346 +1,2351 @@ section \Complex matrices\ theory Complex_Matrix imports "Jordan_Normal_Form.Matrix" "Jordan_Normal_Form.Conjugate" "Jordan_Normal_Form.Jordan_Normal_Form_Existence" begin subsection \Trace of a matrix\ definition trace :: "'a::ring mat \ 'a" where "trace A = (\ i \ {0 ..< dim_row A}. A $$ (i,i))" lemma trace_zero [simp]: "trace (0\<^sub>m n n) = 0" by (simp add: trace_def) lemma trace_id [simp]: "trace (1\<^sub>m n) = n" by (simp add: trace_def) lemma trace_comm: fixes A B :: "'a::comm_ring mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A * B) = trace (B * A)" proof (simp add: trace_def) have "(\i = 0..i = 0..j = 0.. = (\j = 0..i = 0.. = (\j = 0.. row B j)" by (metis (no_types, lifting) A B atLeastLessThan_iff carrier_matD index_col index_row scalar_prod_def sum.cong) also have "\ = (\j = 0.. col A j)" apply (rule sum.cong) apply auto apply (subst comm_scalar_prod[where n=n]) apply auto using assms by auto also have "\ = (\j = 0..i = 0..i = 0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A + B) = trace A + trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" and B: "B \ carrier_mat n n" shows "trace (A - B) = trace A - trace B" (is "?lhs = ?rhs") proof - have "?lhs = (\i=0.. = (\i=0..i=0..i=0..i=0..i=0..i=0.. carrier_mat n n" shows "trace (c \\<^sub>m A) = c * trace A" proof - have "trace (c \\<^sub>m A) = (\i = 0.. = c * (\i = 0.. = c * trace A" unfolding trace_def by auto ultimately show ?thesis by auto qed subsection \Conjugate of a vector\ lemma conjugate_scalar_prod: fixes v w :: "'a::conjugatable_ring vec" assumes "dim_vec v = dim_vec w" shows "conjugate (v \ w) = conjugate v \ conjugate w" using assms by (simp add: scalar_prod_def sum_conjugate conjugate_dist_mul) subsection \Inner product\ abbreviation inner_prod :: "'a vec \ 'a vec \ 'a :: conjugatable_ring" where "inner_prod v w \ w \c v" lemma conjugate_scalar_prod_Im [simp]: "Im (v \c v) = 0" by (simp add: scalar_prod_def conjugate_vec_def sum.neutral) lemma conjugate_scalar_prod_Re [simp]: "Re (v \c v) \ 0" by (simp add: scalar_prod_def conjugate_vec_def sum_nonneg) lemma self_cscalar_prod_geq_0: fixes v :: "'a::conjugatable_ordered_field vec" shows "v \c v \ 0" by (auto simp add: scalar_prod_def, rule sum_nonneg, rule conjugate_square_positive) lemma inner_prod_distrib_left: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod (v + w) u = inner_prod v u + inner_prod w u" (is "?lhs = ?rhs") proof - have dimcv: "conjugate v \ carrier_vec n" and dimcw: "conjugate w \ carrier_vec n" using assms by auto have dimvw: "conjugate (v + w) \ carrier_vec n" using assms by auto have "u \ (conjugate (v + w)) = u \ conjugate v + u \ conjugate w" using dimv dimw dimu dimcv dimcw by (metis conjugate_add_vec scalar_prod_add_distrib) then show ?thesis by auto qed lemma inner_prod_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v + w) = inner_prod u v + inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v + w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v + w) \ (conjugate u) = v \ conjugate u + w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_add_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_minus_distrib_right: fixes u v w :: "('a::conjugatable_field) vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" and dimw: "w \ carrier_vec n" shows "inner_prod u (v - w) = inner_prod u v - inner_prod u w" (is "?lhs = ?rhs") proof - have dimvw: "v - w \ carrier_vec n" using assms by auto have dimcu: "conjugate u \ carrier_vec n" using assms by auto have "(v - w) \ (conjugate u) = v \ conjugate u - w \ conjugate u" apply (simp add: comm_scalar_prod[OF dimvw dimcu]) apply (simp add: scalar_prod_minus_distrib[OF dimcu dimv dimw]) apply (insert dimv dimw dimcu, simp add: comm_scalar_prod[of _ n]) done then show ?thesis by auto qed lemma inner_prod_smult_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv:"v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) v = conjugate a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def conjugate_dist_mul) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod u (a \\<^sub>v v) = a * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_smult_left_right: fixes u v :: "complex vec" assumes dimu: "u \ carrier_vec n" and dimv: "v \ carrier_vec n" shows "inner_prod (a \\<^sub>v u) (b \\<^sub>v v) = conjugate a * b * inner_prod u v" (is "?lhs = ?rhs") using assms apply (simp add: scalar_prod_def) apply (subst sum_distrib_left) by (rule sum.cong, auto) lemma inner_prod_swap: fixes x y :: "complex vec" assumes "y \ carrier_vec n" and "x \ carrier_vec n" shows "inner_prod y x = conjugate (inner_prod x y)" apply (simp add: scalar_prod_def) apply (rule sum.cong) using assms by auto text \Cauchy-Schwarz theorem for complex vectors. This is analogous to aux\_Cauchy and Cauchy\_Schwarz\_ineq in Generalizations2.thy in QR\_Decomposition. Consider merging and moving to Isabelle library.\ lemma aux_Cauchy: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "0 \ inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y))" proof - have "(inner_prod (x+ a \\<^sub>v y) (x+a \\<^sub>v y)) = (inner_prod (x+a \\<^sub>v y) x) + (inner_prod (x+a \\<^sub>v y) (a \\<^sub>v y))" apply (subst inner_prod_distrib_right) using assms by auto also have "\ = inner_prod x x + (a) * (inner_prod x y) + cnj a * ((cnj (inner_prod x y)) + (a) * (inner_prod y y))" apply (subst (1 2) inner_prod_distrib_left[of _ n]) apply (auto simp add: assms) apply (subst (1 2) inner_prod_smult_right[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_smult_left[of _ n]) apply (auto simp add: assms) apply (subst inner_prod_swap[of y n x]) apply (auto simp add: assms) unfolding distrib_left by auto finally show ?thesis by (metis self_cscalar_prod_geq_0) qed lemma Cauchy_Schwarz_complex_vec: fixes x y :: "complex vec" assumes "x \ carrier_vec n" and "y \ carrier_vec n" shows "inner_prod x y * inner_prod y x \ inner_prod x x * inner_prod y y" proof - define cnj_a where "cnj_a = - (inner_prod x y)/ cnj (inner_prod y y)" define a where "a = cnj (cnj_a)" have cnj_rw: "(cnj a) = cnj_a" unfolding a_def by (simp) have rw_0: "cnj (inner_prod x y) + a * (inner_prod y y) = 0" unfolding a_def cnj_a_def using assms(1) assms(2) conjugate_square_eq_0_vec by fastforce have "0 \ (inner_prod x x + a * (inner_prod x y) + (cnj a) * ((cnj (inner_prod x y)) + a * (inner_prod y y)))" using aux_Cauchy assms by auto also have "\ = (inner_prod x x + a * (inner_prod x y))" unfolding rw_0 by auto also have "\ = (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y))" unfolding a_def cnj_a_def by simp finally have " 0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) " . - hence "0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) * (inner_prod y y)" by auto + hence "0 \ (inner_prod x x - (inner_prod x y) * cnj (inner_prod x y) / (inner_prod y y)) * (inner_prod y y)" + by (auto simp: less_eq_complex_def) also have "\ = ((inner_prod x x)*(inner_prod y y) - (inner_prod x y) * cnj (inner_prod x y))" by (smt add.inverse_neutral add_diff_cancel diff_0 diff_divide_eq_iff divide_cancel_right mult_eq_0_iff nonzero_mult_div_cancel_right rw_0) finally have "(inner_prod x y) * cnj (inner_prod x y) \ (inner_prod x x)*(inner_prod y y)" by auto then show ?thesis apply (subst inner_prod_swap[of y n x]) by (auto simp add: assms) qed subsection \Hermitian adjoint of a matrix\ abbreviation adjoint where "adjoint \ mat_adjoint" lemma adjoint_dim_row [simp]: "dim_row (adjoint A) = dim_col A" by (simp add: mat_adjoint_def) lemma adjoint_dim_col [simp]: "dim_col (adjoint A) = dim_row A" by (simp add: mat_adjoint_def) lemma adjoint_dim: "A \ carrier_mat n n \ adjoint A \ carrier_mat n n" using adjoint_dim_col adjoint_dim_row by blast lemma adjoint_def: "adjoint A = mat (dim_col A) (dim_row A) (\(i,j). conjugate (A $$ (j,i)))" unfolding mat_adjoint_def mat_of_rows_def by auto lemma adjoint_eval: assumes "i < dim_col A" "j < dim_row A" shows "(adjoint A) $$ (i,j) = conjugate (A $$ (j,i))" using assms by (simp add: adjoint_def) lemma adjoint_row: assumes "i < dim_col A" shows "row (adjoint A) i = conjugate (col A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) lemma adjoint_col: assumes "i < dim_row A" shows "col (adjoint A) i = conjugate (row A i)" apply (rule eq_vecI) using assms by (auto simp add: adjoint_eval) text \The identity = \ lemma adjoint_def_alter: fixes v w :: "'a::conjugatable_field vec" and A :: "'a::conjugatable_field mat" assumes dims: "v \ carrier_vec n" "w \ carrier_vec m" "A \ carrier_mat n m" shows "inner_prod v (A *\<^sub>v w) = inner_prod (adjoint A *\<^sub>v v) w" (is "?lhs = ?rhs") proof - from dims have "?lhs = (\i=0..j=0..i=0..j=0..m n) = (1\<^sub>m n::complex mat)" apply (rule eq_matI) by (auto simp add: adjoint_eval) lemma adjoint_scale: fixes A :: "'a::conjugatable_field mat" shows "adjoint (a \\<^sub>m A) = (conjugate a) \\<^sub>m adjoint A" apply (rule eq_matI) using conjugatable_ring_class.conjugate_dist_mul by (auto simp add: adjoint_eval) lemma adjoint_add: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A + B) = adjoint A + adjoint B" apply (rule eq_matI) using assms conjugatable_ring_class.conjugate_dist_add by( auto simp add: adjoint_eval) lemma adjoint_minus: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat n m" shows "adjoint (A - B) = adjoint A - adjoint B" apply (rule eq_matI) using assms apply(auto simp add: adjoint_eval) by (metis add_uminus_conv_diff conjugate_dist_add conjugate_neg) lemma adjoint_mult: fixes A B :: "'a::conjugatable_field mat" assumes "A \ carrier_mat n m" "B \ carrier_mat m l" shows "adjoint (A * B) = adjoint B * adjoint A" proof (rule eq_matI, auto simp add: adjoint_eval adjoint_row adjoint_col) fix i j assume "i < dim_col B" "j < dim_row A" show "conjugate (row A j \ col B i) = conjugate (col B i) \ conjugate (row A j)" using assms apply (simp add: conjugate_scalar_prod) apply (subst comm_scalar_prod[where n="dim_row B"]) by (auto simp add: carrier_vecI) qed lemma adjoint_adjoint: fixes A :: "'a::conjugatable_field mat" shows "adjoint (adjoint A) = A" by (rule eq_matI, auto simp add: adjoint_eval) lemma trace_adjoint_positive: fixes A :: "complex mat" shows "trace (A * adjoint A) \ 0" apply (auto simp add: trace_def adjoint_col) apply (rule sum_nonneg) by auto subsection \Algebraic manipulations on matrices\ lemma right_add_zero_mat[simp]: "(A :: 'a :: monoid_add mat) \ carrier_mat nr nc \ A + 0\<^sub>m nr nc = A" by (intro eq_matI, auto) lemma add_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A + B \ carrier_mat nr nc" by simp lemma minus_carrier_mat': "A \ carrier_mat nr nc \ B \ carrier_mat nr nc \ A - B \ carrier_mat nr nc" by auto lemma swap_plus_mat: fixes A B C :: "'a::semiring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" shows "A + B + C = A + C + B" by (metis assms assoc_add_mat comm_add_mat) lemma uminus_mat: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "-A = (-1) \\<^sub>m A" by auto ML_file "mat_alg.ML" method_setup mat_assoc = \mat_assoc_method\ "Normalization of expressions on matrices" lemma mat_assoc_test: fixes A B C D :: "complex mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" shows "(A * B) * (C * D) = A * B * C * D" "adjoint (A * adjoint B) * C = B * (adjoint A * C)" "A * 1\<^sub>m n * 1\<^sub>m n * B * 1\<^sub>m n = A * B" "(A - B) + (B - C) = A + (-B) + B + (-C)" "A + (B - C) = A + B - C" "A - (B + C + D) = A - B - C - D" "(A + B) * (B + C) = A * B + B * B + A * C + B * C" "A - B = A + (-1) \\<^sub>m B" "A * (B - C) * D = A * B * D - A * C * D" "trace (A * B * C) = trace (B * C * A)" "trace (A * B * C * D) = trace (C * D * A * B)" "trace (A + B * C) = trace A + trace (C * B)" "A + B = B + A" "A + B + C = C + B + A" "A + B + (C + D) = A + C + (B + D)" using assms by (mat_assoc n)+ subsection \Hermitian matrices\ text \A Hermitian matrix is a matrix that is equal to its Hermitian adjoint.\ definition hermitian :: "'a::conjugatable_field mat \ bool" where "hermitian A \ (adjoint A = A)" lemma hermitian_one: shows "hermitian ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding hermitian_def proof- have "conjugate (1::'a) = 1" apply (subst mult_1_right[symmetric, of "conjugate 1"]) apply (subst conjugate_id[symmetric, of "conjugate 1 * 1"]) apply (subst conjugate_dist_mul) apply auto done then show "adjoint ((1\<^sub>m n)::('a::conjugatable_field mat)) = (1\<^sub>m n)" by (auto simp add: adjoint_eval) qed subsection \Inverse matrices\ lemma inverts_mat_symm: fixes A B :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" and AB: "inverts_mat A B" shows "inverts_mat B A" proof - have "A * B = 1\<^sub>m n" using dim AB unfolding inverts_mat_def by auto with dim have "B * A = 1\<^sub>m n" by (rule mat_mult_left_right_inverse) then show "inverts_mat B A" using dim inverts_mat_def by auto qed lemma inverts_mat_unique: fixes A B C :: "'a::field mat" assumes dim: "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" and AB: "inverts_mat A B" and AC: "inverts_mat A C" shows "B = C" proof - have AB1: "A * B = 1\<^sub>m n" using AB dim unfolding inverts_mat_def by auto have "A * C = 1\<^sub>m n" using AC dim unfolding inverts_mat_def by auto then have CA1: "C * A = 1\<^sub>m n" using mat_mult_left_right_inverse[of A n C] dim by auto then have "C = C * 1\<^sub>m n" using dim by auto also have "\ = C * (A * B)" using AB1 by auto also have "\ = (C * A) * B" using dim by auto also have "\ = 1\<^sub>m n * B" using CA1 by auto also have "\ = B" using dim by auto finally show "B = C" .. qed subsection \Unitary matrices\ text \A unitary matrix is a matrix whose Hermitian adjoint is also its inverse.\ definition unitary :: "'a::conjugatable_field mat \ bool" where "unitary A \ A \ carrier_mat (dim_row A) (dim_row A) \ inverts_mat A (adjoint A)" lemma unitaryD2: assumes "A \ carrier_mat n n" shows "unitary A \ inverts_mat (adjoint A) A" using assms adjoint_dim inverts_mat_symm unitary_def by blast lemma unitary_simps [simp]: "A \ carrier_mat n n \ unitary A \ adjoint A * A = 1\<^sub>m n" "A \ carrier_mat n n \ unitary A \ A * adjoint A = 1\<^sub>m n" apply (metis adjoint_dim_row carrier_matD(2) inverts_mat_def unitaryD2) by (simp add: inverts_mat_def unitary_def) lemma unitary_adjoint [simp]: assumes "A \ carrier_mat n n" "unitary A" shows "unitary (adjoint A)" unfolding unitary_def using adjoint_dim[OF assms(1)] assms by (auto simp add: unitaryD2[OF assms] adjoint_adjoint) lemma unitary_one: shows "unitary ((1\<^sub>m n)::('a::conjugatable_field mat))" unfolding unitary_def proof - define I where I_def[simp]: "I \ ((1\<^sub>m n)::('a::conjugatable_field mat))" have dim: "I \ carrier_mat n n" by auto have "hermitian I" using hermitian_one by auto hence "adjoint I = I" using hermitian_def by auto with dim show "I \ carrier_mat (dim_row I) (dim_row I) \ inverts_mat I (adjoint I)" unfolding inverts_mat_def using dim by auto qed lemma unitary_zero: fixes A :: "'a::conjugatable_field mat" assumes "A \ carrier_mat 0 0" shows "unitary A" unfolding unitary_def inverts_mat_def Let_def using assms by auto lemma unitary_elim: assumes dims: "A \ carrier_mat n n" "B \ carrier_mat n n" "P \ carrier_mat n n" and uP: "unitary P" and eq: "P * A * adjoint P = P * B * adjoint P" shows "A = B" proof - have dimaP: "adjoint P \ carrier_mat n n" using dims by auto have iv: "inverts_mat P (adjoint P)" using uP unitary_def by auto then have "P * (adjoint P) = 1\<^sub>m n" using inverts_mat_def dims by auto then have aPP: "adjoint P * P = 1\<^sub>m n" using mat_mult_left_right_inverse[OF dims(3) dimaP] by auto have "adjoint P * (P * A * adjoint P) * P = (adjoint P * P) * A * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * A * 1\<^sub>m n" using aPP by auto also have "\ = A" using dims by auto finally have eqA: "A = adjoint P * (P * A * adjoint P) * P" .. have "adjoint P * (P * B * adjoint P) * P = (adjoint P * P) * B * (adjoint P * P)" using dims dimaP by (mat_assoc n) also have "\ = 1\<^sub>m n * B * 1\<^sub>m n" using aPP by auto also have "\ = B" using dims by auto finally have eqB: "B = adjoint P * (P * B * adjoint P) * P" .. then show ?thesis using eqA eqB eq by auto qed lemma unitary_is_corthogonal: fixes U :: "'a::conjugatable_field mat" assumes dim: "U \ carrier_mat n n" and U: "unitary U" shows "corthogonal_mat U" unfolding corthogonal_mat_def Let_def proof (rule conjI) have dima: "adjoint U \ carrier_mat n n" using dim by auto have aUU: "mat_adjoint U * U = (1\<^sub>m n)" apply (insert U[unfolded unitary_def] dim dima, drule conjunct2) apply (drule inverts_mat_symm[of "U", OF dim dima], unfold inverts_mat_def, auto) done then show "diagonal_mat (mat_adjoint U * U)" by (simp add: diagonal_mat_def) show "\i 0" using dim by (simp add: aUU) qed lemma unitary_times_unitary: fixes P Q :: "'a:: conjugatable_field mat" assumes dim: "P \ carrier_mat n n" "Q \ carrier_mat n n" and uP: "unitary P" and uQ: "unitary Q" shows "unitary (P * Q)" proof - have dim_pq: "P * Q \ carrier_mat n n" using dim by auto have "(P * Q) * adjoint (P * Q) = P * (Q * adjoint Q) * adjoint P" using dim by (mat_assoc n) also have "\ = P * (1\<^sub>m n) * adjoint P" using uQ dim by auto also have "\ = P * adjoint P" using dim by (mat_assoc n) also have "\ = 1\<^sub>m n" using uP dim by simp finally have "(P * Q) * adjoint (P * Q) = 1\<^sub>m n" by auto hence "inverts_mat (P * Q) (adjoint (P * Q))" using inverts_mat_def dim_pq by auto thus "unitary (P*Q)" using unitary_def dim_pq by auto qed lemma unitary_operator_keep_trace: fixes U A :: "complex mat" assumes dU: "U \ carrier_mat n n" and dA: "A \ carrier_mat n n" and u: "unitary U" shows "trace A = trace (adjoint U * A * U)" proof - have u': "U * adjoint U = 1\<^sub>m n" using u unfolding unitary_def inverts_mat_def using dU by auto have "trace (adjoint U * A * U) = trace (U * adjoint U * A)" using dU dA by (mat_assoc n) also have "\ = trace A" using u' dA by auto finally show ?thesis by auto qed subsection \Normalization of vectors\ definition vec_norm :: "complex vec \ complex" where "vec_norm v \ csqrt (v \c v)" lemma vec_norm_geq_0: fixes v :: "complex vec" shows "vec_norm v \ 0" - unfolding vec_norm_def by (insert self_cscalar_prod_geq_0[of v], simp) + unfolding vec_norm_def by (insert self_cscalar_prod_geq_0[of v], simp add: less_eq_complex_def) lemma vec_norm_zero: fixes v :: "complex vec" assumes dim: "v \ carrier_vec n" shows "vec_norm v = 0 \ v = 0\<^sub>v n" unfolding vec_norm_def by (subst conjugate_square_eq_0_vec[OF dim, symmetric], rule csqrt_eq_0) lemma vec_norm_ge_0: fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_norm v > 0" proof - have geq: "vec_norm v \ 0" using vec_norm_geq_0 by auto have neq: "vec_norm v \ 0" apply (insert dim_v neq0) apply (drule vec_norm_zero, auto) done show ?thesis using neq geq by (rule dual_order.not_eq_order_implies_strict) qed definition vec_normalize :: "complex vec \ complex vec" where "vec_normalize v = (if (v = 0\<^sub>v (dim_vec v)) then v else 1 / (vec_norm v) \\<^sub>v v)" lemma normalized_vec_dim[simp]: assumes "(v::complex vec) \ carrier_vec n" shows "vec_normalize v \ carrier_vec n" unfolding vec_normalize_def using assms by auto lemma vec_eq_norm_smult_normalized: shows "v = vec_norm v \\<^sub>v vec_normalize v" proof (cases "v = 0\<^sub>v (dim_vec v)") define n where "n = dim_vec v" then have dimv: "v \ carrier_vec n" by auto then have dimnv: "vec_normalize v \ carrier_vec n" by auto { case True then have v0: "v = 0\<^sub>v n" using n_def by auto then have n0: "vec_norm v = 0" using vec_norm_def by auto have "vec_norm v \\<^sub>v vec_normalize v = 0\<^sub>v n" unfolding smult_vec_def by (auto simp add: n0 carrier_vecD[OF dimnv]) then show ?thesis using v0 by auto next case False then have v: "v \ 0\<^sub>v n" using n_def by auto then have ge0: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto have "vec_normalize v = (1 / vec_norm v) \\<^sub>v v" using False vec_normalize_def by auto then have "vec_norm v \\<^sub>v vec_normalize v = (vec_norm v * (1 / vec_norm v)) \\<^sub>v v" using smult_smult_assoc by auto also have "\ = v" using ge0 by auto finally have "v = vec_norm v \\<^sub>v vec_normalize v".. then show "v = vec_norm v \\<^sub>v vec_normalize v" using v by auto } qed lemma normalized_cscalar_prod: fixes v w :: "complex vec" assumes dim_v: "v \ carrier_vec n" and dim_w: "w \ carrier_vec n" shows "v \c w = (vec_norm v * vec_norm w) * (vec_normalize v \c vec_normalize w)" unfolding vec_normalize_def apply (split if_split, split if_split) proof (intro conjI impI) note dim0 = dim_v dim_w have dim: "dim_vec v = n" "dim_vec w = n" using dim0 by auto { assume "w = 0\<^sub>v n" "v = 0\<^sub>v n" then have lhs: "v \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c w) = 0" by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c w)" by auto { assume asm: "w = 0\<^sub>v n" "v \ 0\<^sub>v n" then have w0: "conjugate w = 0\<^sub>v n" by auto with dim0 have "(1 / vec_norm v \\<^sub>v v) \c w = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w) = 0" by auto moreover have "v \c w = 0" using w0 dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto } with dim show "w = 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c w)" by auto { assume asm: "w \ 0\<^sub>v n" "v = 0\<^sub>v n" with dim0 have "v \c (1 / vec_norm w \\<^sub>v w) = 0" by auto then moreover have rhs: "vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w)) = 0" by auto moreover have "v \c w = 0" using asm dim0 by auto ultimately have "v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto } with dim show "w \ 0\<^sub>v (dim_vec w) \ v = 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * (v \c (1 / vec_norm w \\<^sub>v w))" by auto { assume asmw: "w \ 0\<^sub>v n" and asmv: "v \ 0\<^sub>v n" have "vec_norm w > 0" by (insert asmw dim0, rule vec_norm_ge_0, auto) - then have cw: "conjugate (1 / vec_norm w) = 1 / vec_norm w" by (simp add: complex_eq_iff complex_is_Real_iff) + then have cw: "conjugate (1 / vec_norm w) = 1 / vec_norm w" + by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) from dim0 have "((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = 1 / vec_norm v * (v \c (1 / vec_norm w \\<^sub>v w))" by auto also have "\ = 1 / vec_norm v * (v \ (conjugate (1 / vec_norm w) \\<^sub>v conjugate w))" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * conjugate (1 / vec_norm w) * (v \ conjugate w)" using dim by auto also have "\ = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" using vec_norm_ge_0 cw by auto finally have eq1: "(1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w) = 1 / vec_norm v * (1 / vec_norm w) * (v \c w)" . then have "vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w)) = (v \c w)" by (subst eq1, insert vec_norm_ge_0[of v n, OF dim_v asmv] vec_norm_ge_0[of w n, OF dim_w asmw], auto) } with dim show " w \ 0\<^sub>v (dim_vec w) \ v \ 0\<^sub>v (dim_vec v) \ v \c w = vec_norm v * vec_norm w * ((1 / vec_norm v \\<^sub>v v) \c (1 / vec_norm w \\<^sub>v w))" by auto qed lemma normalized_vec_norm : fixes v :: "complex vec" assumes dim_v: "v \ carrier_vec n" and neq0: "v \ 0\<^sub>v n" shows "vec_normalize v \c vec_normalize v = 1" unfolding vec_normalize_def proof (simp, rule conjI) show "v = 0\<^sub>v (dim_vec v) \ v \c v = 1" using neq0 dim_v by auto have dim_a: "(vec_normalize v) \ carrier_vec n" "conjugate (vec_normalize v) \ carrier_vec n" using dim_v vec_normalize_def by auto note dim = dim_v dim_a have nvge0: "vec_norm v > 0" using vec_norm_ge_0 neq0 dim_v by auto then have vvvv: "v \c v = (vec_norm v) * (vec_norm v)" unfolding vec_norm_def by (metis power2_csqrt power2_eq_square) - from nvge0 have "conjugate (vec_norm v) = vec_norm v" by (simp add: complex_eq_iff complex_is_Real_iff) + from nvge0 have "conjugate (vec_norm v) = vec_norm v" + by (simp add: complex_eq_iff complex_is_Real_iff less_complex_def) then have "v \c (1 / vec_norm v \\<^sub>v v) = 1 / vec_norm v * (v \c v)" by (subst conjugate_smult_vec, auto) also have "\ = 1 / vec_norm v * vec_norm v * vec_norm v" using vvvv by auto also have "\ = vec_norm v" by auto finally have "v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v". then show "v \ 0\<^sub>v (dim_vec v) \ vec_norm v \ 0 \ v \c (1 / vec_norm v \\<^sub>v v) = vec_norm v" using neq0 nvge0 by auto qed lemma normalize_zero: assumes "v \ carrier_vec n" shows "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" proof show "v = 0\<^sub>v n \ vec_normalize v = 0\<^sub>v n" unfolding vec_normalize_def by auto next have "v \ 0\<^sub>v n \ vec_normalize v \ 0\<^sub>v n" unfolding vec_normalize_def proof (simp, rule impI) assume asm: "v \ 0\<^sub>v n" then have "vec_norm v > 0" using vec_norm_ge_0 assms by auto - then have nvge0: "1 / vec_norm v > 0" by (simp add: complex_is_Real_iff) + then have nvge0: "1 / vec_norm v > 0" by (simp add: complex_is_Real_iff less_complex_def) have "\k < n. v $ k \ 0" using asm assms by auto then obtain k where kn: "k < n" and vkneq0: "v $ k \ 0" by auto then have "(1 / vec_norm v \\<^sub>v v) $ k = (1 / vec_norm v) * (v $ k)" using assms carrier_vecD index_smult_vec(1) by blast with nvge0 vkneq0 have "(1 / vec_norm v \\<^sub>v v) $ k \ 0" by auto then show "1 / vec_norm v \\<^sub>v v \ 0\<^sub>v n" using assms kn by fastforce qed then show "vec_normalize v = 0\<^sub>v n \ v = 0\<^sub>v n" by auto qed lemma normalize_normalize[simp]: "vec_normalize (vec_normalize v) = vec_normalize v" proof (rule disjE[of "v = 0\<^sub>v (dim_vec v)" "v \ 0\<^sub>v (dim_vec v)"], auto) let ?n = "dim_vec v" { assume "v = 0\<^sub>v ?n" then have "vec_normalize v = v" unfolding vec_normalize_def by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by auto } assume neq0: "v \ 0\<^sub>v ?n" have dim: "v \ carrier_vec ?n" by auto have "vec_norm (vec_normalize v) = 1" unfolding vec_norm_def using normalized_vec_norm[OF dim neq0] by auto then show "vec_normalize (vec_normalize v) = vec_normalize v" by (subst (1) vec_normalize_def, simp) qed subsection \Spectral decomposition of normal complex matrices\ lemma normalize_keep_corthogonal: fixes vs :: "complex vec list" assumes cor: "corthogonal vs" and dims: "set vs \ carrier_vec n" shows "corthogonal (map vec_normalize vs)" unfolding corthogonal_def proof (rule allI, rule impI, rule allI, rule impI, goal_cases) case c: (1 i j) let ?m = "length vs" have len: "length (map vec_normalize vs) = ?m" by auto have dim: "\k. k < ?m \ (vs ! k) \ carrier_vec n" using dims by auto have map: "\k. k < ?m \ map vec_normalize vs ! k = vec_normalize (vs ! k)" by auto have eq1: "\j k. j < ?m \ k < ?m \ ((vs ! j) \c (vs ! k) = 0) = (j \ k)" using assms unfolding corthogonal_def by auto then have "\k. k < ?m \ (vs ! k) \c (vs ! k) \ 0 " by auto then have "\k. k < ?m \ (vs ! k) \ (0\<^sub>v n)" using dim by (auto simp add: conjugate_square_eq_0_vec[of _ n, OF dim]) then have vnneq0: "\k. k < ?m \ vec_norm (vs ! k) \ 0" using vec_norm_zero[OF dim] by auto then have i0: "vec_norm (vs ! i) \ 0" and j0: "vec_norm (vs ! j) \ 0" using c by auto have "(vs ! i) \c (vs ! j) = vec_norm (vs ! i) * vec_norm (vs ! j) * (vec_normalize (vs ! i) \c vec_normalize (vs ! j))" by (subst normalized_cscalar_prod[of "vs ! i" n "vs ! j"], auto, insert dim c, auto) with i0 j0 have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = ((vs ! i) \c (vs ! j) = 0)" by auto with eq1 c have "(vec_normalize (vs ! i) \c vec_normalize (vs ! j) = 0) = (i \ j)" by auto with map c show "(map vec_normalize vs ! i \c map vec_normalize vs ! j = 0) = (i \ j)" by auto qed lemma normalized_corthogonal_mat_is_unitary: assumes W: "set ws \ carrier_vec n" and orth: "corthogonal ws" and len: "length ws = n" shows "unitary (mat_of_cols n (map vec_normalize ws))" (is "unitary ?W") proof - define vs where "vs = map vec_normalize ws" define W where "W = mat_of_cols n vs" have W': "set vs \ carrier_vec n" using assms vs_def by auto then have W'': "\k. k < length vs \ vs ! k \ carrier_vec n" by auto have orth': "corthogonal vs" using assms normalize_keep_corthogonal vs_def by auto have len'[simp]: "length vs = n" using assms vs_def by auto have dimW: "W \ carrier_mat n n" using W_def len by auto have "adjoint W \ carrier_mat n n" using dimW by auto then have dimaW: "mat_adjoint W \ carrier_mat n n" by auto { fix i j assume i: "i < n" and j: "j < n" have dimws: "(ws ! i) \ carrier_vec n" "(ws ! j) \ carrier_vec n" using W len i j by auto have "(ws ! i) \c (ws ! i) \ 0" "(ws ! j) \c (ws ! j) \ 0" using orth corthogonal_def[of ws] len i j by auto then have neq0: "(ws ! i) \ 0\<^sub>v n" "(ws ! j) \ 0\<^sub>v n" by (auto simp add: conjugate_square_eq_0_vec[of "ws ! i" n]) then have "vec_norm (ws ! i) > 0" "vec_norm (ws ! j) > 0" using vec_norm_ge_0 dimws by auto - then have ge0: "vec_norm (ws ! i) * vec_norm (ws ! j) > 0" by auto + then have ge0: "vec_norm (ws ! i) * vec_norm (ws ! j) > 0" by (auto simp: less_complex_def) have ws': "vs ! i = vec_normalize (ws ! i)" "vs ! j = vec_normalize (ws ! j)" using len i j vs_def by auto have ii1: "(vs ! i) \c (vs ! i) = 1" apply (simp add: ws') apply (rule normalized_vec_norm[of "ws ! i"], rule dimws, rule neq0) done have ij0: "i \ j \ (ws ! i) \c (ws ! j) = 0" using i j by (insert orth, auto simp add: corthogonal_def[of ws] len) have "i \ j \ (ws ! i) \c (ws ! j) = (vec_norm (ws ! i) * vec_norm (ws ! j)) * ((vs ! i) \c (vs ! j))" apply (auto simp add: ws') apply (rule normalized_cscalar_prod) apply (rule dimws, rule dimws) done with ij0 have ij0': "i \ j \ (vs ! i) \c (vs ! j) = 0" using ge0 by auto have cWk: "\k. k < n \ col W k = vs ! k" unfolding W_def apply (subst col_mat_of_cols) apply (auto simp add: W'') done have "(mat_adjoint W * W) $$ (j, i) = row (mat_adjoint W) j \ col W i" by (insert dimW i j dimaW, auto) also have "\ = conjugate (col W j) \ col W i" by (insert dimW i j dimaW, auto simp add: mat_adjoint_def) also have "\ = col W i \ conjugate (col W j)" using comm_scalar_prod[of "col W i" n] dimW by auto also have "\ = (vs ! i) \c (vs ! j)" using W_def col_mat_of_cols i j len cWk by auto finally have "(mat_adjoint W * W) $$ (j, i) = (vs ! i) \c (vs ! j)". then have "(mat_adjoint W * W) $$ (j, i) = (if (j = i) then 1 else 0)" by (auto simp add: ii1 ij0') } note maWW = this then have "mat_adjoint W * W = 1\<^sub>m n" unfolding one_mat_def using dimW dimaW by (auto simp add: maWW adjoint_def) then have iv0: "adjoint W * W = 1\<^sub>m n" by auto have dimaW: "adjoint W \ carrier_mat n n" using dimaW by auto then have iv1: "W * adjoint W = 1\<^sub>m n" using mat_mult_left_right_inverse dimW iv0 by auto then show "unitary W" unfolding unitary_def inverts_mat_def using dimW dimaW iv0 iv1 by auto qed lemma normalize_keep_eigenvector: assumes ev: "eigenvector A v e" and dim: "A \ carrier_mat n n" "v \ carrier_vec n" shows "eigenvector A (vec_normalize v) e" unfolding eigenvector_def proof show "vec_normalize v \ carrier_vec (dim_row A)" using dim by auto have eg: "A *\<^sub>v v = e \\<^sub>v v" using ev dim eigenvector_def by auto have vneq0: "v \ 0\<^sub>v n" using ev dim unfolding eigenvector_def by auto then have s0: "vec_normalize v \ 0\<^sub>v n" by (insert dim, subst normalize_zero[of v], auto) from vneq0 have vvge0: "vec_norm v > 0" using vec_norm_ge_0 dim by auto have s1: "A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" unfolding vec_normalize_def using vneq0 dim apply (auto, simp add: mult_mat_vec) apply (subst eg, auto) done with s0 dim show "vec_normalize v \ 0\<^sub>v (dim_row A) \ A *\<^sub>v vec_normalize v = e \\<^sub>v vec_normalize v" by auto qed lemma four_block_mat_adjoint: fixes A B C D :: "'a::conjugatable_field mat" assumes dim: "A \ carrier_mat nr1 nc1" "B \ carrier_mat nr1 nc2" "C \ carrier_mat nr2 nc1" "D \ carrier_mat nr2 nc2" shows "adjoint (four_block_mat A B C D) = four_block_mat (adjoint A) (adjoint C) (adjoint B) (adjoint D)" by (rule eq_matI, insert dim, auto simp add: adjoint_eval) fun unitary_schur_decomposition :: "complex mat \ complex list \ complex mat \ complex mat \ complex mat" where "unitary_schur_decomposition A [] = (A, 1\<^sub>m (dim_row A), 1\<^sub>m (dim_row A))" | "unitary_schur_decomposition A (e # es) = (let n = dim_row A; n1 = n - 1; v' = find_eigenvector A e; v = vec_normalize v'; ws0 = gram_schmidt n (basis_completion v); ws = map vec_normalize ws0; W = mat_of_cols n ws; W' = corthogonal_inv W; A' = W' * A * W; (A1,A2,A0,A3) = split_block A' 1 1; (B,P,Q) = unitary_schur_decomposition A3 es; z_row = (0\<^sub>m 1 n1); z_col = (0\<^sub>m n1 1); one_1 = 1\<^sub>m 1 in (four_block_mat A1 (A2 * P) A0 B, W * four_block_mat one_1 z_row z_col P, four_block_mat one_1 z_row z_col Q * W'))" theorem unitary_schur_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P Q \ upper_triangular B \ diag_mat B = es \ unitary P \ (Q = adjoint P)" using assms proof (induct es arbitrary: n A B P Q) case Nil with degree_monic_char_poly[of A n] show ?case by (auto intro: similar_mat_wit_refl simp: diag_mat_def unitary_zero) next case (Cons e es n A C P Q) let ?n1 = "n - 1" from Cons have A: "A \ carrier_mat n n" and dim: "dim_row A = n" by auto let ?cp = "char_poly A" from Cons(3) have cp: "?cp = [: -e, 1 :] * (\e \ es. [:- e, 1:])" by auto have mon: "monic (\e\ es. [:- e, 1:])" by (rule monic_prod_list, auto) have deg: "degree ?cp = Suc (degree (\e\ es. [:- e, 1:]))" unfolding cp by (subst degree_mult_eq, insert mon, auto) with degree_monic_char_poly[OF A] have n: "n \ 0" by auto define v' where "v' = find_eigenvector A e" define v where "v = vec_normalize v'" define b where "b = basis_completion v" define ws0 where "ws0 = gram_schmidt n b" define ws where "ws = map vec_normalize ws0" define W where "W = mat_of_cols n ws" define W' where "W' = corthogonal_inv W" define A' where "A' = W' * A * W" obtain A1 A2 A0 A3 where splitA': "split_block A' 1 1 = (A1,A2,A0,A3)" by (cases "split_block A' 1 1", auto) obtain B P' Q' where schur: "unitary_schur_decomposition A3 es = (B,P',Q')" by (cases "unitary_schur_decomposition A3 es", auto) let ?P' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) P'" let ?Q' = "four_block_mat (1\<^sub>m 1) (0\<^sub>m 1 ?n1) (0\<^sub>m ?n1 1) Q'" have C: "C = four_block_mat A1 (A2 * P') A0 B" and P: "P = W * ?P'" and Q: "Q = ?Q' * W'" using Cons(4) unfolding unitary_schur_decomposition.simps Let_def list.sel dim v'_def[symmetric] v_def[symmetric] b_def[symmetric] ws0_def[symmetric] ws_def[symmetric] W'_def[symmetric] W_def[symmetric] A'_def[symmetric] split splitA' schur by auto have e: "eigenvalue A e" unfolding eigenvalue_root_char_poly[OF A] cp by simp from find_eigenvector[OF A e] have ev': "eigenvector A v' e" unfolding v'_def . then have "v' \ carrier_vec n" unfolding eigenvector_def using A by auto with ev' have ev: "eigenvector A v e" unfolding v_def using A dim normalize_keep_eigenvector by auto from this[unfolded eigenvector_def] have v[simp]: "v \ carrier_vec n" and v0: "v \ 0\<^sub>v n" using A by auto interpret cof_vec_space n "TYPE(complex)" . from basis_completion[OF v v0, folded b_def] have span_b: "span (set b) = carrier_vec n" and dist_b: "distinct b" and indep: "\ lin_dep (set b)" and b: "set b \ carrier_vec n" and hdb: "hd b = v" and len_b: "length b = n" by auto from hdb len_b n obtain vs where bv: "b = v # vs" by (cases b, auto) from gram_schmidt_result[OF b dist_b indep refl, folded ws0_def] have ws0: "set ws0 \ carrier_vec n" "corthogonal ws0" "length ws0 = n" by (auto simp: len_b) then have ws: "set ws \ carrier_vec n" "corthogonal ws" "length ws = n" unfolding ws_def using normalize_keep_corthogonal by auto have ws0ne: "ws0 \ []" using \length ws0 = n\ n by auto from gram_schmidt_hd[OF v, of vs, folded bv] have hdws0: "hd ws0 = (vec_normalize v')" unfolding ws0_def v_def . have "hd ws = vec_normalize (hd ws0)" unfolding ws_def using hd_map[OF ws0ne] by auto then have hdws: "hd ws = v" unfolding v_def using normalize_normalize[of v'] hdws0 by auto have orth_W: "corthogonal_mat W" using orthogonal_mat_of_cols ws unfolding W_def. have W: "W \ carrier_mat n n" using ws unfolding W_def using mat_of_cols_carrier(1)[of n ws] by auto have W': "W' \ carrier_mat n n" unfolding W'_def corthogonal_inv_def using W by (auto simp: mat_of_rows_def) from corthogonal_inv_result[OF orth_W] have W'W: "inverts_mat W' W" unfolding W'_def . hence WW': "inverts_mat W W'" using mat_mult_left_right_inverse[OF W' W] W' W unfolding inverts_mat_def by auto have A': "A' \ carrier_mat n n" using W W' A unfolding A'_def by auto have A'A_wit: "similar_mat_wit A' A W' W" by (rule similar_mat_witI[of _ _ n], insert W W' A A' W'W WW', auto simp: A'_def inverts_mat_def) hence A'A: "similar_mat A' A" unfolding similar_mat_def by blast from similar_mat_wit_sym[OF A'A_wit] have simAA': "similar_mat_wit A A' W W'" by auto have eigen[simp]: "A *\<^sub>v v = e \\<^sub>v v" and v0: "v \ 0\<^sub>v n" using v_def v'_def find_eigenvector[OF A e] A normalize_keep_eigenvector unfolding eigenvector_def by auto let ?f = "(\ i. if i = 0 then e else 0)" have col0: "col A' 0 = vec n ?f" unfolding A'_def W'_def W_def using corthogonal_col_ev_0[OF A v v0 eigen n hdws ws]. from A' n have "dim_row A' = 1 + ?n1" "dim_col A' = 1 + ?n1" by auto from split_block[OF splitA' this] have A2: "A2 \ carrier_mat 1 ?n1" and A3: "A3 \ carrier_mat ?n1 ?n1" and A'block: "A' = four_block_mat A1 A2 A0 A3" by auto have A1id: "A1 = mat 1 1 (\ _. e)" using splitA'[unfolded split_block_def Let_def] arg_cong[OF col0, of "\ v. v $ 0"] A' n by (auto simp: col_def) have A1: "A1 \ carrier_mat 1 1" unfolding A1id by auto { fix i assume "i < ?n1" with arg_cong[OF col0, of "\ v. v $ Suc i"] A' have "A' $$ (Suc i, 0) = 0" by auto } note A'0 = this have A0id: "A0 = 0\<^sub>m ?n1 1" using splitA'[unfolded split_block_def Let_def] A'0 A' by auto have A0: "A0 \ carrier_mat ?n1 1" unfolding A0id by auto from cp char_poly_similar[OF A'A] have cp: "char_poly A' = [: -e,1 :] * (\ e \ es. [:- e, 1:])" by simp also have "char_poly A' = char_poly A1 * char_poly A3" unfolding A'block A0id by (rule char_poly_four_block_zeros_col[OF A1 A2 A3]) also have "char_poly A1 = [: -e,1 :]" by (simp add: A1id char_poly_defs det_def) finally have cp: "char_poly A3 = (\ e \ es. [:- e, 1:])" by (metis mult_cancel_left pCons_eq_0_iff zero_neq_one) from Cons(1)[OF A3 cp schur] have simIH: "similar_mat_wit A3 B P' Q'" and ut: "upper_triangular B" and diag: "diag_mat B = es" and uP': "unitary P'" and Q'P': "Q' = adjoint P'" by auto from similar_mat_witD2[OF A3 simIH] have B: "B \ carrier_mat ?n1 ?n1" and P': "P' \ carrier_mat ?n1 ?n1" and Q': "Q' \ carrier_mat ?n1 ?n1" and PQ': "P' * Q' = 1\<^sub>m ?n1" by auto have A0_eq: "A0 = P' * A0 * 1\<^sub>m 1" unfolding A0id using P' by auto have simA'C: "similar_mat_wit A' C ?P' ?Q'" unfolding A'block C by (rule similar_mat_wit_four_block[OF similar_mat_wit_refl[OF A1] simIH _ A0_eq A1 A3 A0], insert PQ' A2 P' Q', auto) have ut1: "upper_triangular A1" unfolding A1id by auto have ut: "upper_triangular C" unfolding C A0id by (intro upper_triangular_four_block[OF _ B ut1 ut], auto simp: A1id) from A1id have diagA1: "diag_mat A1 = [e]" unfolding diag_mat_def by auto from diag_four_block_mat[OF A1 B] have diag: "diag_mat C = e # es" unfolding diag diagA1 C by simp have aW: "adjoint W \ carrier_mat n n" using W by auto have aW': "adjoint W' \ carrier_mat n n" using W' by auto have "unitary W" using W_def ws_def ws0 normalized_corthogonal_mat_is_unitary by auto then have ivWaW: "inverts_mat W (adjoint W)" using unitary_def W aW by auto with WW' have W'aW: "W' = (adjoint W)" using inverts_mat_unique W W' aW by auto then have "adjoint W' = W" using adjoint_adjoint by auto with ivWaW have "inverts_mat W' (adjoint W')" using inverts_mat_symm W aW W'aW by auto then have "unitary W'" using unitary_def W' by auto have newP': "P' \ carrier_mat (n - Suc 0) (n - Suc 0)" using P' by auto have rl: "\ x1 x2 x3 x4 y1 y2 y3 y4 f. x1 = y1 \ x2 = y2 \ x3 = y3 \ x4 = y4 \ f x1 x2 x3 x4 = f y1 y2 y3 y4" by simp have Q'aP': "?Q' = adjoint ?P'" apply (subst four_block_mat_adjoint, auto simp add: newP') apply (rule rl[where f2 = four_block_mat]) apply (auto simp add: eq_matI adjoint_eval Q'P') done have "adjoint P = adjoint ?P' * adjoint W" using W newP' n apply (simp add: P) apply (subst adjoint_mult[of W, symmetric]) apply (auto simp add: W P' carrier_matD[of W n n]) done also have "\ = ?Q' * W'" using Q'aP' W'aW by auto also have "\ = Q" using Q by auto finally have QaP: "Q = adjoint P" .. from similar_mat_wit_trans[OF simAA' simA'C, folded P Q] have smw: "similar_mat_wit A C P Q" by blast then have dimP: "P \ carrier_mat n n" and dimQ: "Q \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto from smw have "P * Q = 1\<^sub>m n" unfolding similar_mat_wit_def using A by auto then have "inverts_mat P Q" using inverts_mat_def dimP by auto then have uP: "unitary P" using QaP unitary_def dimP by auto from ut similar_mat_wit_trans[OF simAA' simA'C, folded P Q] diag uP QaP show ?case by blast qed lemma complex_mat_char_poly_factorizable: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" proof - let ?ca = "char_poly A" have ex0: "\bs. Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by (simp add: fundamental_theorem_algebra_factorized) then obtain bs where " Polynomial.smult (lead_coeff ?ca) (\b\bs. [:- b, 1:]) = ?ca \ length bs = degree ?ca" by auto moreover have "lead_coeff ?ca = (1::complex)" using assms degree_monic_char_poly by blast ultimately have ex1: "?ca = (\b\bs. [:- b, 1:]) \ length bs = degree ?ca" by auto moreover have "degree ?ca = n" by (simp add: assms degree_monic_char_poly) ultimately show ?thesis by auto qed lemma complex_mat_has_unitary_schur_decomposition: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\B P es. similar_mat_wit A B P (adjoint P) \ unitary P \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" proof - have "\es. char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" using assms by (simp add: complex_mat_char_poly_factorizable) then obtain es where es: "char_poly A = (\ e \ es. [:- e, 1:]) \ length es = n" by auto obtain B P Q where B: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) have "similar_mat_wit A B P Q \ upper_triangular B \ unitary P \ (Q = adjoint P) \ char_poly A = (\ (e :: complex) \ es. [:- e, 1:]) \ diag_mat B = es" using assms es B by (auto simp add: unitary_schur_decomposition) then show ?thesis by auto qed lemma normal_upper_triangular_matrix_is_diagonal: fixes A :: "'a::conjugatable_ordered_field mat" assumes "A \ carrier_mat n n" and tri: "upper_triangular A" and norm: "A * adjoint A = adjoint A * A" shows "diagonal_mat A" proof (rule disjE[of "n = 0" "n > 0"], blast) have dim: "dim_row A = n" "dim_col A = n" using assms by auto from norm have eq0: "\i j. (A * adjoint A)$$(i,j) = (adjoint A * A)$$(i,j)" by auto have nat_induct_strong: "\P. (P::nat\bool) 0 \ (\i. i < n \ (\k. k < i \ P k) \ P i) \ (\i. i < n \ P i)" by (metis dual_order.strict_trans infinite_descent0 linorder_neqE_nat) show "n = 0 \ ?thesis" using dim unfolding diagonal_mat_def by auto show "n > 0 \ ?thesis" unfolding diagonal_mat_def dim apply (rule allI, rule impI) apply (rule nat_induct_strong) proof (rule allI, rule impI, rule impI) assume asm: "n > 0" from tri upper_triangularD[of A 0 j] dim have z0: "\j. 0< j \ j < n \ A$$(j, 0) = 0" by auto then have ada00: "(adjoint A * A)$$(0,0) = conjugate (A$$(0,0)) * A$$(0,0)" using asm dim by (auto simp add: scalar_prod_def adjoint_eval sum.atLeast_Suc_lessThan) have aad00: "(A * adjoint A)$$(0,0) = (\k=0.. = A$$(0,0) * conjugate (A$$(0,0)) + (\k=1..k. A$$(0, k) * conjugate (A$$(0, k))"], auto) ultimately have f1tneq0: "(\k=(Suc 0)..k. k < n \ A$$(0, k) * conjugate (A$$(0, k)) \ 0" using conjugate_square_positive by auto have "\k. 1 \ k \ k < n \ A$$(0, k) * conjugate (A$$(0, k)) = 0" by (rule sum_nonneg_0[of "{1..j. 0 < n \ j < n \ 0 \ j \ A $$ (0, j) = 0" by auto { fix i assume asm: "n > 0" "i < n" "i > 0" and ih: "\k. k < i \ \j j \ A $$ (k, j) = 0" then have "\j. j i \ j \ A $$ (i, j) = 0" proof - have inter_part: "\b m e. (b::nat) < e \ b < m \ m < e \ {b.. {m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k\{b..{m..b m e f. (b::nat) < e \ b < m \ m < e \ (\k=b..k=b..k=m..j. j < i \ A$$(i, j) = 0" by auto from tri upper_triangularD[of A j i] asm dim have zsi1: "\k. i < k \ k < n \ A$$(k, i) = 0" by auto have "(A * adjoint A)$$(i, i) = (\k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i) + (\k=(Suc i)..k=(Suc i)..k=0.. = (\k=0..k=i.. = (\k=i.. = conjugate (A$$(i, i)) * A$$(i, i)" using asm zsi1 by (auto simp add: sum.atLeast_Suc_lessThan) finally have "(adjoint A * A)$$(i, i) = conjugate (A$$(i, i)) * A$$(i, i)" . with adaii eq0 have fsitoneq0: "(\k=(Suc i)..k. k i < k \ conjugate (A$$(i, k)) * A$$(i, k) = 0" by (rule sum_nonneg_0[of "{(Suc i)..k. k i A $$ (i, k) = 0" by auto with zsi0 show "\j. j i \ j \ A $$ (i, j) = 0" by (metis linorder_neqE_nat) qed } with case0 show "\i ia. 0 < n \ i < n \ ia < n \ (\k. k < ia \ \j j \ A $$ (k, j) = 0) \ \j j \ A $$ (ia, j) = 0" by auto qed qed lemma normal_complex_mat_has_spectral_decomposition: assumes A: "(A::complex mat) \ carrier_mat n n" and normal: "A * adjoint A = adjoint A * A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" proof - have smw: "similar_mat_wit A B P (adjoint P)" and ut: "upper_triangular B" and uP: "unitary P" and dB: "diag_mat B = es" and "(Q = adjoint P)" using assms by (auto simp add: unitary_schur_decomposition) from smw have dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def using A by auto have dimaB: "adjoint B \ carrier_mat n n" using dimB by auto note dims = dimP dimB dimaP dimaB have "inverts_mat P (adjoint P)" using unitary_def uP dims by auto then have iaPP: "inverts_mat (adjoint P) P" using inverts_mat_symm using dims by auto have aPP: "adjoint P * P = 1\<^sub>m n" using dims iaPP unfolding inverts_mat_def by auto from smw have A: "A = P * B * (adjoint P)" unfolding similar_mat_wit_def Let_def by auto then have aA: "adjoint A = P * adjoint B * adjoint P" by (insert A dimP dimB dimaP, auto simp add: adjoint_mult[of _ n n _ n] adjoint_adjoint) have "A * adjoint A = (P * B * adjoint P) * (P * adjoint B * adjoint P)" using A aA by auto also have "\ = P * B * (adjoint P * P) * (adjoint B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * B * 1\<^sub>m n * (adjoint B * adjoint P)" using dims aPP by (auto) also have "\ = P * B * adjoint B * adjoint P" using dims by (mat_assoc n) finally have "A * adjoint A = P * B * adjoint B * adjoint P". then have "adjoint P * (A * adjoint A) * P = (adjoint P * P) * B * adjoint B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * B * adjoint B * 1\<^sub>m n" using aPP by auto also have "\ = B * adjoint B" using dims by auto finally have eq0: "adjoint P * (A * adjoint A) * P = B * adjoint B". have "adjoint A * A = (P * adjoint B * adjoint P) * (P * B * adjoint P)" using A aA by auto also have "\ = P * adjoint B * (adjoint P * P) * (B * adjoint P)" using dims by (mat_assoc n) also have "\ = P * adjoint B * 1\<^sub>m n * (B * adjoint P)" using dims aPP by (auto) also have "\ = P * adjoint B * B * adjoint P" using dims by (mat_assoc n) finally have "adjoint A * A = P * adjoint B * B * adjoint P" by auto then have "adjoint P * (adjoint A * A) * P = (adjoint P * P) * adjoint B * B * (adjoint P * P)" using dims by (simp add: assoc_mult_mat[of _ n n _ n _ n]) also have "\ = 1\<^sub>m n * adjoint B * B * 1\<^sub>m n" using aPP by auto also have "\ = adjoint B * B" using dims by auto finally have eq1: "adjoint P * (adjoint A * A) * P = adjoint B * B". from normal have "adjoint P * (adjoint A * A) * P = adjoint P * (A * adjoint A) * P" by auto with eq0 eq1 have "B * adjoint B = adjoint B * B" by auto with ut dims have "diagonal_mat B" using normal_upper_triangular_matrix_is_diagonal by auto with smw uP dB show "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" by auto qed lemma complex_mat_has_jordan_nf: fixes A :: "complex mat" assumes "A \ carrier_mat n n" shows "\n_as. jordan_nf A n_as" proof - have "\as. char_poly A = (\ a \ as. [:- a, 1:]) \ length as = n" using assms by (simp add: complex_mat_char_poly_factorizable) then show ?thesis using assms by (auto simp add: jordan_nf_iff_linear_factorization) qed lemma hermitian_is_normal: assumes "hermitian A" shows "A * adjoint A = adjoint A * A" using assms by (auto simp add: hermitian_def) lemma hermitian_eigenvalue_real: assumes dim: "(A::complex mat) \ carrier_mat n n" and hA: "hermitian A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" proof - have normal: "A * adjoint A = adjoint A * A" using hA hermitian_is_normal by auto then have schur: "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P" using normal_complex_mat_has_spectral_decomposition[OF dim normal c B] by (simp) then have "similar_mat_wit A B P (adjoint P)" and uP: "unitary P" and dB: "diag_mat B = es" using assms by auto then have A: "A = P * B * (adjoint P)" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dim by auto then have dimaB: "adjoint B \ carrier_mat n n" by auto have "adjoint A = adjoint (adjoint P) * adjoint (P * B)" apply (subst A) apply (subst adjoint_mult[of "P * B" n n "adjoint P" n]) apply (insert dimB dimP, auto) done also have "\ = P * adjoint (P * B)" by (auto simp add: adjoint_adjoint) also have "\ = P * (adjoint B * adjoint P)" using dimB dimP by (auto simp add: adjoint_mult) also have "\ = P * adjoint B * adjoint P" using dimB dimP by (subst assoc_mult_mat[symmetric, of P n n "adjoint B" n "adjoint P" n], auto) finally have aA: "adjoint A = P * adjoint B * adjoint P" . have "A = adjoint A" using hA hermitian_def[of A] by auto then have "P * B * adjoint P = P * adjoint B * adjoint P" using A aA by auto then have BaB: "B = adjoint B" using unitary_elim[OF dimB dimaB dimP] uP by auto { fix i assume "i < n" then have "B$$(i, i) = conjugate (B$$(i, i))" apply (subst BaB) by (insert dimB, simp add: adjoint_eval) then have "B$$(i, i) \ Reals" unfolding conjugate_complex_def using Reals_cnj_iff by auto } then have "\i Reals" by auto with schur show ?thesis by auto qed lemma hermitian_inner_prod_real: assumes dimA: "(A::complex mat) \ carrier_mat n n" and dimv: "v \ carrier_vec n" and hA: "hermitian A" shows "inner_prod v (A *\<^sub>v v) \ Reals" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto define w where "w = (adjoint P) *\<^sub>v v" then have dimw: "w \ carrier_vec n" using dimaP dimv by auto from A have "inner_prod v (A *\<^sub>v v) = inner_prod v ((P * B * (adjoint P)) *\<^sub>v v)" by auto also have "\ = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod w (B *\<^sub>v w)" unfolding w_def apply (rule adjoint_def_alter[OF _ _ dimP]) apply (insert mult_mat_vec_carrier[OF dimB mult_mat_vec_carrier[OF dimaP dimv]], auto simp add: dimv) done also have "\ = (\i=0..j=0.. = (\i=0..v v) = (\i=0..i. i < n \ B$$(i, i) * conjugate (w$i) * w$i \ Reals" using Bii by (simp add: Reals_cnj_iff) then have "(\i=0.. Reals" by auto then show ?thesis using sum by auto qed lemma unit_vec_bracket: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and i: "i < n" shows "inner_prod (unit_vec n i) (A *\<^sub>v (unit_vec n i)) = A$$(i, i)" proof - define w where "(w::complex vec) = unit_vec n i" have "A *\<^sub>v w = col A i" using i dimA w_def by auto then have 1: "inner_prod w (A *\<^sub>v w) = inner_prod w (col A i)" using w_def by auto have "conjugate w = w" unfolding w_def unit_vec_def conjugate_vec_def using i by auto then have 2: "inner_prod w (col A i) = A$$(i, i)" using i dimA w_def by auto from 1 2 show "inner_prod w (A *\<^sub>v w) = A$$(i, i)" by auto qed lemma spectral_decomposition_extract_diag: fixes P B :: "complex mat" assumes dimP: "P \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and uP: "unitary P" and dB: "diagonal_mat B" and i: "i < n" shows "inner_prod (col P i) (P * B * (adjoint P) *\<^sub>v (col P i)) = B$$(i, i)" proof - have dimaP: "adjoint P\ carrier_mat n n" using dimP by auto have uaP: "unitary (adjoint P)" using unitary_adjoint uP dimP by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "(adjoint P) * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto define v where "v = col P i" then have dimv: "v \ carrier_vec n" using dimP by auto define w where "(w::complex vec) = unit_vec n i" then have dimw: "w \ carrier_vec n" by auto have BaPv: "B *\<^sub>v (adjoint P *\<^sub>v v) \ carrier_vec n" using dimB dimaP dimv by auto have "(adjoint P) *\<^sub>v v = (col (adjoint P * P) i)" by (simp add: col_mult2[OF dimaP dimP i, symmetric] v_def) then have aPv: "(adjoint P) *\<^sub>v v = w" by (auto simp add: iv i w_def) have "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = inner_prod v ((P * B) *\<^sub>v ((adjoint P) *\<^sub>v v))" using dimP dimB dimv by (subst assoc_mult_mat_vec[of _ n n "adjoint P" n], auto) also have "\ = inner_prod v (P *\<^sub>v (B *\<^sub>v ((adjoint P) *\<^sub>v v)))" using dimP dimB dimv dimaP by (subst assoc_mult_mat_vec[of _ n n "B" n], auto) also have "\ = inner_prod (adjoint P *\<^sub>v v) (B *\<^sub>v (adjoint P *\<^sub>v v))" by (simp add: adjoint_def_alter[OF dimv BaPv dimP]) also have "\ = inner_prod w (B *\<^sub>v w)" using aPv by auto also have "\ = B$$(i, i)" using w_def unit_vec_bracket dimB i by auto finally show "inner_prod v (P * B * (adjoint P) *\<^sub>v v) = B$$(i, i)". qed lemma hermitian_inner_prod_zero: fixes A :: "complex mat" assumes dimA: "A \ carrier_mat n n" and hA: "hermitian A" and zero: "\v\carrier_vec n. inner_prod v (A *\<^sub>v v) = 0" shows "A = 0\<^sub>m n n" proof - obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto then have uaP: "unitary (adjoint P)" using unitary_adjoint by auto then have "inverts_mat (adjoint P) P" by (simp add: unitary_def adjoint_adjoint) then have iv: "adjoint P * P = 1\<^sub>m n" using dimaP inverts_mat_def by auto have "B = 0\<^sub>m n n" proof- { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) = 0" using dimv zero by auto ultimately have "B$$(i, i) = 0" by auto } note zB = this show "B = 0\<^sub>m n n" by (insert zB dB dimB, rule eq_matI, auto simp add: diagonal_mat_def) qed then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma complex_mat_decomposition_to_hermitian: fixes A :: "complex mat" assumes dim: "A \ carrier_mat n n" shows "\B C. hermitian B \ hermitian C \ A = B + \ \\<^sub>m C \ B \ carrier_mat n n \ C \ carrier_mat n n" proof - obtain B C where B: "B = (1 / 2) \\<^sub>m (A + adjoint A)" and C: "C = (-\ / 2) \\<^sub>m (A - adjoint A)" by auto then have dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" using dim by auto have "hermitian B" unfolding B hermitian_def using dim by (auto simp add: adjoint_eval) moreover have "hermitian C" unfolding C hermitian_def using dim apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done moreover have "A = B + \ \\<^sub>m C" using dim B C apply (subst eq_matI) apply (auto simp add: adjoint_eval algebra_simps) done ultimately show ?thesis using dimB dimC by auto qed subsection \Outer product\ definition outer_prod :: "'a::conjugatable_field vec \ 'a vec \ 'a mat" where "outer_prod v w = mat (dim_vec v) 1 (\(i, j). v $ i) * mat 1 (dim_vec w) (\(i, j). (conjugate w) $ j)" lemma outer_prod_dim[simp]: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" shows "outer_prod v w \ carrier_mat n m" unfolding outer_prod_def using assms mat_of_cols_carrier mat_of_rows_carrier by auto lemma mat_of_vec_mult_eq_scalar_prod: fixes v w :: "'a::conjugatable_field vec" assumes "v \ carrier_vec n" and "w \ carrier_vec n" shows "mat 1 (dim_vec v) (\(i, j). (conjugate v) $ j) * mat (dim_vec w) 1 (\(i, j). w $ i) = mat 1 1 (\k. inner_prod v w)" apply (rule eq_matI) using assms apply (simp add: scalar_prod_def) apply (rule sum.cong) by auto lemma one_dim_mat_mult_is_scale: fixes A B :: "('a::conjugatable_field mat)" assumes "B \ carrier_mat 1 n" shows "(mat 1 1 (\k. a)) * B = a \\<^sub>m B" apply (rule eq_matI) using assms by (auto simp add: scalar_prod_def) lemma outer_prod_mult_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" and d: "d \ carrier_vec d3" shows "outer_prod a b * outer_prod c d = inner_prod b c \\<^sub>m outer_prod a d" proof - let ?ma = "mat (dim_vec a) 1 (\(i, j). a $ i)" let ?mb = "mat 1 (dim_vec b) (\(i, j). (conjugate b) $ j)" let ?mc = "mat (dim_vec c) 1 (\(i, j). c $ i)" let ?md = "mat 1 (dim_vec d) (\(i, j). (conjugate d) $ j)" have "(?ma * ?mb) * (?mc * ?md) = ?ma * (?mb * (?mc * ?md))" apply (subst assoc_mult_mat[of "?ma" d1 1 "?mb" d2 "?mc * ?md" d3] ) using assms by auto also have "\ = ?ma * ((?mb * ?mc) * ?md)" apply (subst assoc_mult_mat[symmetric, of "?mb" 1 d2 "?mc" 1 "?md" d3]) using assms by auto also have "\ = ?ma * ((mat 1 1 (\k. inner_prod b c)) * ?md)" apply (subst mat_of_vec_mult_eq_scalar_prod[of b d2 c]) using assms by auto also have "\ = ?ma * (inner_prod b c \\<^sub>m ?md)" apply (subst one_dim_mat_mult_is_scale) using assms by auto also have "\ = (inner_prod b c) \\<^sub>m (?ma * ?md)" using assms by auto finally show ?thesis unfolding outer_prod_def by auto qed lemma index_outer_prod: fixes v w :: "'a::conjugatable_field vec" assumes v: "v \ carrier_vec n" and w: "w \ carrier_vec m" and ij: "i < n" "j < m" shows "(outer_prod v w)$$(i, j) = v $ i * conjugate (w $ j)" unfolding outer_prod_def using assms by (simp add: scalar_prod_def) lemma mat_of_vec_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" and b: "b \ carrier_vec d" shows "mat 1 d (\(i, j). (conjugate a) $ j) *\<^sub>v b = vec 1 (\k. inner_prod a b)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_vecD[OF a] carrier_vecD[OF b]) apply (rule sum.cong) by auto lemma mat_of_vec_mult_one_dim_vec: fixes a b :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d" shows "mat d 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. c) = c \\<^sub>v a" apply (rule eq_vecI) by (auto simp add: scalar_prod_def carrier_vecD[OF a]) lemma outer_prod_mult_vec: fixes a b c :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec d1" and b: "b \ carrier_vec d2" and c: "c \ carrier_vec d2" shows "outer_prod a b *\<^sub>v c = inner_prod b c \\<^sub>v a" proof - have "outer_prod a b *\<^sub>v c = mat d1 1 (\(i, j). a $ i) * mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c" unfolding outer_prod_def using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v (mat 1 d2 (\(i, j). (conjugate b) $ j) *\<^sub>v c)" apply (subst assoc_mult_mat_vec) using assms by auto also have "\ = mat d1 1 (\(i, j). a $ i) *\<^sub>v vec 1 (\k. inner_prod b c)" using mat_of_vec_mult_vec[of b] assms by auto also have "\ = inner_prod b c \\<^sub>v a" using mat_of_vec_mult_one_dim_vec assms by auto finally show ?thesis by auto qed lemma trace_outer_prod_right: fixes A :: "'a::conjugatable_field mat" and v w :: "'a vec" assumes A: "A \ carrier_mat n n" and v: "v \ carrier_vec n" and w: "w \ carrier_vec n" shows "trace (A * outer_prod v w) = inner_prod w (A *\<^sub>v v)" (is "?lhs = ?rhs") proof - define B where "B = outer_prod v w" then have B: "B \ carrier_mat n n" using assms by auto have "trace(A * B) = (\i = 0..j = 0.. = (\i = 0..j = 0..i = 0..j = 0..i = 0..j = 0.. carrier_vec n" and w: "w \ carrier_vec n" shows "trace (outer_prod v w) = inner_prod w v" (is "?lhs = ?rhs") proof - have "(1\<^sub>m n) * (outer_prod v w) = outer_prod v w" apply (subst left_mult_one_mat) using outer_prod_dim assms by auto moreover have "1\<^sub>m n *\<^sub>v v = v" using assms by auto ultimately show ?thesis using trace_outer_prod_right[of "1\<^sub>m n" n v w] assms by auto qed lemma inner_prod_outer_prod: fixes a b c d :: "'a::conjugatable_field vec" assumes a: "a \ carrier_vec n" and b: "b \ carrier_vec n" and c: "c \ carrier_vec m" and d: "d \ carrier_vec m" shows "inner_prod a (outer_prod b c *\<^sub>v d) = inner_prod a b * inner_prod c d" (is "?lhs = ?rhs") proof - define P where "P = outer_prod b c" then have dimP: "P \ carrier_mat n m" using assms by auto have "inner_prod a (P *\<^sub>v d) = (\i=0..j=0.. = (\i=0..j=0..i=0..j=0..i=0..j=0.. = (\i=0..j=0..Semi-definite matrices\ definition positive :: "complex mat \ bool" where "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) \ 0)" lemma positive_iff_normalized_vec: "positive A \ A \ carrier_mat (dim_col A) (dim_col A) \ (\v. (dim_vec v = dim_col A \ vec_norm v = 1) \ inner_prod v (A *\<^sub>v v) \ 0)" proof (rule) assume "positive A" then show "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" unfolding positive_def by auto next define n where "n = dim_col A" assume "A \ carrier_mat (dim_col A) (dim_col A) \ (\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v))" then have A: "A \ carrier_mat (dim_col A) (dim_col A)" and geq0: "\v. dim_vec v = dim_col A \ vec_norm v = 1 \ 0 \ inner_prod v (A *\<^sub>v v)" by auto then have dimA: "A \ carrier_mat n n" using n_def[symmetric] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have "0 \ inner_prod v (A *\<^sub>v v)" proof (cases "v = 0\<^sub>v n") case True then show "0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto next case False then have 1: "vec_norm v > 0" using vec_norm_ge_0 dimv by auto - then have cnv: "cnj (vec_norm v) = vec_norm v" using Reals_cnj_iff complex_is_Real_iff by auto + then have cnv: "cnj (vec_norm v) = vec_norm v" + using Reals_cnj_iff complex_is_Real_iff less_complex_def by auto define w where "w = vec_normalize v" then have dimw: "w \ carrier_vec n" using dimv by auto have nvw: "v = vec_norm v \\<^sub>v w" using w_def vec_eq_norm_smult_normalized by auto have "vec_norm w = 1" using normalized_vec_norm[OF dimv False] vec_norm_def w_def by auto then have 2: "0 \ inner_prod w (A *\<^sub>v w)" using geq0 dimw dimA by auto have "inner_prod v (A *\<^sub>v v) = vec_norm v * vec_norm v * inner_prod w (A *\<^sub>v w)" using dimA dimv dimw apply (subst (1 2) nvw) apply (subst mult_mat_vec, simp, simp) apply (subst scalar_prod_smult_left[of "(A *\<^sub>v w)" "conjugate (vec_norm v \\<^sub>v w)" "vec_norm v"], simp) apply (simp add: conjugate_smult_vec cnv) done also have "\ \ 0" using 1 2 by auto finally show "0 \ inner_prod v (A *\<^sub>v v)" by auto qed } then have geq: "\v. dim_vec v = dim_col A \ 0 \ inner_prod v (A *\<^sub>v v)" using dimA by auto show "positive A" unfolding positive_def by (rule, simp add: A, rule geq) qed lemma positive_is_hermitian: fixes A :: "complex mat" assumes pA: "positive A" shows "hermitian A" proof - define n where "n = dim_col A" then have dimA: "A \ carrier_mat n n" using positive_def pA by auto obtain B C where B: "hermitian B" and C: "hermitian C" and A: "A = B + \ \\<^sub>m C" and dimB: "B \ carrier_mat n n" and dimC: "C \ carrier_mat n n" and dimiC: "\ \\<^sub>m C \ carrier_mat n n" using complex_mat_decomposition_to_hermitian[OF dimA] by auto { fix v :: "complex vec" assume dimv: "v \ carrier_vec n" have dimvA: "dim_vec v = dim_col A" using dimv dimA by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + inner_prod v ((\ \\<^sub>m C) *\<^sub>v v)" unfolding A using dimB dimiC dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) moreover have "inner_prod v ((\ \\<^sub>m C) *\<^sub>v v) = \ * inner_prod v (C *\<^sub>v v)" using dimv dimC apply (simp add: scalar_prod_def sum_distrib_left cong: sum.cong) apply (rule sum.cong, auto) done ultimately have ABC: "inner_prod v (A *\<^sub>v v) = inner_prod v (B *\<^sub>v v) + \ * inner_prod v (C *\<^sub>v v)" by auto moreover have "inner_prod v (B *\<^sub>v v) \ Reals" using B dimB dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (C *\<^sub>v v) \ Reals" using C dimC dimv hermitian_inner_prod_real by auto moreover have "inner_prod v (A *\<^sub>v v) \ Reals" using pA unfolding positive_def apply (rule) apply (fold n_def) apply (simp add: complex_is_Real_iff[of "inner_prod v (A *\<^sub>v v)"]) - apply (auto simp add: dimvA) + apply (auto simp add: dimvA less_complex_def less_eq_complex_def) done ultimately have "inner_prod v (C *\<^sub>v v) = 0" using of_real_Re by fastforce } then have "C = 0\<^sub>m n n" using hermitian_inner_prod_zero dimC C by auto then have "A = B" using A dimC dimB by auto then show "hermitian A" using B by auto qed lemma positive_eigenvalue_positive: assumes dimA: "(A::complex mat) \ carrier_mat n n" and pA: "positive A" and c: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" and B: "unitary_schur_decomposition A es = (B,P,Q)" shows "\i. i < n \ B$$(i, i) \ 0" proof - have hA: "hermitian A" using positive_is_hermitian pA by auto have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA hA B c by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" and uP: "unitary P" unfolding similar_mat_wit_def Let_def unitary_def using dimA by auto { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v (A *\<^sub>v v) = B$$(i, i)" unfolding A v_def using spectral_decomposition_extract_diag[OF dimP dimB uP dB i] by auto moreover have "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA dimA positive_def by auto ultimately show "B$$(i, i) \ 0" by auto } qed lemma diag_mat_mult_diag_mat: fixes B D :: "'a::semiring_0 mat" assumes dimB: "B \ carrier_mat n n" and dimD: "D \ carrier_mat n n" and dB: "diagonal_mat B" and dD: "diagonal_mat D" shows "B * D = mat n n (\(i,j). (if i = j then (B$$(i, i)) * (D$$(i, i)) else 0))" proof(rule eq_matI, auto) have Bij: "\x y. x < n \ y < n \ x \ y \ B$$(x, y) = 0" using dB diagonal_mat_def dimB by auto have Dij: "\x y. x < n \ y < n \ x \ y \ D$$(x, y) = 0" using dD diagonal_mat_def dimD by auto { fix i j assume ij: "i < n" "j < n" have "(B * D) $$ (i, j) = (\k=0.. = B$$(i, i) * D$$(i, j)" apply (simp add: sum.remove[of _i] ij) apply (simp add: Bij Dij ij) done finally have "(B * D) $$ (i, j) = B$$(i, i) * D$$(i, j)". } note BDij = this from BDij show "\j. j < n \ (B * D) $$ (j, j) = B $$ (j, j) * D $$ (j, j)" by auto from BDij show "\i j. i < n \ j < n \ i \ j \ (B * D) $$ (i, j) = 0" using Bij Dij by auto from assms show "dim_row B = n" "dim_col D = n" by auto qed lemma positive_only_if_decomp: assumes dimA: "A \ carrier_mat n n" and pA: "positive A" shows "\M \ carrier_mat n n. M * adjoint M = A" proof - from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ diag_mat B = es \ unitary P \ (\i < n. B$$(i, i) \ Reals)" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and Bii: "\i. i < n \ B$$(i, i) \ Reals" and dimB: "B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto have Bii: "\i. i < n \ B$$(i, i) \ 0" using pA dimA es schur positive_eigenvalue_positive by auto define D where "D = mat n n (\(i, j). (if (i = j) then csqrt (B$$(i, i)) else 0))" then have dimD: "D \ carrier_mat n n" and dimaD: "adjoint D \ carrier_mat n n" using dimB by auto have dD: "diagonal_mat D" using dB D_def unfolding diagonal_mat_def by auto then have daD: "diagonal_mat (adjoint D)" by (simp add: adjoint_eval diagonal_mat_def) have Dii: "\i. i < n \ D$$(i, i) = csqrt (B$$(i, i))" using dimD D_def by auto { fix i assume i: "i < n" define c where "c = csqrt (B$$(i, i))" - have c: "c \ 0" using Bii i c_def by auto + have c: "c \ 0" using Bii i c_def by (auto simp: less_complex_def less_eq_complex_def) then have "conjugate c = c" - using Reals_cnj_iff complex_is_Real_iff by auto + using Reals_cnj_iff complex_is_Real_iff unfolding less_complex_def less_eq_complex_def by auto then have "c * cnj c = B$$(i, i)" using c_def c unfolding conjugate_complex_def by (metis power2_csqrt power2_eq_square) } note cBii = this have "D * adjoint D = mat n n (\(i,j). (if (i = j) then B$$(i, i) else 0))" apply (simp add: diag_mat_mult_diag_mat[OF dimD dimaD dD daD]) apply (rule eq_matI, auto simp add: D_def adjoint_eval cBii) done also have "\ = B" using dimB dB[unfolded diagonal_mat_def] by auto finally have DaDB: "D * adjoint D = B". define M where "M = P * D" then have dimM: "M \ carrier_mat n n" using dimP dimD by auto have "M * adjoint M = (P * D) * (adjoint D * adjoint P)" using M_def adjoint_mult[OF dimP dimD] by auto also have "\ = P * (D * adjoint D) * (adjoint P)" using dimP dimD by (mat_assoc n) also have "\ = P * B * (adjoint P)" using DaDB by auto finally have "M * adjoint M = A" using A by auto with dimM show "\M \ carrier_mat n n. M * adjoint M = A" by auto qed lemma positive_if_decomp: assumes dimA: "A \ carrier_mat n n" and "\M. M * adjoint M = A" shows "positive A" proof - from assms obtain M where M: "M * adjoint M = A" by auto define m where "m = dim_col M" have dimM: "M \ carrier_mat n m" using M dimA m_def by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have dimaM: "adjoint M \ carrier_mat m n" using dimM by auto have dimaMv: "(adjoint M) *\<^sub>v v \ carrier_vec m" using dimaM dimv by auto have "inner_prod v (A *\<^sub>v v) = inner_prod v (M * adjoint M *\<^sub>v v)" using M by auto also have "\ = inner_prod v (M *\<^sub>v (adjoint M *\<^sub>v v))" using assoc_mult_mat_vec dimM dimaM dimv by auto also have "\ = inner_prod (adjoint M *\<^sub>v v) (adjoint M *\<^sub>v v)" using adjoint_def_alter[OF dimv dimaMv dimM] by auto also have "\ \ 0" using self_cscalar_prod_geq_0 by auto finally have "inner_prod v (A *\<^sub>v v) \ 0". } note geq0 = this from dimA geq0 show "positive A" using positive_def by auto qed lemma positive_iff_decomp: assumes dimA: "A \ carrier_mat n n" shows "positive A \ (\M\carrier_mat n n. M * adjoint M = A)" proof assume pA: "positive A" then show "\M\carrier_mat n n. M * adjoint M = A" using positive_only_if_decomp assms by auto next assume "\M\carrier_mat n n. M * adjoint M = A" then obtain M where M: "M * adjoint M = A" by auto then show "positive A" using M positive_if_decomp assms by auto qed lemma positive_dim_eq: assumes "positive A" shows "dim_row A = dim_col A" using carrier_matD(1)[of A "dim_col A" "dim_col A"] assms[unfolded positive_def] by simp lemma positive_zero: "positive (0\<^sub>m n n)" by (simp add: positive_def zero_mat_def mult_mat_vec_def scalar_prod_def) lemma positive_one: "positive (1\<^sub>m n)" proof (rule positive_if_decomp) show "1\<^sub>m n \ carrier_mat n n" by auto have "adjoint (1\<^sub>m n) = 1\<^sub>m n" using hermitian_one hermitian_def by auto then have "1\<^sub>m n * adjoint (1\<^sub>m n) = 1\<^sub>m n" by auto then show "\M. M * adjoint M = 1\<^sub>m n" by fastforce qed lemma positive_antisym: assumes pA: "positive A" and pnA: "positive (-A)" shows "A = 0\<^sub>m (dim_col A) (dim_col A)" proof - define n where "n = dim_col A" from pA have dimA: "A \ carrier_mat n n" and dimnA: "-A \ carrier_mat n n" using positive_def n_def by auto from pA have hA: "hermitian A" using positive_is_hermitian by auto obtain es where es: "char_poly A = (\ (e :: complex) \ es. [:- e, 1:])" using complex_mat_char_poly_factorizable dimA by auto obtain B P Q where schur: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) then have "similar_mat_wit A B P (adjoint P) \ diagonal_mat B \ unitary P" using hermitian_eigenvalue_real dimA es hA by auto then have A: "A = P * B * (adjoint P)" and dB: "diagonal_mat B" and uP: "unitary P" and dimB: "B \ carrier_mat n n" and dimnB: "-B \ carrier_mat n n" and dimP: "P \ carrier_mat n n" and dimaP: "adjoint P \ carrier_mat n n" unfolding similar_mat_wit_def Let_def using dimA by auto from es schur have geq0: "\i. i < n \ B$$(i, i) \ 0" using positive_eigenvalue_positive dimA pA by auto from A have nA: "-A = P * (-B) * (adjoint P)" using mult_smult_assoc_mat dimB dimP dimaP by auto from dB have dnB: "diagonal_mat (-B)" by (simp add: diagonal_mat_def) { fix i assume i: "i < n" define v where "v = col P i" then have dimv: "v \ carrier_vec n" using v_def dimP by auto have "inner_prod v ((-A) *\<^sub>v v) = (-B)$$(i, i)" unfolding nA v_def using spectral_decomposition_extract_diag[OF dimP dimnB uP dnB i] by auto moreover have "inner_prod v ((-A) *\<^sub>v v) \ 0" using dimv pnA dimnA positive_def by auto ultimately have "B$$(i, i) \ 0" using dimB i by auto moreover have "B$$(i, i) \ 0" using i geq0 by auto ultimately have "B$$(i, i) = 0" by (metis no_atp(10)) } then have "B = 0\<^sub>m n n" using dimB dB[unfolded diagonal_mat_def] by (subst eq_matI, auto) then show "A = 0\<^sub>m n n" using A dimB dimP dimaP by auto qed lemma positive_add: assumes pA: "positive A" and pB: "positive B" and dimA: "A \ carrier_mat n n" and dimB: "B \ carrier_mat n n" shows "positive (A + B)" unfolding positive_def proof have dimApB: "A + B \ carrier_mat n n" using dimA dimB by auto then show "A + B \ carrier_mat (dim_col (A + B)) (dim_col (A + B))" using carrier_matD[of "A+B"] by auto { fix v assume dimv: "(v::complex vec) \ carrier_vec n" have 1: "inner_prod v (A *\<^sub>v v) \ 0" using dimv pA[unfolded positive_def] dimA by auto have 2: "inner_prod v (B *\<^sub>v v) \ 0" using dimv pB[unfolded positive_def] dimB by auto have "inner_prod v ((A + B) *\<^sub>v v) = inner_prod v (A *\<^sub>v v) + inner_prod v (B *\<^sub>v v)" using dimA dimB dimv by (simp add: add_mult_distrib_mat_vec inner_prod_distrib_right) also have "\ \ 0" using 1 2 by auto finally have "inner_prod v ((A + B) *\<^sub>v v) \ 0". } note geq0 = this then have "\v. dim_vec v = n \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" by auto then show "\v. dim_vec v = dim_col (A + B) \ 0 \ inner_prod v ((A + B) *\<^sub>v v)" using dimApB by auto qed lemma positive_trace: assumes "A \ carrier_mat n n" and "positive A" shows "trace A \ 0" using assms positive_iff_decomp trace_adjoint_positive by auto lemma positive_close_under_left_right_mult_adjoint: fixes M A :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and pA: "positive A" shows "positive (M * A * adjoint M)" unfolding positive_def proof (rule, simp add: mult_carrier_mat[OF mult_carrier_mat[OF dM dA] adjoint_dim[OF dM]] carrier_matD[OF dM], rule, rule) have daM: "adjoint M \ carrier_mat n n" using dM by auto fix v::"complex vec" assume "dim_vec v = dim_col (M * A * adjoint M)" then have dv: "v \ carrier_vec n" using assms by auto then have "adjoint M *\<^sub>v v \ carrier_vec n" using daM by auto have assoc: "M * A * adjoint M *\<^sub>v v = M *\<^sub>v (A *\<^sub>v (adjoint M *\<^sub>v v))" using dA dM daM dv by (auto simp add: assoc_mult_mat_vec[of _ n n _ n]) have "inner_prod v (M * A * adjoint M *\<^sub>v v) = inner_prod (adjoint M *\<^sub>v v) (A *\<^sub>v (adjoint M *\<^sub>v v))" apply (subst assoc) apply (subst adjoint_def_alter[where ?A = "M"]) by (auto simp add: dv dA daM dM carrier_matD[OF dM] mult_mat_vec_carrier[of _ n n]) also have "\ \ 0" using dA dv daM pA positive_def by auto finally show "inner_prod v (M * A * adjoint M *\<^sub>v v) \ 0" by auto qed lemma positive_same_outer_prod: fixes v w :: "complex vec" assumes v: "v \ carrier_vec n" shows "positive (outer_prod v v)" proof - have d1: "adjoint (mat (dim_vec v) 1 (\(i, j). v $ i)) \ carrier_mat 1 n" using assms by auto have d2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) \ carrier_mat 1 n" using assms by auto have dv: "dim_vec v = n" using assms by auto have "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) = adjoint (mat (dim_vec v) 1 (\(i, j). v $ i))" (is "?r = adjoint ?l") apply (rule eq_matI) subgoal for i j by (simp add: dv adjoint_eval) using d1 d2 by auto then have "outer_prod v v = ?l * adjoint ?l" unfolding outer_prod_def by auto then have "\M. M * adjoint M = outer_prod v v" by auto then show "positive (outer_prod v v)" using positive_if_decomp[OF outer_prod_dim[OF v v]] by auto qed lemma smult_smult_mat: fixes k :: complex and l :: complex assumes "A \ carrier_mat nr n" shows "k \\<^sub>m (l \\<^sub>m A) = (k * l) \\<^sub>m A" by auto lemma positive_smult: assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" proof - - have sc: "csqrt c \ 0" using assms(3) by fastforce + have sc: "csqrt c \ 0" using assms(3) by (fastforce simp: less_eq_complex_def) obtain M where dimM: "M \ carrier_mat n n" and A: "M * adjoint M = A" using assms(1-2) positive_iff_decomp by auto have "c \\<^sub>m A = c \\<^sub>m (M * adjoint M)" using A by auto - have ccsq: "conjugate (csqrt c) = (csqrt c)" using sc Reals_cnj_iff[of "csqrt c"] complex_is_Real_iff by auto + have ccsq: "conjugate (csqrt c) = (csqrt c)" using sc Reals_cnj_iff[of "csqrt c"] complex_is_Real_iff + by (auto simp: less_eq_complex_def) have MM: "(M * adjoint M) \ carrier_mat n n" using A assms by fastforce have leftd: "c \\<^sub>m (M * adjoint M) \ carrier_mat n n" using A assms by fastforce have rightd: "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M))\ carrier_mat n n" using A assms by fastforce have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = (csqrt c \\<^sub>m M) * ((conjugate (csqrt c)) \\<^sub>m adjoint M)" using adjoint_scale assms(1) by (metis adjoint_scale) also have "\ = (csqrt c \\<^sub>m M) * (csqrt c \\<^sub>m adjoint M)" using sc ccsq by fastforce also have "\ = csqrt c \\<^sub>m (M * (csqrt c \\<^sub>m adjoint M))" using mult_smult_assoc_mat index_smult_mat(2,3) by fastforce also have "\ = csqrt c \\<^sub>m ((csqrt c) \\<^sub>m (M * adjoint M))" using mult_smult_distrib by fastforce also have "\ = c \\<^sub>m (M * adjoint M)" using smult_smult_mat[of "M * adjoint M" n n "(csqrt c)" "(csqrt c)"] MM sc by (metis power2_csqrt power2_eq_square ) also have "\ = c \\<^sub>m A" using A by auto finally have "(csqrt c \\<^sub>m M) * (adjoint (csqrt c \\<^sub>m M)) = c \\<^sub>m A" by auto moreover have "c \\<^sub>m A \ carrier_mat n n" using assms(1) by auto moreover have "csqrt c \\<^sub>m M \ carrier_mat n n" using dimM by auto ultimately show ?thesis using positive_iff_decomp by auto qed text \Version of previous theorem for real numbers\ lemma positive_scale: fixes c :: real assumes "A \ carrier_mat n n" and "positive A" and "c \ 0" shows "positive (c \\<^sub>m A)" - apply (rule positive_smult) using assms by auto + apply (rule positive_smult) using assms by (auto simp: less_eq_complex_def) subsection \L\"{o}wner partial order\ definition lowner_le :: "complex mat \ complex mat \ bool" (infix "\\<^sub>L" 50) where "A \\<^sub>L B \ dim_row A = dim_row B \ dim_col A = dim_col B \ positive (B - A)" lemma lowner_le_refl: assumes "A \ carrier_mat n n" shows "A \\<^sub>L A" unfolding lowner_le_def apply (simp add: minus_r_inv_mat[OF assms]) by (rule positive_zero) lemma lowner_le_antisym: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L A" shows "A = B" proof - from L1 have P1: "positive (B - A)" by (simp add: lowner_le_def) from L2 have P2: "positive (A - B)" by (simp add: lowner_le_def) have "A - B = - (B - A)" using A B by auto then have P3: "positive (- (B - A))" using P2 by auto have BA: "B - A \ carrier_mat n n" using A B by auto have "B - A = 0\<^sub>m n n" using BA by (subst positive_antisym[OF P1 P3], auto) then have "B + (-A) + A = 0\<^sub>m n n + A" using A B minus_add_uminus_mat[OF B A] by auto then have "B + (-A + A) = 0\<^sub>m n n + A" using A B by auto then show "A = B" using A B BA uminus_l_inv_mat[OF A] by auto qed lemma lowner_le_inner_prod_le: fixes A B :: "complex mat" and v :: "complex vec" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and v: "v \ carrier_vec n" and "A \\<^sub>L B" shows "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" proof - from assms have "positive (B-A)" by (auto simp add: lowner_le_def) with assms have geq: "inner_prod v ((B-A) *\<^sub>v v) \ 0" unfolding positive_def by auto have "inner_prod v ((B-A) *\<^sub>v v) = inner_prod v (B *\<^sub>v v) - inner_prod v (A *\<^sub>v v)" unfolding minus_add_uminus_mat[OF B A] by (subst add_mult_distrib_mat_vec[OF B _ v], insert A B v, auto simp add: inner_prod_distrib_right[OF v]) then show ?thesis using geq by auto qed lemma lowner_le_trans: fixes A B C :: "complex mat" assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and C: "C \ carrier_mat n n" and L1: "A \\<^sub>L B" and L2: "B \\<^sub>L C" shows "A \\<^sub>L C" unfolding lowner_le_def proof (auto simp add: carrier_matD[OF A] carrier_matD[OF C]) have dim: "C - A \ carrier_mat n n" using A C by auto { fix v assume v: "(v::complex vec) \ carrier_vec n" from L1 have "inner_prod v (A *\<^sub>v v) \ inner_prod v (B *\<^sub>v v)" using lowner_le_inner_prod_le A B v by auto also from L2 have "\ \ inner_prod v (C *\<^sub>v v)" using lowner_le_inner_prod_le B C v by auto finally have "inner_prod v (A *\<^sub>v v) \ inner_prod v (C *\<^sub>v v)". then have "inner_prod v (C *\<^sub>v v) - inner_prod v (A *\<^sub>v v) \ 0" by auto then have "inner_prod v ((C - A) *\<^sub>v v) \ 0" using A C v apply (subst minus_add_uminus_mat[OF C A]) apply (subst add_mult_distrib_mat_vec[OF C _ v], simp) apply (simp add: inner_prod_distrib_right[OF v]) done } note leq = this show "positive (C - A)" unfolding positive_def apply (rule, simp add: carrier_matD[OF A] dim) apply (subst carrier_matD[OF dim], insert leq, auto) done qed lemma lowner_le_imp_trace_le: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "A \\<^sub>L B" shows "trace A \ trace B" proof - have "positive (B - A)" using assms lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "trace (B - A) \ 0" using positive_trace by auto moreover have "trace (B - A) = trace B - trace A" using trace_minus_linear assms by auto ultimately have "trace B - trace A \ 0" by auto then show "trace A \ trace B" by auto qed lemma lowner_le_add: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A + C \\<^sub>L B + D" proof - have "B + D - (A + C) = B - A + (D - C) " using assms by auto then have "positive (B + D - (A + C))" using assms unfolding lowner_le_def using positive_add by (metis minus_carrier_mat) then show "A + C \\<^sub>L B + D" unfolding lowner_le_def using assms by fastforce qed lemma lowner_le_swap: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" and "A \\<^sub>L B" shows "-B \\<^sub>L -A" proof - have "positive (B - A)" using assms lowner_le_def by fastforce moreover have "B - A = (-A) - (-B)" using assms by fastforce ultimately have "positive ((-A) - (-B))" by auto then show ?thesis using lowner_le_def assms by fastforce qed lemma lowner_le_minus: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "C \ carrier_mat n n" "D \ carrier_mat n n" and "A \\<^sub>L B" "C \\<^sub>L D" shows "A - D \\<^sub>L B - C" proof - have "positive (D - C)" using assms lowner_le_def by auto then have "-D \\<^sub>L -C" using lowner_le_swap assms by auto then have "A + (-D) \\<^sub>L B + (-C)" using lowner_le_add[of "A" n "B"] assms by auto moreover have "A + (-D) = A - D" and "B + (-C) = B - C" by auto ultimately show ?thesis by auto qed lemma outer_prod_le_one: assumes "v \ carrier_vec n" and "inner_prod v v \ 1" shows "outer_prod v v \\<^sub>L 1\<^sub>m n" proof - let ?o = "outer_prod v v" have do: "?o \ carrier_mat n n" using assms by auto { fix u :: "complex vec" assume "dim_vec u = n" then have du: "u \ carrier_vec n" by auto have r: "inner_prod u u \ Reals" apply (simp add: scalar_prod_def carrier_vecD[OF du]) using complex_In_mult_cnj_zero complex_is_Real_iff by blast have geq0: "inner_prod u u \ 0" using self_cscalar_prod_geq_0 by auto have "inner_prod u (?o *\<^sub>v u) = inner_prod u v * inner_prod v u" apply (subst inner_prod_outer_prod) using du assms by auto also have "\ \ inner_prod u u * inner_prod v v" using Cauchy_Schwarz_complex_vec du assms by auto also have "\ \ inner_prod u u" using assms(2) r geq0 - by (simp add: mult_right_le_one_le) + by (simp add: mult_right_le_one_le less_eq_complex_def) finally have le: "inner_prod u (?o *\<^sub>v u) \ inner_prod u u". have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) = inner_prod u ((1\<^sub>m n *\<^sub>v u) - ?o *\<^sub>v u)" apply (subst minus_mult_distrib_mat_vec) using do du by auto also have "\ = inner_prod u u - inner_prod u (?o *\<^sub>v u)" apply (subst inner_prod_minus_distrib_right) using du do by auto also have "\ \ 0" using le by auto finally have "inner_prod u ((1\<^sub>m n - ?o) *\<^sub>v u) \ 0" by auto } then have "positive (1\<^sub>m n - outer_prod v v)" unfolding positive_def using do by auto then show ?thesis unfolding lowner_le_def using do by auto qed lemma zero_lowner_le_positiveD: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "0\<^sub>m n n \\<^sub>L A" shows "positive A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma zero_lowner_le_positiveI: fixes A :: "complex mat" assumes dA: "A \ carrier_mat n n" and le: "positive A" shows "0\<^sub>m n n \\<^sub>L A" using assms unfolding lowner_le_def by (subgoal_tac "A - 0\<^sub>m n n = A", auto) lemma lowner_le_trans_positiveI: fixes A B :: "complex mat" assumes dA: "A \ carrier_mat n n" and pA: "positive A" and le: "A \\<^sub>L B" shows "positive B" proof - have dB: "B \ carrier_mat n n" using le dA lowner_le_def by auto have "0\<^sub>m n n \\<^sub>L A" using zero_lowner_le_positiveI dA pA by auto then have "0\<^sub>m n n \\<^sub>L B" using dA dB le by (simp add: lowner_le_trans[of _ n A B]) then show ?thesis using dB zero_lowner_le_positiveD by auto qed lemma lowner_le_keep_under_measurement: fixes M A B :: "complex mat" assumes dM: "M \ carrier_mat n n" and dA: "A \ carrier_mat n n" and dB: "B \ carrier_mat n n" and le: "A \\<^sub>L B" shows "adjoint M * A * M \\<^sub>L adjoint M * B * M" unfolding lowner_le_def proof (rule conjI, fastforce)+ have daM: "adjoint M \ carrier_mat n n" using dM by auto have dBmA: "B - A \ carrier_mat n n" using dB dA by fastforce have "positive (B - A)" using le lowner_le_def by auto then have p: "positive (adjoint M * (B - A) * M)" using positive_close_under_left_right_mult_adjoint[OF daM dBmA] adjoint_adjoint[of M] by auto moreover have e: "adjoint M * (B - A) * M = adjoint M * B * M - adjoint M * A * M" using dM dB dA by (mat_assoc n) ultimately show "positive (adjoint M * B * M - adjoint M * A * M)" by auto qed lemma smult_distrib_left_minus_mat: fixes A B :: "'a::comm_ring_1 mat" assumes "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: minus_add_uminus_mat add_smult_distrib_left_mat) lemma lowner_le_smultc: fixes c :: complex assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" proof - have eqBA: "c \\<^sub>m (B - A) = c \\<^sub>m B - c \\<^sub>m A" using assms by (auto simp add: smult_distrib_left_minus_mat) have "positive (B - A)" using assms(2) unfolding lowner_le_def by auto then have "positive (c \\<^sub>m (B - A))" using positive_smult[of "B-A" n c] assms by fastforce moreover have "c \\<^sub>m A \ carrier_mat n n" using index_smult_mat(2,3) assms(3) by auto moreover have "c \\<^sub>m B \ carrier_mat n n" using index_smult_mat(2,3) assms(4) by auto ultimately show ?thesis unfolding lowner_le_def using eqBA by fastforce qed lemma lowner_le_smult: fixes c :: real assumes "c \ 0" "A \\<^sub>L B" "A \ carrier_mat n n" "B \ carrier_mat n n" shows "c \\<^sub>m A \\<^sub>L c \\<^sub>m B" - apply (rule lowner_le_smultc) using assms by auto + apply (rule lowner_le_smultc) using assms by (auto simp: less_eq_complex_def) lemma minus_smult_vec_distrib: fixes w :: "'a::comm_ring_1 vec" shows "(a - b) \\<^sub>v w = a \\<^sub>v w - b \\<^sub>v w" apply (rule eq_vecI) by (auto simp add: scalar_prod_def algebra_simps) lemma smult_mat_mult_mat_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "a \\<^sub>m A *\<^sub>v w = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma mult_mat_vec_smult_vec_assoc: fixes A :: "'a::comm_ring_1 mat" assumes A: "A \ carrier_mat n m" and w: "w \ carrier_vec m" shows "A *\<^sub>v (a \\<^sub>v w) = a \\<^sub>v (A *\<^sub>v w)" apply (rule eq_vecI) apply (simp add: scalar_prod_def carrier_matD[OF A] carrier_vecD[OF w]) apply (subst sum_distrib_left) apply (rule sum.cong, simp) by auto lemma outer_prod_left_right_mat: fixes A B :: "complex mat" assumes du: "u \ carrier_vec d2" and dv: "v \ carrier_vec d3" and dA: "A \ carrier_mat d1 d2" and dB: "B \ carrier_mat d3 d4" shows "A * (outer_prod u v) * B = (outer_prod (A *\<^sub>v u) (adjoint B *\<^sub>v v))" unfolding outer_prod_def proof - have eq1: "A * (mat (dim_vec u) 1 (\(i, j). u $ i)) = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i)" apply (rule eq_matI) by (auto simp add: dA du scalar_prod_def) have conj: "conjugate a * b = conjugate ((a::complex) * conjugate b) " for a b by auto have eq2: "mat 1 (dim_vec v) (\(i, y). conjugate v $ y) * B = mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" apply (rule eq_matI) apply (auto simp add: carrier_matD[OF dB] carrier_vecD[OF dv] scalar_prod_def adjoint_def conjugate_vec_def sum_conjugate ) apply (rule sum.cong) by (auto simp add: conj) have "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *(mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B" using dA du dv dB assoc_mult_mat[OF dA, of "mat (dim_vec u) 1 (\(i, j). u $ i)" 1 "mat 1 (dim_vec v) (\(i, y). conjugate v $ y)"] by fastforce also have "\ = (A * (mat (dim_vec u) 1 (\(i, j). u $ i))) *((mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B)" using dA du dv dB assoc_mult_mat[OF _ _ dB, of "(A * (mat (dim_vec u) 1 (\(i, j). u $ i)))" d1 1] by fastforce finally show "A * (mat (dim_vec u) 1 (\(i, j). u $ i) * mat 1 (dim_vec v) (\(i, y). conjugate v $ y)) * B = mat (dim_vec (A *\<^sub>v u)) 1 (\(i, j). (A *\<^sub>v u) $ i) * mat 1 (dim_vec (adjoint B *\<^sub>v v)) (\(i, y). conjugate (adjoint B *\<^sub>v v) $ y)" using eq1 eq2 by auto qed subsection \Density operators\ definition density_operator :: "complex mat \ bool" where "density_operator A \ positive A \ trace A = 1" definition partial_density_operator :: "complex mat \ bool" where "partial_density_operator A \ positive A \ trace A \ 1" lemma pure_state_self_outer_prod_is_partial_density_operator: fixes v :: "complex vec" assumes dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" shows "partial_density_operator (outer_prod v v)" unfolding partial_density_operator_def proof have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto show "positive (outer_prod v v)" unfolding positive_def proof (rule, simp add: carrier_matD(2)[OF dimov] dimov, rule allI, rule impI) fix w assume "dim_vec (w::complex vec) = dim_col (outer_prod v v)" then have dimw: "w \ carrier_vec n" using dimov carrier_vecI by auto then have "inner_prod w ((outer_prod v v) *\<^sub>v w) = inner_prod w v * inner_prod v w" using inner_prod_outer_prod dimw dimv by auto also have "\ = inner_prod w v * conjugate (inner_prod w v)" using dimw dimv apply (subst conjugate_scalar_prod[of v "conjugate w"], simp) apply (subst conjugate_vec_sprod_comm[of "conjugate v" _ "conjugate w"], auto) apply (rule carrier_vec_conjugate[OF dimv]) apply (rule carrier_vec_conjugate[OF dimw]) done - also have "\ \ 0" by auto + also have "\ \ 0" by (auto simp: less_eq_complex_def) finally show "inner_prod w ((outer_prod v v) *\<^sub>v w) \ 0". qed have eq: "trace (outer_prod v v) = (\i=0..i=0..i=0.. 1" by auto qed (* Lemma 2.1 *) lemma lowner_le_trace: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" shows "A \\<^sub>L B \ (\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \))" proof (rule iffI) have dimBmA: "B - A \ carrier_mat n n" using A B by auto { assume "A \\<^sub>L B" then have pBmA: "positive (B - A)" using lowner_le_def by auto moreover have "B - A \ carrier_mat n n" using assms by auto ultimately have "\M\carrier_mat n n. M * adjoint M = B - A" using positive_iff_decomp[of "B - A"] by auto then obtain M where dimM: "M \ carrier_mat n n" and M: "M * adjoint M = B - A" by auto { fix \ assume dimr: "\ \ carrier_mat n n" and pdr: "partial_density_operator \" have eq: "trace(B * \) - trace(A * \) = trace((B - A) * \)" using A B dimr apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done have pr: "positive \" using pdr partial_density_operator_def by auto then have "\P\carrier_mat n n. \ = P * adjoint P" using positive_iff_decomp dimr by auto then obtain P where dimP: "P \ carrier_mat n n" and P: "\ = P * adjoint P" by auto have "trace((B - A) * \) = trace(M * adjoint M * (P * adjoint P))" using P M by auto also have "\ = trace((adjoint P * M) * adjoint (adjoint P * M))" using dimM dimP by (mat_assoc n) also have "\ \ 0" using trace_adjoint_positive by auto finally have "trace((B - A) * \) \ 0". with eq have " trace (B * \) - trace (A * \) \ 0" by auto } then show "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" by auto } { assume asm: "\\\carrier_mat n n. partial_density_operator \ \ trace (A * \) \ trace (B * \)" have "positive (B - A)" proof - { fix v assume "dim_vec (v::complex vec) = dim_col (B - A) \ vec_norm v = 1" then have dimv: "v \ carrier_vec n" and nv: "vec_norm v = 1" using carrier_matD[OF dimBmA] by (auto intro: carrier_vecI) have dimov: "outer_prod v v \ carrier_mat n n" using dimv by auto then have "partial_density_operator (outer_prod v v)" using dimv nv pure_state_self_outer_prod_is_partial_density_operator by auto then have leq: "trace(A * (outer_prod v v)) \ trace(B * (outer_prod v v))" using asm dimov by auto have "trace((B - A) * (outer_prod v v)) = trace(B * (outer_prod v v)) - trace(A * (outer_prod v v))" using A B dimov apply (subst minus_mult_distrib_mat, auto) apply (subst trace_minus_linear, auto) done then have "trace((B - A) * (outer_prod v v)) \ 0" using leq by auto then have "inner_prod v ((B - A) *\<^sub>v v) \ 0" using trace_outer_prod_right[OF dimBmA dimv dimv] by auto } then show "positive (B - A)" using positive_iff_normalized_vec[of "B - A"] dimBmA A by simp qed then show "A \\<^sub>L B" using lowner_le_def A B by auto } qed lemma lowner_le_traceI: assumes "A \ carrier_mat n n" and "B \ carrier_mat n n" and "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) \ trace (B * \)" shows "A \\<^sub>L B" using lowner_le_trace assms by auto lemma trace_pdo_eq_imp_eq: assumes A: "A \ carrier_mat n n" and B: "B \ carrier_mat n n" and teq: "\\. \ \ carrier_mat n n \ partial_density_operator \ \ trace (A * \) = trace (B * \)" shows "A = B" proof - from teq have "A \\<^sub>L B" using lowner_le_trace[OF A B] teq by auto moreover from teq have "B \\<^sub>L A" using lowner_le_trace[OF B A] teq by auto ultimately show "A = B" using lowner_le_antisym A B by auto qed lemma lowner_le_traceD: assumes "A \ carrier_mat n n" "B \ carrier_mat n n" "\ \ carrier_mat n n" and "A \\<^sub>L B" and "partial_density_operator \" shows "trace (A * \) \ trace (B * \)" using lowner_le_trace assms by blast lemma sum_only_one_neq_0: assumes "finite A" and "j \ A" and "\i. i \ A \ i \ j \ g i = 0" shows "sum g A = g j" proof - have "{j} \ A" using assms by auto moreover have "\i\A - {j}. g i = 0" using assms by simp ultimately have "sum g A = sum g {j}" using assms by (auto simp add: comm_monoid_add_class.sum.mono_neutral_right[of A "{j}" g]) moreover have "sum g {j} = g j" by simp ultimately show ?thesis by auto qed end diff --git a/thys/Stirling_Formula/Gamma_Asymptotics.thy b/thys/Stirling_Formula/Gamma_Asymptotics.thy --- a/thys/Stirling_Formula/Gamma_Asymptotics.thy +++ b/thys/Stirling_Formula/Gamma_Asymptotics.thy @@ -1,1893 +1,1893 @@ (* File: Gamma_Asymptotics.thy Author: Manuel Eberl The complete asymptotics of the real and complex logarithmic Gamma functions. Also of the real Polygamma functions (could be extended to the complex ones fairly easily if needed). *) section \Complete asymptotics of the logarithmic Gamma function\ theory Gamma_Asymptotics imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" Bernoulli.Bernoulli_FPS Bernoulli.Periodic_Bernpoly Stirling_Formula begin subsection \Auxiliary Facts\ (* TODO: could be automated with Laurent series expansions in the future *) lemma stirling_limit_aux1: "((\y. Ln (1 + z * of_real y) / of_real y) \ z) (at_right 0)" for z :: complex proof (cases "z = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + z * of_real y)) has_vector_derivative 1 * z) (at 0)" by (rule has_vector_derivative_real_field) (auto intro!: derivative_eq_intros) then have "(\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \0\ 0" by (auto simp add: has_vector_derivative_def has_derivative_def netlimit_at scaleR_conv_of_real field_simps) then have "((\y. (Ln (1 + z * of_real y) - of_real y * z) / of_real \y\) \ 0) (at_right 0)" by (rule filterlim_mono[OF _ _ at_le]) simp_all also have "?this \ ((\y. Ln (1 + z * of_real y) / (of_real y) - z) \ 0) (at_right 0)" using eventually_at_right_less[of "0::real"] by (intro filterlim_cong refl) (auto elim!: eventually_mono simp: field_simps) finally show ?thesis by (simp only: LIM_zero_iff) qed lemma stirling_limit_aux2: "((\y. y * Ln (1 + z / of_real y)) \ z) at_top" for z :: complex using stirling_limit_aux1[of z] by (subst filterlim_at_top_to_right) (simp add: field_simps) lemma Union_atLeastAtMost: assumes "N > 0" shows "(\n\{0.. {0..real N}" thus "x \ (\n\{0.. 0" "x < real N" by simp_all hence "x \ real (nat \x\)" "x \ real (nat \x\ + 1)" by linarith+ moreover from x have "nat \x\ < N" by linarith ultimately have "\n\{0.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto subsection \Cones in the complex plane\ definition complex_cone :: "real \ real \ complex set" where "complex_cone a b = {z. \y\{a..b}. z = rcis (norm z) y}" abbreviation complex_cone' :: "real \ complex set" where "complex_cone' a \ complex_cone (-a) a" lemma zero_in_complex_cone [simp, intro]: "a \ b \ 0 \ complex_cone a b" by (auto simp: complex_cone_def) lemma complex_coneE: assumes "z \ complex_cone a b" obtains r \ where "r \ 0" "\ \ {a..b}" "z = rcis r \" proof - from assms obtain y where "y \ {a..b}" "z = rcis (norm z) y" unfolding complex_cone_def by auto thus ?thesis using that[of "norm z" y] by auto qed lemma arg_cis [simp]: assumes "x \ {-pi<..pi}" shows "Arg (cis x) = x" using assms by (intro cis_Arg_unique) auto lemma arg_mult_of_real_left [simp]: assumes "r > 0" shows "Arg (of_real r * z) = Arg z" proof (cases "z = 0") case False thus ?thesis using Arg_bounded[of z] assms by (intro cis_Arg_unique) (auto simp: sgn_mult sgn_of_real cis_Arg) qed auto lemma arg_mult_of_real_right [simp]: assumes "r > 0" shows "Arg (z * of_real r) = Arg z" by (subst mult.commute, subst arg_mult_of_real_left) (simp_all add: assms) lemma arg_rcis [simp]: assumes "x \ {-pi<..pi}" "r > 0" shows "Arg (rcis r x) = x" using assms by (simp add: rcis_def) lemma rcis_in_complex_cone [intro]: assumes "\ \ {a..b}" "r \ 0" shows "rcis r \ \ complex_cone a b" using assms by (auto simp: complex_cone_def) lemma arg_imp_in_complex_cone: assumes "Arg z \ {a..b}" shows "z \ complex_cone a b" proof - have "z = rcis (norm z) (Arg z)" by (simp add: rcis_cmod_Arg) also have "\ \ complex_cone a b" using assms by auto finally show ?thesis . qed lemma complex_cone_altdef: assumes "-pi < a" "a \ b" "b \ pi" shows "complex_cone a b = insert 0 {z. Arg z \ {a..b}}" proof (intro equalityI subsetI) fix z assume "z \ complex_cone a b" then obtain r \ where *: "r \ 0" "\ \ {a..b}" "z = rcis r \" by (auto elim: complex_coneE) have "Arg z \ {a..b}" if [simp]: "z \ 0" proof - have "r > 0" using that * by (subst (asm) *) auto hence "\ \ {a..b}" using *(1,2) assms by (auto simp: *(1)) moreover from assms *(2) have "\ \ {-pi<..pi}" by auto ultimately show ?thesis using *(3) \r > 0\ by (subst *) auto qed thus "z \ insert 0 {z. Arg z \ {a..b}}" by auto qed (use assms in \auto intro: arg_imp_in_complex_cone\) lemma nonneg_of_real_in_complex_cone [simp, intro]: assumes "x \ 0" "a \ 0" "0 \ b" shows "of_real x \ complex_cone a b" proof - from assms have "rcis x 0 \ complex_cone a b" by (intro rcis_in_complex_cone) auto thus ?thesis by simp qed lemma one_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ 1 \ complex_cone a b" using nonneg_of_real_in_complex_cone[of 1] by (simp del: nonneg_of_real_in_complex_cone) lemma of_nat_in_complex_cone [simp, intro]: "a \ 0 \ 0 \ b \ of_nat n \ complex_cone a b" using nonneg_of_real_in_complex_cone[of "real n"] by (simp del: nonneg_of_real_in_complex_cone) subsection \Another integral representation of the Beta function\ lemma complex_cone_inter_nonpos_Reals: assumes "-pi < a" "a \ b" "b < pi" shows "complex_cone a b \ \\<^sub>\\<^sub>0 = {0}" proof (safe elim!: nonpos_Reals_cases) fix x :: real assume "complex_of_real x \ complex_cone a b" "x \ 0" hence "\(x < 0)" using assms by (intro notI) (auto simp: complex_cone_altdef) with \x \ 0\ show "complex_of_real x = 0" by auto qed (use assms in auto) theorem assumes a: "a > 0" and b: "b > (0 :: real)" shows has_integral_Beta_real': "((\u. u powr (b - 1) / (1 + u) powr (a + b)) has_integral Beta a b) {0<..}" and Beta_conv_nn_integral: "Beta a b = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" proof - define I where "I = (\\<^sup>+u. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) \lborel)" have "Gamma (a + b) > 0" "Beta a b > 0" using assms by (simp_all add: add_pos_pos Beta_def) from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add mult_ac) also have "\ = (\\<^sup>+t. indicator {0<..} t * (\\<^sup>+u. indicator {0..} u * t powr (a - 1) * u powr (b - 1) / exp (t + u) \(density (distr lborel borel ((*) t)) (\x. ennreal \t\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(t::real). indicator {0<..} t * (\\<^sup>+u. indicator {0..} (u * t) * t powr a * (u * t) powr (b - 1) / exp (t + t * u) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_density nn_integral_distr algebra_simps powr_diff simp flip: ennreal_mult) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr a * (u * t) powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def zero_le_mult_iff algebra_simps) also have "\ = (\\<^sup>+(t::real). \\<^sup>+u. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: powr_add powr_diff indicator_def powr_mult field_simps) also have "\ = (\\<^sup>+(u::real). \\<^sup>+t. indicator ({0<..}\{0..}) (t, u) * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel \lborel)" by (rule lborel_pair.Fubini') auto also have "\ = (\\<^sup>+(u::real). indicator {0..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \lborel) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0}"]) (auto simp: indicator_def) also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) * u powr (b - 1) / exp (t * (u + 1)) \(density (distr lborel borel ((*) (1/(1+u)))) (\x. ennreal \1/(1+u)\))) \lborel)" by (intro nn_integral_cong mult_indicator_cong, subst lborel_distr_mult' [symmetric]) auto also have "\ = (\\<^sup>+(u::real). indicator {0<..} u * (\\<^sup>+t. ennreal (1 / (u + 1)) * ennreal (indicator {0<..} (t / (u + 1)) * (t / (1+u)) powr (a + b - 1) * u powr (b - 1) / exp t) \lborel) \lborel)" by (intro nn_integral_cong mult_indicator_cong) (auto simp: nn_integral_distr nn_integral_density add_ac) also have "\ = (\\<^sup>+u. \\<^sup>+t. indicator ({0<..}\{0<..}) (u, t) * 1/(u+1) * (t / (u+1)) powr (a + b - 1) * u powr (b - 1) / exp t \lborel \lborel)" by (subst nn_integral_cmult [symmetric], simp, intro nn_integral_cong) (auto simp: indicator_def field_simps divide_ennreal simp flip: ennreal_mult ennreal_mult') also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0<..} u * u powr (b - 1) / (1 + u) powr (a + b)) * ennreal (indicator {0<..} t * t powr (a + b - 1) / exp t) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_diff powr_divide powr_minus divide_simps add_ac simp flip: ennreal_mult) also have "\ = I * (\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel)" by (simp add: nn_integral_cmult nn_integral_multc I_def) also have "(\\<^sup>+t. indicator {0<..} t * t powr (a + b - 1) / exp t \lborel) = ennreal (Gamma (a + b))" using assms by (subst Gamma_conv_nn_integral_real) (auto intro!: nn_integral_cong_AE[OF AE_I[of _ _ "{0}"]] simp: indicator_def split: if_splits) finally have "ennreal (Gamma a * Gamma b) = I * ennreal (Gamma (a + b))" . hence "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = I * ennreal (Gamma (a + b)) / ennreal (Gamma (a + b))" by simp also have "\ = I" using \Gamma (a + b) > 0\ by (intro ennreal_mult_divide_eq) (auto simp: ) also have "ennreal (Gamma a * Gamma b) / ennreal (Gamma (a + b)) = ennreal (Gamma a * Gamma b / Gamma (a + b))" using assms by (intro divide_ennreal) auto also have "\ = ennreal (Beta a b)" by (simp add: Beta_def) finally show *: "ennreal (Beta a b) = I" . define f where "f = (\u. u powr (b - 1) / (1 + u) powr (a + b))" have "((\u. indicator {0<..} u * f u) has_integral Beta a b) UNIV" using * \Beta a b > 0\ by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: f_def measurable_completion nn_integral_completion I_def mult_ac) also have "(\u. indicator {0<..} u * f u) = (\u. if u \ {0<..} then f u else 0)" by (auto simp: fun_eq_iff) also have "(\ has_integral Beta a b) UNIV \ (f has_integral Beta a b) {0<..}" by (rule has_integral_restrict_UNIV) finally show \ by (simp add: f_def) qed lemma has_integral_Beta2: fixes a :: real assumes "a < -1/2" shows "((\x. (1 + x ^ 2) powr a) has_integral Beta (- a - 1 / 2) (1 / 2) / 2) {0<..}" proof - define f where "f = (\u. u powr (-1/2) / (1 + u) powr (-a))" define C where "C = Beta (-a-1/2) (1/2)" have I: "(f has_integral C) {0<..}" using has_integral_Beta_real'[of "-a-1/2" "1/2"] assms by (simp_all add: diff_divide_distrib f_def C_def) define g where "g = (\x. x ^ 2 :: real)" have bij: "bij_betw g {0<..} {0<..}" by (intro bij_betwI[of _ _ _ sqrt]) (auto simp: g_def) have "(f absolutely_integrable_on g ` {0<..} \ integral (g ` {0<..}) f = C)" using I bij by (simp add: bij_betw_def has_integral_iff absolutely_integrable_on_def f_def) also have "?this \ ((\x. \2 * x\ *\<^sub>R f (g x)) absolutely_integrable_on {0<..} \ integral {0<..} (\x. \2 * x\ *\<^sub>R f (g x)) = C)" using bij by (intro has_absolute_integral_change_of_variables_1' [symmetric]) (auto intro!: derivative_eq_intros simp: g_def bij_betw_def) finally have "((\x. \2 * x\ * f (g x)) has_integral C) {0<..}" by (simp add: absolutely_integrable_on_def f_def has_integral_iff) also have "?this \ ((\x::real. 2 * (1 + x\<^sup>2) powr a) has_integral C) {0<..}" by (intro has_integral_cong) (auto simp: f_def g_def powr_def exp_minus ln_realpow field_simps) finally have "((\x::real. 1/2 * (2 * (1 + x\<^sup>2) powr a)) has_integral 1/2 * C) {0<..}" by (intro has_integral_mult_right) thus ?thesis by (simp add: C_def) qed lemma has_integral_Beta3: fixes a b :: real assumes "a < -1/2" and "b > 0" shows "((\x. (b + x ^ 2) powr a) has_integral Beta (-a - 1/2) (1/2) / 2 * b powr (a + 1/2)) {0<..}" proof - define C where "C = Beta (- a - 1 / 2) (1 / 2) / 2" have int: "nn_integral lborel (\x. indicator {0<..} x * (1 + x ^ 2) powr a) = C" using nn_integral_has_integral_lebesgue[OF _ has_integral_Beta2[OF assms(1)]] by (auto simp: C_def) have "nn_integral lborel (\x. indicator {0<..} x * (b + x ^ 2) powr a) = (\\<^sup>+x. ennreal (indicat_real {0<..} (x * sqrt b) * (b + (x * sqrt b)\<^sup>2) powr a * sqrt b) \lborel)" using assms by (subst lborel_distr_mult'[of "sqrt b"]) (auto simp: nn_integral_density nn_integral_distr mult_ac simp flip: ennreal_mult) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * (b * (1 + x ^ 2)) powr a * sqrt b) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def field_simps zero_less_mult_iff) also have "\ = (\\<^sup>+x. ennreal (indicat_real {0<..} x * b powr (a + 1/2) * (1 + x ^ 2) powr a) \lborel)" using assms by (intro nn_integral_cong) (auto simp: indicator_def powr_add powr_half_sqrt powr_mult) also have "\ = b powr (a + 1/2) * (\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel)" using assms by (subst nn_integral_cmult [symmetric]) (simp_all add: mult_ac flip: ennreal_mult) also have "(\\<^sup>+x. ennreal (indicat_real {0<..} x * (1 + x ^ 2) powr a) \lborel) = C" using int by simp also have "ennreal (b powr (a + 1/2)) * ennreal C = ennreal (C * b powr (a + 1/2))" using assms by (subst ennreal_mult) (auto simp: C_def mult_ac Beta_def) finally have *: "(\\<^sup>+ x. ennreal (indicat_real {0<..} x * (b + x\<^sup>2) powr a) \lborel) = \" . hence "((\x. indicator {0<..} x * (b + x^2) powr a) has_integral C * b powr (a + 1/2)) UNIV" using assms by (subst has_integral_iff_nn_integral_lebesgue) (auto simp: C_def measurable_completion nn_integral_completion Beta_def) also have "(\x. indicator {0<..} x * (b + x^2) powr a) = (\x. if x \ {0<..} then (b + x^2) powr a else 0)" by (auto simp: fun_eq_iff) finally show ?thesis by (subst (asm) has_integral_restrict_UNIV) (auto simp: C_def) qed subsection \Asymptotics of the real $\log\Gamma$ function and its derivatives\ text \ This is the error term that occurs in the expansion of @{term ln_Gamma}. It can be shown to be of order $O(s^{-n})$. \ definition stirling_integral :: "nat \ 'a :: {real_normed_div_algebra, banach} \ 'a" where "stirling_integral n s = lim (\N. integral {0..N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" context fixes s :: complex assumes s: "s \ \\<^sub>\\<^sub>0" fixes approx :: "nat \ complex" defines "approx \ (\N. (\n = 1.. \\\ ln_Gamma s\\ (ln_Gamma (of_nat N) - ln (2 * pi / of_nat N) / 2 - of_nat N * ln (of_nat N) + of_nat N) - \ \\\ 0\\ s * (harm (N - 1) - ln (of_nat (N - 1)) - euler_mascheroni) + \ \\\ 0\\ s * (ln (of_nat N + s) - ln (of_nat (N - 1))) - \ \\\ 0\\ (1/2) * (ln (of_nat N + s) - ln (of_nat N)) + \ \\\ 0\\ of_nat N * (ln (of_nat N + s) - ln (of_nat N)) - \ \\\ s\\ (s - 1/2) * ln s - ln (2 * pi) / 2)" begin qualified lemma assumes N: "N > 0" shows integrable_pbernpoly_1: "(\x. of_real (-pbernpoly 1 x) / (of_real x + s)) integrable_on {0..real N}" and integral_pbernpoly_1_aux: "integral {0..real N} (\x. -of_real (pbernpoly 1 x) / (of_real x + s)) = approx N" and has_integral_pbernpoly_1: "((\x. pbernpoly 1 x /(x + s)) has_integral (\mx. -pbernpoly 1 x / (x + s)) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" for n proof (rule has_integral_spike) - have "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral + have "((\x. (of_nat n + 1/2 + s) * (1 / (of_real x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_real (real (n + 1)) + s) - ln (of_real (real n) + s)) - 1) {of_nat n..of_nat (n + 1)}" using s has_integral_const_real[of 1 "of_nat n" "of_nat (n + 1)"] by (intro has_integral_diff has_integral_mult_right fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: has_field_derivative_iff_has_vector_derivative [symmetric] field_simps complex_nonpos_Reals_iff) - thus "((\x. (of_nat n + 1/2 + s) * (1 / (x + s)) - 1) has_integral + thus "((\x. (of_nat n + 1/2 + s) * (1 / (of_real x + s)) - 1) has_integral (of_nat n + 1/2 + s) * (ln (of_nat (n + 1) + s) - ln (of_nat n + s)) - 1) {of_nat n..of_nat (n + 1)}" by simp show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" if "x \ {of_nat n..of_nat (n + 1)} - {of_nat (n + 1)}" for x proof - have x: "x \ real n" "x < real (n + 1)" using that by simp_all hence "floor x = int n" by linarith moreover from s x have "complex_of_real x \ -s" by (auto simp add: complex_eq_iff complex_nonpos_Reals_iff simp del: of_nat_Suc) ultimately show "-pbernpoly 1 x / (x + s) = (of_nat n + 1/2 + s) * (1 / (x + s)) - 1" by (auto simp: pbernpoly_def bernpoly_def frac_def divide_simps add_eq_0_iff2) qed qed simp_all hence *: "\I. I\?A \ ((\x. -pbernpoly 1 x / (x + s)) has_integral (Inf I + 1/2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1) I" by (auto simp: add_ac) have "((\x. - pbernpoly 1 x / (x + s)) has_integral (\I\?A. (Inf I + 1 / 2 + s) * (ln (Inf I + 1 + s) - ln (Inf I + s)) - 1)) (\n\{0..x. - pbernpoly 1 x / (x + s)) has_integral ?i) {0..real N}" by (subst has_integral_spike_set_eq) (use Union_atLeastAtMost assms in \auto simp: intro!: empty_imp_negligible\) hence "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" and integral: "integral {0..real N} (\x. - pbernpoly 1 x / (x + s)) = ?i" by (simp_all add: has_integral_iff) show "(\x. - pbernpoly 1 x / (x + s)) integrable_on {0..real N}" by fact note has_integral_neg[OF has_integral] also have "-?i = (\xx. of_real (pbernpoly 1 x) / (of_real x + s)) has_integral \) {0..real N}" by simp note integral also have "?i = (\nnnn=1.. = (\n=1..n=1..n=1..n=1.. - (\n = 1..n=1.. 0" for x by (auto simp: complex_nonpos_Reals_iff complex_eq_iff) hence "(\n=1..n=1.. = ln (fact (N - 1)) + (\n=1..n=1..n=1..n=1..n = 1..n = 1..x. -of_real (pbernpoly 1 x) / (of_real x + s)) = \" by simp qed lemma integrable_ln_Gamma_aux: shows "(\x. of_real (pbernpoly n x) / (of_real x + s) ^ n) integrable_on {0..real N}" proof (cases "n = 1") case True with s show ?thesis using integrable_neg[OF integrable_pbernpoly_1[of N]] by (cases "N = 0") (simp_all add: integrable_negligible) next case False from s have "of_real x + s \ 0" if "x \ 0" for x using that by (auto simp: complex_eq_iff add_eq_0_iff2 complex_nonpos_Reals_iff) with False s show ?thesis by (auto intro!: integrable_continuous_real continuous_intros) qed text \ This following proof is based on ``Rudiments of the theory of the gamma function'' by Bruce Berndt~\cite{berndt}. \ lemma tendsto_of_real_0_I: "(f \ 0) G \ ((\x. (of_real (f x))) \ (0 ::'a::real_normed_div_algebra)) G" using tendsto_of_real_iff by force qualified lemma integral_pbernpoly_1: "(\N. integral {0..real N} (\x. pbernpoly 1 x / (x + s))) \ -ln_Gamma s - s + (s - 1 / 2) * ln s + ln (2 * pi) / 2" proof - have neq: "s + of_real x \ 0" if "x \ 0" for x :: real using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) have "(approx \ ln_Gamma s - 0 - 0 + 0 - 0 + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" unfolding approx_def proof (intro tendsto_add tendsto_diff) from s have s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff elim!: nonpos_Ints_cases) have "(\n. \i=1.. ln_Gamma s + euler_mascheroni * s + ln s" (is "?f \ _") using ln_Gamma_series'_aux[OF s'] unfolding sums_def by (subst filterlim_sequentially_Suc [symmetric], subst (asm) sum.atLeast1_atMost_eq [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) thus "((\n. ?f n - (euler_mascheroni * s + ln s)) \ ln_Gamma s) at_top" by (auto intro: tendsto_eq_intros) next show "(\x. complex_of_real (ln_Gamma (real x) - ln (2 * pi / real x) / 2 - real x * ln (real x) + real x)) \ 0" proof (intro tendsto_of_real_0_I filterlim_compose[OF tendsto_sandwich filterlim_real_sequentially]) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 0) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(1), simp add: algebra_simps) show "eventually (\x::real. ln_Gamma x - ln (2 * pi / x) / 2 - x * ln x + x \ 1 / 12 * inverse x) at_top" using eventually_ge_at_top[of "1::real"] by eventually_elim (insert ln_Gamma_bounds(2), simp add: field_simps) show "((\x::real. 1 / 12 * inverse x) \ 0) at_top" by (intro tendsto_mult_right_zero tendsto_inverse_0_at_top filterlim_ident) qed simp_all next have "(\x. s * of_real (harm (x - 1) - ln (real (x - 1)) - euler_mascheroni)) \ s * of_real (euler_mascheroni - euler_mascheroni)" by (subst filterlim_sequentially_Suc [symmetric], intro tendsto_intros) (insert euler_mascheroni_LIMSEQ, simp_all) also have "?this \ (\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" by (intro filterlim_cong refl eventually_mono[OF eventually_gt_at_top[of "1::nat"]]) (auto simp: Ln_of_nat of_real_harm) finally show "(\x. s * (harm (x - 1) - ln (of_nat (x - 1)) - euler_mascheroni)) \ 0" . next have "((\x. ln (1 + (s + 1) / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real (x + 1) + s) - ln (complex_of_real x) = ln (1 + (s + 1) / of_real x)" if "x > 1" for x using that s using Ln_divide_of_real[of x "of_real (x + 1) + s", symmetric] neq[of "x+1"] by (simp add: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real (x + 1) + s) - ln (of_real x)) \ 0) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "((\n. ln (of_real (real n + 1) + s) - ln (of_real (real n))) \ 0) at_top" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) hence "((\n. ln (of_nat n + s) - ln (of_nat (n - 1))) \ 0) at_top" by (subst filterlim_sequentially_Suc [symmetric]) (simp add: add_ac) thus "(\x. s * (ln (of_nat x + s) - ln (of_nat (x - 1)))) \ 0" by (rule tendsto_mult_right_zero) next have "((\x. ln (1 + s / of_real x)) \ ln (1 + 0)) at_top" (is ?P) by (intro tendsto_intros tendsto_divide_0[OF tendsto_const]) (simp_all add: filterlim_ident filterlim_at_infinity_conv_norm_at_top filterlim_abs_real) also have "ln (of_real x + s) - ln (of_real x) = ln (1 + s / of_real x)" if "x > 0" for x using Ln_divide_of_real[of x "of_real x + s"] neq[of x] that by (auto simp: field_simps Ln_of_real) hence "?P \ ((\x. ln (of_real x + s) - ln (of_real x)) \ 0) at_top" using s by (intro filterlim_cong refl) (auto intro: eventually_mono [OF eventually_gt_at_top[of "1::real"]]) finally have "(\x. (1/2) * (ln (of_real (real x) + s) - ln (of_real (real x)))) \ 0" by (rule tendsto_mult_right_zero[OF filterlim_compose[OF _ filterlim_real_sequentially]]) thus "(\x. (1/2) * (ln (of_nat x + s) - ln (of_nat x))) \ 0" by simp next have "((\x. x * (ln (1 + s / of_real x))) \ s) at_top" (is ?P) by (rule stirling_limit_aux2) also have "ln (1 + s / of_real x) = ln (of_real x + s) - ln (of_real x)" if "x > 1" for x using that s Ln_divide_of_real [of x "of_real x + s", symmetric] neq[of x] by (auto simp: Ln_of_real field_simps) hence "?P \ ((\x. of_real x * (ln (of_real x + s) - ln (of_real x))) \ s) at_top" by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "1::real"]]) finally have "(\n. of_real (real n) * (ln (of_real (real n) + s) - ln (of_real (real n)))) \ s" by (rule filterlim_compose[OF _ filterlim_real_sequentially]) thus "(\n. of_nat n * (ln (of_nat n + s) - ln (of_nat n))) \ s" by simp qed simp_all also have "?this \ ((\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) \ ln_Gamma s + s - (s - 1/2) * ln s - ln (2 * pi) / 2) at_top" using integral_pbernpoly_1_aux by (intro filterlim_cong refl) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::nat"]]) also have "(\N. integral {0..real N} (\x. -pbernpoly 1 x / (x + s))) = (\N. -integral {0..real N} (\x. pbernpoly 1 x / (x + s)))" by (simp add: fun_eq_iff) finally show ?thesis by (simp add: tendsto_minus_cancel_left [symmetric] algebra_simps) qed qualified lemma pbernpoly_integral_conv_pbernpoly_integral_Suc: assumes "n \ 1" shows "integral {0..real N} (\x. pbernpoly n x / (x + s) ^ n) = of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * integral {0..real N} (\x. of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)" proof - note [derivative_intros] = has_field_derivative_pbernpoly_Suc' define I where "I = -of_real (pbernpoly (Suc n) (of_nat N)) / (of_nat (Suc n) * (of_nat N + s) ^ n) + of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n + integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)" have "((\x. (-of_nat n * inverse (of_real x + s) ^ Suc n) * (of_real (pbernpoly (Suc n) x) / (of_nat (Suc n)))) has_integral -I) {0..real N}" proof (rule integration_by_parts_interior_strong[OF bounded_bilinear_mult]) fix x :: real assume x: "x \ {0<.. \" proof assume "x \ \" then obtain n where "x = of_int n" by (auto elim!: Ints_cases) with x have x': "x = of_nat (nat n)" by simp from x show False by (auto simp: x') qed hence "((\x. of_real (pbernpoly (Suc n) x / of_nat (Suc n))) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by (intro has_vector_derivative_of_real) (auto intro!: derivative_eq_intros) thus "((\x. of_real (pbernpoly (Suc n) x) / of_nat (Suc n)) has_vector_derivative complex_of_real (pbernpoly n x)) (at x)" by simp from x s have "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "((\x. inverse (of_real x + s) ^ n) has_vector_derivative - of_nat n * inverse (of_real x + s) ^ Suc n) (at x)" using x s assms by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: divide_simps power_add [symmetric] simp del: power_Suc) next have "complex_of_real x + s \ 0" if "x \ 0" for x using that s by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) thus "continuous_on {0..real N} (\x. inverse (of_real x + s) ^ n)" "continuous_on {0..real N} (\x. complex_of_real (pbernpoly (Suc n) x) / of_nat (Suc n))" using assms s by (auto intro!: continuous_intros simp del: of_nat_Suc) next have "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral pbernpoly (Suc n) (of_nat N) / (of_nat (Suc n) * (of_nat N + s) ^ n) - of_real (bernoulli (Suc n) / real (Suc n)) / s ^ n - -I) {0..real N}" using integrable_ln_Gamma_aux[of n N] assms by (auto simp: I_def has_integral_integral divide_simps) thus "((\x. inverse (of_real x + s) ^ n * of_real (pbernpoly n x)) has_integral inverse (of_real (real N) + s) ^ n * (of_real (pbernpoly (Suc n) (real N)) / of_nat (Suc n)) - inverse (of_real 0 + s) ^ n * (of_real (pbernpoly (Suc n) 0) / of_nat (Suc n)) - - I) {0..real N}" by (simp_all add: field_simps) qed simp_all also have "(\x. - of_nat n * inverse (of_real x + s) ^ Suc n * (of_real (pbernpoly (Suc n) x) / of_nat (Suc n))) = (\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))" by (simp add: divide_simps fun_eq_iff) finally have "((\x. - (of_nat n / of_nat (Suc n) * of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n)) has_integral - I) {0..real N}" . from has_integral_neg[OF this] show ?thesis by (auto simp add: I_def has_integral_iff algebra_simps integral_mult_right [symmetric] simp del: power_Suc of_nat_Suc ) qed lemma pbernpoly_over_power_tendsto_0: assumes "n > 0" shows "(\x. of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ 0" proof - from s have neq: "s + of_nat n \ 0" for n by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) obtain c where c: "\x. norm (pbernpoly (Suc n) x) \ c" using bounded_pbernpoly by auto have "eventually (\x. real x + Re s > 0) at_top" by real_asymp hence "eventually (\x. norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / (real x + Re s) ^ n) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim x) have "norm (of_real (pbernpoly (Suc n) (real x)) / (of_nat (Suc n) * (s + of_nat x) ^ n)) \ (c / real (Suc n)) / norm (s + of_nat x) ^ n" (is "_ \ ?rhs") using c[of x] by (auto simp: norm_divide norm_mult norm_power neq field_simps simp del: of_nat_Suc) also have "(real x + Re s) \ cmod (s + of_nat x)" using complex_Re_le_cmod[of "s + of_nat x"] s by (auto simp add: complex_nonpos_Reals_iff) hence "?rhs \ (c / real (Suc n)) / (real x + Re s) ^ n" using s elim c[of 0] neq[of x] by (intro divide_left_mono power_mono mult_pos_pos divide_nonneg_pos zero_less_power) auto finally show ?case . qed moreover have "(\x. (c / real (Suc n)) / (real x + Re s) ^ n) \ 0" using \n > 0\ by real_asymp ultimately show ?thesis by (rule Lim_null_comparison) qed lemma convergent_stirling_integral: assumes "n > 0" shows "convergent (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n))" (is "convergent (?f n)") proof - have "convergent (?f (Suc n))" for n proof (induction n) case 0 thus ?case using integral_pbernpoly_1 by (auto intro!: convergentI) next case (Suc n) have "convergent (\N. ?f (Suc n) N - of_real (pbernpoly (Suc (Suc n)) (real N)) / (of_nat (Suc (Suc n)) * (s + of_nat N) ^ Suc n) + of_real (bernoulli (Suc (Suc n)) / (real (Suc (Suc n)))) / s ^ Suc n)" (is "convergent ?g") by (intro convergent_add convergent_diff Suc convergent_const convergentI[OF pbernpoly_over_power_tendsto_0]) simp_all also have "?g = (\N. of_nat (Suc n) / of_nat (Suc (Suc n)) * ?f (Suc (Suc n)) N)" using s by (subst pbernpoly_integral_conv_pbernpoly_integral_Suc) (auto simp: fun_eq_iff field_simps simp del: of_nat_Suc power_Suc) also have "convergent \ \ convergent (?f (Suc (Suc n)))" by (intro convergent_mult_const_iff) (simp_all del: of_nat_Suc) finally show ?case . qed from this[of "n - 1"] assms show ?thesis by simp qed lemma stirling_integral_conv_stirling_integral_Suc: assumes "n > 0" shows "stirling_integral n s = of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" proof - have "(\N. of_real (pbernpoly (Suc n) (real N)) / (of_nat (Suc n) * (s + of_nat N) ^ n) - of_real (bernoulli (Suc n)) / (real (Suc n) * s ^ n) + integral {0..real N} (\x. of_nat n / of_nat (Suc n) * (of_real (pbernpoly (Suc n) x) / (of_real x + s) ^ Suc n))) \ 0 - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n) + of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s" (is "?f \ _") unfolding stirling_integral_def integral_mult_right using convergent_stirling_integral[of "Suc n"] assms s by (intro tendsto_intros pbernpoly_over_power_tendsto_0) (auto simp: convergent_LIMSEQ_iff simp del: of_nat_Suc) also have "?this \ (\N. integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" using eventually_gt_at_top[of "0::nat"] pbernpoly_integral_conv_pbernpoly_integral_Suc[of n] assms unfolding integral_mult_right by (intro filterlim_cong refl) (auto elim!: eventually_mono simp del: power_Suc) finally show ?thesis unfolding stirling_integral_def[of n] by (rule limI) qed lemma stirling_integral_1_unfold: assumes "m > 0" shows "stirling_integral 1 s = stirling_integral m s / of_nat m - (\k=1..k=1..k = 1.. 0" shows "ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..k = 1.. 0 \ (\x. integral {0..real x} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ stirling_integral n s" unfolding stirling_integral_def using convergent_stirling_integral[of n] by (simp only: convergent_LIMSEQ_iff) end lemmas has_integral_of_real = has_integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemmas integral_of_real = integral_linear[OF _ bounded_linear_of_real, unfolded o_def] lemma integrable_ln_Gamma_aux_real: assumes "0 < s" shows "(\x. pbernpoly n x / (x + s) ^ n) integrable_on {0..real N}" proof - have "(\x. complex_of_real (pbernpoly n x / (x + s) ^ n)) integrable_on {0..real N}" using integrable_ln_Gamma_aux[of "of_real s" n N] assms by simp from integrable_linear[OF this bounded_linear_Re] show ?thesis by (simp only: o_def Re_complex_of_real) qed lemma assumes "x > 0" "n > 0" shows stirling_integral_complex_of_real: "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" and LIMSEQ_stirling_integral_real: "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" and stirling_integral_real_convergent: "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" proof - have "(\N. integral {0..real N} (\t. of_real (pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using LIMSEQ_stirling_integral[of "complex_of_real x" n] assms by simp hence "(\N. of_real (integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))) \ stirling_integral n (complex_of_real x)" using integrable_ln_Gamma_aux_real[OF assms(1), of n] by (subst (asm) integral_of_real) simp from tendsto_Re[OF this] have "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ Re (stirling_integral n (complex_of_real x))" by simp thus "convergent (\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n))" by (rule convergentI) thus "(\N. integral {0..real N} (\t. pbernpoly n t / (t + x) ^ n)) \ stirling_integral n x" unfolding stirling_integral_def by (simp add: convergent_LIMSEQ_iff) from tendsto_of_real[OF this, where 'a = complex] integrable_ln_Gamma_aux_real[OF assms(1), of n] have "(\xa. integral {0..real xa} (\xa. complex_of_real (pbernpoly n xa) / (complex_of_real xa + x) ^ n)) \ complex_of_real (stirling_integral n x)" by (subst (asm) integral_of_real [symmetric]) simp_all from LIMSEQ_unique[OF this LIMSEQ_stirling_integral[of "complex_of_real x" n]] assms show "stirling_integral n (complex_of_real x) = of_real (stirling_integral n x)" by simp qed lemma ln_Gamma_stirling_real: assumes "x > (0 :: real)" "m > (0::nat)" shows "ln_Gamma x = (x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1..k = 1.. (1::nat)" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp have "norm (stirling_integral n s) \ c / (real n - 1) / Re s ^ (n - 1)" if s: "Re s > 0" for s proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) have pos: "x + norm s > 0" if "x \ 0" for x using s that by (intro add_nonneg_pos) auto have nz: "of_real x + s \ 0" if "x \ 0" for x using s that by (auto simp: complex_eq_iff) let ?bound = "\N. c / (Re s ^ (n - 1) * (real n - 1)) - c / ((real N + Re s) ^ (n - 1) * (real n - 1))" show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) let ?F = "\x. -c / ((x + Re s) ^ (n - 1) * (real n - 1))" from n s have "((\x. c / (x + Re s) ^ n) has_integral (?F (real N) - ?F 0)) {0..real N}" by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 has_field_derivative_iff_has_vector_derivative [symmetric]) also have "?F (real N) - ?F 0 = ?bound N" by simp finally have *: "((\x. c / (x + Re s) ^ n) has_integral ?bound N) {0..real N}" . have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / (x + Re s) ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) fix x assume x: "x \ {0..real N}" have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / (x + Re s) ^ n" using x c c_nonneg s nz[of x] complex_Re_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power add_nonneg_pos) auto finally show "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ \" . qed (insert n s * pos nz c, auto simp: complex_nonpos_Reals_iff) also have "\ = ?bound N" using * by (simp add: has_integral_iff) also have "\ \ c / (Re s ^ (n - 1) * (real n - 1))" using c_nonneg elim s n by simp also have "\ = c / (real n - 1) / (Re s ^ (n - 1))" by simp finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ c / (real n - 1) / Re s ^ (n - 1)" . qed qed (insert s n, simp_all add: complex_nonpos_Reals_iff) thus ?thesis by (rule that) qed lemma stirling_integral_bound_aux_integral1: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" "l < a - b" "r > a + b" shows "((\x. c / max b \x - a\ ^ n) has_integral 2*c*(n / (n - 1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))) {l..r}" proof - define x1 x2 where "x1 = a - b" and "x2 = a + b" define F1 where "F1 = (\x::real. c / (a - x) ^ (n - 1) / (n - 1))" define F3 where "F3 = (\x::real. -c / (x - a) ^ (n - 1) / (n - 1))" have deriv: "(F1 has_vector_derivative (c / (a - x) ^ n)) (at x within A)" "(F3 has_vector_derivative (c / (x - a) ^ n)) (at x within A)" if "x \ a" for x :: real and A unfolding F1_def F3_def using assms that by (auto intro!: derivative_eq_intros simp: divide_simps power_diff add_eq_0_iff2 simp flip: has_field_derivative_iff_has_vector_derivative) from assms have "((\x. c / (a - x) ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x1_def max_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" using assms by (intro has_integral_spike_finite_eq[of "{l}"]) (auto simp: x1_def max_def split: if_splits) finally have I1: "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l)) {l..x1}" . have "((\x. c / b ^ n) has_integral (x2 - x1) * c / b ^ n) {x1..x2}" using has_integral_const_real[of "c / b ^ n" x1 x2] assms by (simp add: x1_def x2_def) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" by (intro has_integral_spike_finite_eq[of "{x1, x2}"]) (auto simp: x1_def x2_def split: if_splits) finally have I2: "((\x. c / max b \x - a\ ^ n) has_integral ((x2 - x1) * c / b ^ n)) {x1..x2}" . from assms have I3: "((\x. c / (x - a) ^ n) has_integral (F3 r - F3 x2)) {x2..r}" by (intro fundamental_theorem_of_calculus deriv) (auto simp: x2_def min_def split: if_splits) also have "?this \ ((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" using assms by (intro has_integral_spike_finite_eq[of "{r}"]) (auto simp: x2_def min_def split: if_splits) finally have I3: "((\x. c / max b \x - a\ ^ n) has_integral (F3 r - F3 x2)) {x2..r}" . have "((\x. c / max b \x - a\ ^ n) has_integral (F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2)) {l..r}" using assms by (intro has_integral_combine[OF _ _ has_integral_combine[OF _ _ I1 I2] I3]) (auto simp: x1_def x2_def) also have "(F1 x1 - F1 l) + ((x2 - x1) * c / b ^ n) + (F3 r - F3 x2) = F1 x1 - F1 l + F3 r - F3 x2 + (x2 - x1) * c / b ^ n" by (simp add: algebra_simps) also have "x2 - x1 = 2 * b" using assms by (simp add: x2_def x1_def min_def max_def) also have "2 * b * c / b ^ n = 2 * c / b ^ (n - 1)" using assms by (simp add: power_diff field_simps) also have "F1 x1 - F1 l + F3 r - F3 x2 = c/(n-1) * (2/b^(n-1) - 1/(a-l)^(n-1) - 1/(r-a)^(n-1))" using assms by (simp add: x1_def x2_def F1_def F3_def field_simps) also have "\ + 2 * c / b ^ (n - 1) = 2*c*(1 + 1/(n-1))/b^(n-1) - c/(n-1) * (1/(a-l)^(n-1) + 1/(r-a)^(n-1))" using assms by (simp add: field_simps) also have "1 + 1 / (n - 1) = n / (n - 1)" using assms by (simp add: field_simps) finally show ?thesis . qed lemma stirling_integral_bound_aux_integral2: fixes a b c :: real and n :: nat assumes "a \ 0" "b > 0" "c \ 0" "n > 1" obtains I where "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" "I \ 2 * c * (n / (n - 1)) / b ^ (n-1)" proof - define l' where "l' = min l (a - b - 1)" define r' where "r' = max r (a + b + 1)" define A where "A = 2 * c * (n / (n - 1)) / b ^ (n - 1)" define B where "B = c / real (n - 1) * (1 / (a - l') ^ (n - 1) + 1 / (r' - a) ^ (n - 1))" have has_int: "((\x. c / max b \x - a\ ^ n) has_integral (A - B)) {l'..r'}" using assms unfolding A_def B_def by (intro stirling_integral_bound_aux_integral1) (auto simp: l'_def r'_def) have "(\x. c / max b \x - a\ ^ n) integrable_on {l..r}" by (rule integrable_on_subinterval[OF has_integral_integrable[OF has_int]]) (auto simp: l'_def r'_def) then obtain I where has_int': "((\x. c / max b \x - a\ ^ n) has_integral I) {l..r}" by (auto simp: integrable_on_def) from assms have "I \ A - B" by (intro has_integral_subset_le[OF _ has_int' has_int]) (auto simp: l'_def r'_def) also have "\ \ A" using assms by (simp add: B_def l'_def r'_def) finally show ?thesis using that[of I] has_int' unfolding A_def by blast qed lemma stirling_integral_bound_aux': assumes n: "n > (1::nat)" and \: "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ (n - 1)" proof - obtain c where c: "norm (pbernpoly n x) \ c" for x by (rule bounded_pbernpoly[of n]) blast have c': "pbernpoly n x \ c" for x using c[of x] by (simp add: abs_real_def split: if_splits) from c[of 0] have c_nonneg: "c \ 0" by simp define D where "D = c * Beta (- (real_of_int (- int n) / 2) - 1 / 2) (1 / 2) / 2" define C where "C = max D (2*c*(n/(n-1))/sin \^(n-1))" have *: "norm (stirling_integral n s) \ C / norm s ^ (n - 1)" if s: "s \ complex_cone' \ - {0}" for s :: complex proof (rule Lim_norm_ubound[OF _ LIMSEQ_stirling_integral]) from s \ have Arg: "\Arg s\ \ \" by (auto simp: complex_cone_altdef) have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ s by auto from s have [simp]: "s \ 0" by auto show "eventually (\N. norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ C / norm s ^ (n - 1)) at_top" using eventually_gt_at_top[of "0::nat"] proof eventually_elim case (elim N) show ?case proof (cases "Re s > 0") case True have int: "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0<..}" using has_integral_mult_left[OF has_integral_Beta3[of "-n/2" "norm s ^ 2"], of c] assms True unfolding D_def by (simp add: algebra_simps) hence int': "((\x. c * (x^2 + norm s^2) powr (-n / 2)) has_integral D * (norm s ^ 2) powr (-n / 2 + 1 / 2)) {0..}" by (subst has_integral_interior [symmetric]) simp_all hence integrable: "(\x. c * (x^2 + norm s^2) powr (-n / 2)) integrable_on {0..}" by (simp add: has_integral_iff) have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c * (x^2 + norm s^2) powr (-n / 2))" proof (intro integral_norm_bound_integral s ballI integrable_ln_Gamma_aux) have [simp]: "{0<..} - {0::real..} = {}" "{0..} - {0<..} = {0::real}" by auto have "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0<..}" using int by (simp add: has_integral_iff) also have "?this \ (\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..}" by (intro integrable_spike_set_eq) auto finally show "(\x. c * (x\<^sup>2 + (cmod s)\<^sup>2) powr (real_of_int (- int n) / 2)) integrable_on {0..real N}" by (rule integrable_on_subinterval) auto next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" using True x by (auto simp: complex_eq_iff) have "norm (of_real (pbernpoly n x) / (of_real x + s) ^ n) \ c / norm (of_real x + s) ^ n" unfolding norm_divide norm_power using c by (intro divide_right_mono) simp_all also have "\ \ c / sqrt (x ^ 2 + norm s ^ 2) ^ n" proof (intro divide_left_mono mult_pos_pos zero_less_power power_mono) show "sqrt (x\<^sup>2 + (cmod s)\<^sup>2) \ cmod (complex_of_real x + s)" using x True by (simp add: cmod_def algebra_simps power2_eq_square) qed (use x True c_nonneg assms nz in \auto simp: add_nonneg_pos\) also have "sqrt (x ^ 2 + norm s ^ 2) ^ n = (x ^ 2 + norm s ^ 2) powr (1/2 * n)" by (subst powr_powr [symmetric], subst powr_realpow) (auto simp: powr_half_sqrt add_nonneg_pos) also have "c / \ = c * (x^2 + norm s^2) powr (-n / 2)" by (simp add: powr_minus field_simps) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed fact+ also have "\ \ integral {0..} (\x. c * (x^2 + norm s^2) powr (-n / 2))" using c_nonneg by (intro integral_subset_le integrable integrable_on_subinterval[OF integrable]) auto also have "\ = D * (norm s ^ 2) powr (-n / 2 + 1 / 2)" using int' by (simp add: has_integral_iff) also have "(norm s ^ 2) powr (-n / 2 + 1 / 2) = norm s powr (2 * (-n / 2 + 1 / 2))" by (subst powr_powr [symmetric]) auto also have "\ = norm s powr (-real (n - 1))" using assms by (simp add: of_nat_diff) also have "D * \ = D / norm s ^ (n - 1)" by (auto simp: powr_minus powr_realpow field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ \" . next case False have "cos \Arg s\ = cos (Arg s)" by (simp add: abs_if) also have "cos (Arg s) = Re (rcis (norm s) (Arg s)) / norm s" by (subst Re_rcis) auto also have "\ = Re s / norm s" by (subst rcis_cmod_Arg) auto also have "\ \ cos (pi / 2)" using False by (auto simp: field_simps) finally have "\Arg s\ \ pi / 2" using Arg \ by (subst (asm) cos_mono_le_eq) auto have "sin \ * norm s = sin (pi - \) * norm s" by simp also have "\ \ sin (pi - \Arg s\) * norm s" using \ Arg \\Arg s\ \ pi / 2\ by (intro mult_right_mono sin_monotone_2pi_le) auto also have "sin \Arg s\ \ 0" using Arg_bounded[of s] by (intro sin_ge_zero) auto hence "sin (pi - \Arg s\) = \sin \Arg s\\" by simp also have "\ = \sin (Arg s)\" by (simp add: abs_if) also have "\ * norm s = \Im (rcis (norm s) (Arg s))\" by (simp add: abs_mult) also have "\ = \Im s\" by (subst rcis_cmod_Arg) auto finally have abs_Im_ge: "\Im s\ \ sin \ * norm s" . have [simp]: "Im s \ 0" "s \ 0" using s \s \ \\<^sub>\\<^sub>0\ False by (auto simp: cmod_def zero_le_mult_iff complex_nonpos_Reals_iff) have "sin \ > 0" using assms by (intro sin_gt_zero) auto obtain I where I: "((\x. c / max \Im s\ \x + Re s\ ^ n) has_integral I) {0..real N}" "I \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using s c_nonneg assms False stirling_integral_bound_aux_integral2[of "-Re s" "\Im s\" c n 0 "real N"] by auto have "norm (integral {0..real N} (\x. of_real (pbernpoly n x) / (of_real x + s) ^ n)) \ integral {0..real N} (\x. c / max \Im s\ \x + Re s\ ^ n)" proof (intro integral_norm_bound_integral integrable_ln_Gamma_aux s ballI) show "(\x. c / max \Im s\ \x + Re s\ ^ n) integrable_on {0..real N}" using I(1) by (simp add: has_integral_iff) next fix x assume x: "x \ {0..real N}" have nz: "complex_of_real x + s \ 0" by (auto simp: complex_eq_iff) have "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ c / norm (complex_of_real x + s) ^ n" unfolding norm_divide norm_power using c[of x] by (intro divide_right_mono) simp_all also have "\ \ c / max \Im s\ \x + Re s\ ^ n" using c_nonneg nz abs_Re_le_cmod[of "of_real x + s"] abs_Im_le_cmod[of "of_real x + s"] by (intro divide_left_mono power_mono mult_pos_pos zero_less_power) (auto simp: less_max_iff_disj) finally show "norm (complex_of_real (pbernpoly n x) / (complex_of_real x + s) ^ n) \ \" . qed (auto simp: complex_nonpos_Reals_iff) also have "\ \ 2*c*(n/(n-1)) / \Im s\ ^ (n - 1)" using I by (simp add: has_integral_iff) also have "\ \ 2*c*(n/(n-1)) / (sin \ * norm s) ^ (n - 1)" using \sin \ > 0\ s c_nonneg abs_Im_ge by (intro divide_left_mono mult_pos_pos zero_less_power power_mono mult_nonneg_nonneg) auto also have "\ = 2*c*(n/(n-1))/sin \^(n-1) / norm s ^ (n - 1)" by (simp add: field_simps) also have "\ \ C / norm s ^ (n - 1)" by (intro divide_right_mono) (auto simp: C_def) finally show ?thesis . qed qed qed (use that assms complex_cone_inter_nonpos_Reals[of "-\" \] \ in auto) thus ?thesis by (rule that) qed lemma stirling_integral_bound: assumes "n > 0" obtains c where "\s. Re s > 0 \ norm (stirling_integral n s) \ c / Re s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux[of "Suc n"] assms obtain c where c: "\s. Re s > 0 \ norm (stirling_integral (Suc n) s) \ c / Re s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "Re s > 0" hence s': "s \ \\<^sub>\\<^sub>0" by (auto simp: complex_nonpos_Reals_iff) have "stirling_integral n s = ?f s" using s' assms by (rule stirling_integral_conv_stirling_integral_Suc) also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / Re s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / Re s ^ n" by (simp add: c1_def) also have "c2 / norm s ^ n \ c2 / Re s ^ n" using s c2_nonneg by (intro divide_left_mono power_mono complex_Re_le_cmod mult_pos_pos zero_less_power) auto also have "c1 / Re s ^ n + c2 / Re s ^ n = (c1 + c2) / Re s ^ n" using s by (simp add: field_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / Re s ^ n" by - simp_all qed qed lemma stirling_integral_bound': assumes "n > 0" and "\ \ {0<..s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral n s) \ c / norm s ^ n" proof - let ?f = "\s. of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s - of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n)" from stirling_integral_bound_aux'[of "Suc n"] assms obtain c where c: "\s::complex. s \ complex_cone' \ - {0} \ norm (stirling_integral (Suc n) s) \ c / norm s ^ n" by auto define c1 where "c1 = real n / real (Suc n) * c" define c2 where "c2 = \bernoulli (Suc n)\ / real (Suc n)" have c2_nonneg: "c2 \ 0" by (simp add: c2_def) show ?thesis proof (rule that) fix s :: complex assume s: "s \ complex_cone' \ - {0}" have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] assms s by auto have "stirling_integral n s = ?f s" using s' assms by (intro stirling_integral_conv_stirling_integral_Suc) auto also have "norm \ \ norm (of_nat n / of_nat (Suc n) * stirling_integral (Suc n) s) + norm (of_real (bernoulli (Suc n)) / (of_nat (Suc n) * s ^ n))" by (rule norm_triangle_ineq4) also have "\ = real n / real (Suc n) * norm (stirling_integral (Suc n) s) + c2 / norm s ^ n" (is "_ = ?A + ?B") by (simp add: norm_divide norm_mult norm_power c2_def field_simps del: of_nat_Suc) also have "?A \ real n / real (Suc n) * (c / norm s ^ n)" by (intro mult_left_mono c s) simp_all also have "\ = c1 / norm s ^ n" by (simp add: c1_def) also have "c1 / norm s ^ n + c2 / norm s ^ n = (c1 + c2) / norm s ^ n" using s by (simp add: divide_simps) finally show "norm (stirling_integral n s) \ (c1 + c2) / norm s ^ n" by - simp_all qed qed lemma stirling_integral_holomorphic [holomorphic_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "stirling_integral m holomorphic_on A" proof - from assms have [simp]: "z \ \\<^sub>\\<^sub>0" if "z \ A" for z using that by auto let ?f = "\s::complex. of_nat m * ((s - 1 / 2) * Ln s - s + of_real (ln (2 * pi) / 2) + (\k=1.. stirling_integral m holomorphic_on A" using assms by (intro holomorphic_cong refl) (simp_all add: field_simps ln_Gamma_stirling_complex) finally show "stirling_integral m holomorphic_on A" . qed lemma stirling_integral_continuous_on_complex [continuous_intros]: assumes m: "m > 0" and "A \ \\<^sub>\\<^sub>0 = {}" shows "continuous_on A (stirling_integral m :: _ \ complex)" by (intro holomorphic_on_imp_continuous_on stirling_integral_holomorphic assms) lemma has_field_derivative_stirling_integral_complex: fixes x :: complex assumes "x \ \\<^sub>\\<^sub>0" "n > 0" shows "(stirling_integral n has_field_derivative deriv (stirling_integral n) x) (at x)" using assms by (intro holomorphic_derivI[OF stirling_integral_holomorphic, of n "-\\<^sub>\\<^sub>0"]) auto lemma assumes n: "n > 0" and "x > 0" shows deriv_stirling_integral_complex_of_real: "(deriv ^^ j) (stirling_integral n) (complex_of_real x) = complex_of_real ((deriv ^^ j) (stirling_integral n) x)" (is "?lhs x = ?rhs x") and differentiable_stirling_integral_real: "(deriv ^^ j) (stirling_integral n) field_differentiable at x" (is ?thesis2) proof - let ?A = "{s. Re s > 0}" let ?f = "\j x. (deriv ^^ j) (stirling_integral n) (complex_of_real x)" let ?f' = "\j x. complex_of_real ((deriv ^^ j) (stirling_integral n) x)" have [simp]: "open ?A" by (simp add: open_halfspace_Re_gt) have "?lhs x = ?rhs x \ (deriv ^^ j) (stirling_integral n) field_differentiable at x" if "x > 0" for x using that proof (induction j arbitrary: x) case 0 have "((\x. Re (stirling_integral n (of_real x))) has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using 0 n by (auto intro!: derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic simp: complex_nonpos_Reals_iff) also have "?this \ (stirling_integral n has_field_derivative Re (deriv (\x. stirling_integral n x) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] 0 n by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono simp: stirling_integral_complex_of_real) finally have "stirling_integral n field_differentiable at x" by (auto simp: field_differentiable_def) with 0 n show ?case by (auto simp: stirling_integral_complex_of_real) next case (Suc j x) note IH = conjunct1[OF Suc.IH] conjunct2[OF Suc.IH] have *: "(deriv ^^ Suc j) (stirling_integral n) (complex_of_real x) = of_real ((deriv ^^ Suc j) (stirling_integral n) x)" if x: "x > 0" for x proof - have "deriv ((deriv ^^ j) (stirling_integral n)) (complex_of_real x) = vector_derivative (\x. (deriv ^^ j) (stirling_integral n) (of_real x)) (at x)" using n x by (intro vector_derivative_of_real_right [symmetric] holomorphic_on_imp_differentiable_at[of _ ?A] holomorphic_higher_deriv stirling_integral_holomorphic) (auto simp: complex_nonpos_Reals_iff) also have "\ = vector_derivative (\x. of_real ((deriv ^^ j) (stirling_integral n) x)) (at x)" using eventually_nhds_in_open[of "{0<..}" x] x by (intro vector_derivative_cong_eq) (auto elim!: eventually_mono simp: IH(1)) also have "\ = of_real (deriv ((deriv ^^ j) (stirling_integral n)) x)" by (intro vector_derivative_of_real_left holomorphic_on_imp_differentiable_at[of _ ?A] field_differentiable_imp_differentiable IH(2) x) finally show ?thesis by simp qed have "((\x. Re ((deriv ^^ Suc j) (stirling_integral n) (of_real x))) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using Suc.prems n by (intro derivative_intros has_vector_derivative_real_field field_differentiable_derivI holomorphic_on_imp_differentiable_at[of _ ?A] stirling_integral_holomorphic holomorphic_higher_deriv) (auto simp: complex_nonpos_Reals_iff) also have "?this \ ((deriv ^^ Suc j) (stirling_integral n) has_field_derivative Re (deriv ((deriv ^^ Suc j) (stirling_integral n)) (of_real x))) (at x)" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems * by (intro has_field_derivative_cong_ev refl) (auto elim!: eventually_mono) finally have "(deriv ^^ Suc j) (stirling_integral n) field_differentiable at x" by (auto simp: field_differentiable_def) with *[OF Suc.prems] show ?case by blast qed from this[OF assms(2)] show "?lhs x = ?rhs x" ?thesis2 by blast+ qed text \ Unfortunately, asymptotic power series cannot, in general, be differentiated. However, since @{term ln_Gamma} is holomorphic on the entire positive real half-space, we can differentiate its asymptotic expansion after all. To do this, we use an ad-hoc version of the more general approach outlined in Erdelyi's ``Asymptotic Expansions'' for holomorphic functions: We bound the value of the $j$-th derivative of the remainder term at some value $x$ by applying Cauchy's integral formula along a circle centred at $x$ with radius $\frac{1}{2} x$. \ lemma deriv_stirling_integral_real_bound: assumes m: "m > 0" shows "(deriv ^^ j) (stirling_integral m) \ O(\x::real. 1 / x ^ (m + j))" proof - obtain c where c: "\s. 0 < Re s \ cmod (stirling_integral m s) \ c / Re s ^ m" using stirling_integral_bound[OF m] by auto have "0 \ cmod (stirling_integral m 1)" by simp also have "\ \ c" using c[of 1] by simp finally have c_nonneg: "c \ 0" . define B where "B = c * 2 ^ (m + Suc j)" define B' where "B' = B * fact j / 2" have "eventually (\x::real. norm ((deriv ^^ j) (stirling_integral m) x) \ B' * norm (1 / x ^ (m+ j))) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) have "s \ \\<^sub>\\<^sub>0" if "s \ cball (of_real x) (x/2)" for s :: complex proof - have "x - Re s \ norm (of_real x - s)" using complex_Re_le_cmod[of "of_real x - s"] by simp also from that have "\ \ x/2" by (simp add: dist_complex_def) finally show ?thesis using elim by (auto simp: complex_nonpos_Reals_iff) qed hence "((\u. stirling_integral m u / (u - of_real x) ^ Suc j) has_contour_integral complex_of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) (circlepath (of_real x) (x/2))" using m elim by (intro Cauchy_has_contour_integral_higher_derivative_circlepath stirling_integral_continuous_on_complex stirling_integral_holomorphic) auto hence "norm (of_real (2 * pi) * \ / fact j * (deriv ^^ j) (stirling_integral m) (of_real x)) \ B / x ^ (m + Suc j) * (2 * pi * (x / 2))" proof (rule has_contour_integral_bound_circlepath) fix u :: complex assume dist: "norm (u - of_real x) = x / 2" have "Re (of_real x - u) \ norm (of_real x - u)" by (rule complex_Re_le_cmod) also have "\ = x / 2" using dist by (simp add: norm_minus_commute) finally have Re_u: "Re u \ x/2" using elim by simp have "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ c / Re u ^ m / (x / 2) ^ Suc j" using Re_u elim unfolding norm_divide norm_power dist by (intro divide_right_mono zero_le_power c) simp_all also have "\ \ c / (x/2) ^ m / (x / 2) ^ Suc j" using c_nonneg elim Re_u by (intro divide_right_mono divide_left_mono power_mono) simp_all also have "\ = B / x ^ (m + Suc j)" using elim by (simp add: B_def field_simps power_add) finally show "norm (stirling_integral m u / (u - of_real x) ^ Suc j) \ B / x ^ (m + Suc j)" . qed (insert elim c_nonneg, auto simp: B_def simp del: power_Suc) hence "cmod ((deriv ^^ j) (stirling_integral m) (of_real x)) \ B' / x ^ (j + m)" using elim by (simp add: field_simps norm_divide norm_mult norm_power B'_def) with elim m show ?case by (simp_all add: add_ac deriv_stirling_integral_complex_of_real) qed thus ?thesis by (rule bigoI) qed definition stirling_sum where "stirling_sum j m x = (-1) ^ j * (\k = 1..k\m. (of_real (bernoulli' k) * pochhammer (of_nat (Suc k)) (j - 1) * inverse x ^ (k + j)))" lemma stirling_sum_complex_of_real: "stirling_sum j m (complex_of_real x) = complex_of_real (stirling_sum j m x)" by (simp add: stirling_sum_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma stirling_sum'_complex_of_real: "stirling_sum' j m (complex_of_real x) = complex_of_real (stirling_sum' j m x)" by (simp add: stirling_sum'_def pochhammer_of_real [symmetric] del: of_nat_Suc) lemma has_field_derivative_stirling_sum_complex [derivative_intros]: "Re x > 0 \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum_real [derivative_intros]: "x > (0::real) \ (stirling_sum j m has_field_derivative stirling_sum (Suc j) m x) (at x)" unfolding stirling_sum_def [abs_def] sum_distrib_left by (rule DERIV_sum) (auto intro!: derivative_eq_intros simp del: of_nat_Suc simp: pochhammer_Suc power_diff) lemma has_field_derivative_stirling_sum'_complex [derivative_intros]: assumes "j > 0" "Re x > 0" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * complex_of_real (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma has_field_derivative_stirling_sum'_real [derivative_intros]: assumes "j > 0" "x > (0::real)" shows "(stirling_sum' j m has_field_derivative stirling_sum' (Suc j) m x) (at x)" proof (cases j) case (Suc j') from assms have [simp]: "x \ 0" by auto define c where "c = (\n. (-1) ^ Suc j * (bernoulli' n) * pochhammer (of_nat (Suc n)) j')" define T where "T = (\n x. c n * inverse x ^ (j + n))" define T' where "T' = (\n x. - (of_nat (j + n)) * c n * inverse x ^ (Suc (j + n)))" have "((\x. \k\m. T k x) has_field_derivative (\k\m. T' k x)) (at x)" using assms Suc by (intro DERIV_sum) (auto simp: T_def T'_def intro!: derivative_eq_intros simp: field_simps power_add [symmetric] simp del: of_nat_Suc power_Suc of_nat_add) also have "(\x. (\k\m. T k x)) = stirling_sum' j m" by (simp add: Suc T_def c_def stirling_sum'_def fun_eq_iff add_ac mult.assoc sum_distrib_left) also have "(\k\m. T' k x) = stirling_sum' (Suc j) m x" by (simp add: T'_def c_def Suc stirling_sum'_def sum_distrib_left sum_distrib_right algebra_simps pochhammer_Suc) finally show ?thesis . qed (insert assms, simp_all) lemma higher_deriv_stirling_sum_complex: "Re x > 0 \ (deriv ^^ i) (stirling_sum j m) x = stirling_sum (i + j) m x" proof (induction i arbitrary: x) case (Suc i) have "deriv ((deriv ^^ i) (stirling_sum j m)) x = deriv (stirling_sum (i + j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: open_halfspace_Re_gt Suc.IH) also from Suc.prems have "\ = stirling_sum (Suc (i + j)) m x" by (intro DERIV_imp_deriv has_field_derivative_stirling_sum_complex) finally show ?case by simp qed simp_all definition Polygamma_approx :: "nat \ nat \ 'a \ 'a :: {real_normed_field, ln}" where "Polygamma_approx j m = (deriv ^^ j) (\x. (x - 1 / 2) * ln x - x + of_real (ln (2 * pi)) / 2 + stirling_sum 0 m x)" lemma Polygamma_approx_Suc: "Polygamma_approx (Suc j) m = deriv (Polygamma_approx j m)" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_0: "Polygamma_approx 0 m x = (x - 1/2) * ln x - x + of_real (ln (2*pi)) / 2 + stirling_sum 0 m x" by (simp add: Polygamma_approx_def) lemma Polygamma_approx_1_complex: "Re x > 0 \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma Polygamma_approx_1_real: "x > (0 :: real) \ Polygamma_approx (Suc 0) m x = ln x - 1 / (2*x) + stirling_sum (Suc 0) m x" unfolding Polygamma_approx_Suc Polygamma_approx_0 by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps) lemma stirling_sum_2_conv_stirling_sum'_1: fixes x :: "'a :: {real_div_algebra, field_char_0}" assumes "m > 0" "x \ 0" shows "stirling_sum' 1 m x = 1 / x + 1 / (2 * x^2) + stirling_sum 2 m x" proof - have pochhammer_2: "pochhammer (of_nat k) 2 = of_nat k * of_nat (Suc k)" for k by (simp add: pochhammer_Suc eval_nat_numeral add_ac) have "stirling_sum 2 m x = (\k = Suc 0.. = (\k=0.. = (\k=0.. = (\k\m. of_real (bernoulli' k) * inverse x ^ Suc k)" by (intro sum.cong) auto also have "\ = stirling_sum' 1 m x" by (simp add: stirling_sum'_def) finally show ?thesis by (simp add: add_ac) qed lemma Polygamma_approx_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{0<..}" x] assms by (intro deriv_cong_ev) (auto elim!: eventually_mono simp: Polygamma_approx_1_real) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (simp add: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc 0)) m x = stirling_sum' 1 m x" proof - have "Polygamma_approx (Suc (Suc 0)) m x = deriv (Polygamma_approx (Suc 0) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (\x. ln x - 1 / (2*x) + stirling_sum (Suc 0) m x) x" using eventually_nhds_in_open[of "{s. Re s > 0}" x] assms by (intro deriv_cong_ev) (auto simp: open_halfspace_Re_gt elim!: eventually_mono simp: Polygamma_approx_1_complex) also have "\ = 1 / x + 1 / (2*x^2) + stirling_sum (Suc (Suc 0)) m x" using assms by (intro DERIV_imp_deriv) (auto intro!: derivative_eq_intros elim!: nonpos_Reals_cases simp: field_simps power2_eq_square) also have "\ = stirling_sum' 1 m x" using stirling_sum_2_conv_stirling_sum'_1[of m x] assms by (subst stirling_sum_2_conv_stirling_sum'_1) (auto simp: eval_nat_numeral) finally show ?thesis . qed lemma Polygamma_approx_ge_2_real: assumes "x > (0::real)" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_real) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{0<..}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_ge_2_complex: assumes "Re x > 0" "m > 0" shows "Polygamma_approx (Suc (Suc j)) m x = stirling_sum' (Suc j) m x" using assms(1) proof (induction j arbitrary: x) case (0 x) with assms show ?case by (simp add: Polygamma_approx_2_complex) next case (Suc j x) have "Polygamma_approx (Suc (Suc (Suc j))) m x = deriv (Polygamma_approx (Suc (Suc j)) m) x" by (simp add: Polygamma_approx_Suc) also have "\ = deriv (stirling_sum' (Suc j) m) x" using eventually_nhds_in_open[of "{x. Re x > 0}" x] Suc.prems by (intro deriv_cong_ev refl) (auto elim!: eventually_mono simp: Suc.IH open_halfspace_Re_gt) also have "\ = stirling_sum' (Suc (Suc j)) m x" using Suc.prems by (intro DERIV_imp_deriv derivative_intros) simp_all finally show ?case . qed lemma Polygamma_approx_complex_of_real: assumes "x > 0" "m > 0" shows "Polygamma_approx j m (complex_of_real x) = of_real (Polygamma_approx j m x)" proof (cases j) case 0 with assms show ?thesis by (simp add: Polygamma_approx_0 Ln_of_real stirling_sum_complex_of_real) next case [simp]: (Suc j') thus ?thesis proof (cases j') case 0 with assms show ?thesis by (simp add: Polygamma_approx_1_complex Polygamma_approx_1_real stirling_sum_complex_of_real Ln_of_real) next case (Suc j'') with assms show ?thesis by (simp add: Polygamma_approx_ge_2_complex Polygamma_approx_ge_2_real stirling_sum'_complex_of_real) qed qed lemma higher_deriv_Polygamma_approx [simp]: "(deriv ^^ j) (Polygamma_approx i m) = Polygamma_approx (j + i) m" by (simp add: Polygamma_approx_def funpow_add) lemma stirling_sum_holomorphic [holomorphic_intros]: "0 \ A \ stirling_sum j m holomorphic_on A" unfolding stirling_sum_def by (intro holomorphic_intros) auto lemma Polygamma_approx_holomorphic [holomorphic_intros]: "Polygamma_approx j m holomorphic_on {s. Re s > 0}" unfolding Polygamma_approx_def by (intro holomorphic_intros) (auto simp: open_halfspace_Re_gt elim!: nonpos_Reals_cases) lemma higher_deriv_lnGamma_stirling: assumes m: "m > 0" shows "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ O(\x. 1 / x ^ (m + j))" proof - have "eventually (\x. \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\ = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\) at_top" using eventually_gt_at_top[of "0::real"] proof eventually_elim case (elim x) note x = this have "\\<^sub>F y in nhds (complex_of_real x). y \ - \\<^sub>\\<^sub>0" using elim by (intro eventually_nhds_in_open) auto hence "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) (\x. (-inverse (of_nat m)) * stirling_integral m x) (complex_of_real x)" using x m by (intro higher_deriv_cong_ev refl) (auto elim!: eventually_mono simp: ln_Gamma_stirling_complex Polygamma_approx_def field_simps open_halfspace_Re_gt stirling_sum_def) also have "\ = - inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (of_real x)" using x m by (intro higher_deriv_cmult[of _ "-\\<^sub>\\<^sub>0"] stirling_integral_holomorphic) (auto simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (\x. ln_Gamma x - Polygamma_approx 0 m x) (complex_of_real x) = (deriv ^^ j) ln_Gamma (of_real x) - (deriv ^^ j) (Polygamma_approx 0 m) (of_real x)" using x by (intro higher_deriv_diff[of _ "{s. Re s > 0}"]) (auto intro!: holomorphic_intros elim!: nonpos_Reals_cases simp: open_halfspace_Re_gt) also have "(deriv ^^ j) (Polygamma_approx 0 m) (complex_of_real x) = of_real (Polygamma_approx j m x)" using x m by (simp add: Polygamma_approx_complex_of_real) also have "norm (- inverse (of_nat m) * (deriv ^^ j) (stirling_integral m) (complex_of_real x)) = inverse (real m) * \(deriv ^^ j) (stirling_integral m) x\" using x m by (simp add: norm_mult norm_inverse deriv_stirling_integral_complex_of_real) also have "(deriv ^^ j) ln_Gamma (complex_of_real x) = of_real ((deriv ^^ j) ln_Gamma x)" using x by (simp add: higher_deriv_ln_Gamma_complex_of_real) also have "norm (\ - of_real (Polygamma_approx j m x)) = \(deriv ^^ j) ln_Gamma x - Polygamma_approx j m x\" by (simp only: of_real_diff [symmetric] norm_of_real) finally show ?case . qed from bigthetaI_cong[OF this] m have "(\x::real. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x) \ \(\x. (deriv ^^ j) (stirling_integral m) x)" by simp also have "(\x::real. (deriv ^^ j) (stirling_integral m) x) \ O(\x. 1 / x ^ (m + j))" using m by (rule deriv_stirling_integral_real_bound) finally show ?thesis . qed lemma Polygamma_approx_1_real': assumes x: "(x::real) > 0" and m: "m > 0" shows "Polygamma_approx 1 m x = ln x - (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" proof - have "Polygamma_approx 1 m x = ln x - (1 / (2 * x) + (\k=Suc 0..k=Suc 0.. = (\k=0.. = (\k = Suc 0..m. bernoulli' k * inverse x ^ k / real k)" using assms by (subst sum.shift_bounds_Suc_ivl [symmetric]) (simp add: atLeastLessThanSuc_atLeastAtMost) finally show ?thesis . qed theorem assumes m: "m > 0" shows ln_Gamma_real_asymptotics: "(\x. ln_Gamma x - ((x - 1 / 2) * ln x - x + ln (2 * pi) / 2 + (\k = 1.. O(\x. 1 / x ^ m)" (is ?th1) and Digamma_real_asymptotics: "(\x. Digamma x - (ln x - (\k=1..m. bernoulli' k / real k / x ^ k))) \ O(\x. 1 / (x ^ Suc m))" (is ?th2) and Polygamma_real_asymptotics: "j > 0 \ (\x. Polygamma j x - (- 1) ^ Suc j * (\k\m. bernoulli' k * pochhammer (real (Suc k)) (j - 1) / x ^ (k + j))) \ O(\x. 1 / x ^ (m+j+1))" (is "_ \ ?th3") proof - define G :: "nat \ real \ real" where "G = (\m. if m = 0 then ln_Gamma else Polygamma (m - 1))" have *: "(\x. G j x - h x) \ O(\x. 1 / x ^ (m + j))" if "\x::real. x > 0 \ Polygamma_approx j m x = h x" for j h proof - have "(\x. G j x - h x) \ \(\x. (deriv ^^ j) ln_Gamma x - Polygamma_approx j m x)" (is "_ \ \(?f)") using that by (intro bigthetaI_cong) (auto intro: eventually_mono[OF eventually_gt_at_top[of "0::real"]] simp del: funpow.simps simp: higher_deriv_ln_Gamma_real G_def) also have "?f \ O(\x::real. 1 / x ^ (m + j))" using m by (rule higher_deriv_lnGamma_stirling) finally show ?thesis . qed note [[simproc del: simplify_landau_sum]] from *[OF Polygamma_approx_0] assms show ?th1 by (simp add: G_def Polygamma_approx_0 stirling_sum_def field_simps) from *[OF Polygamma_approx_1_real'] assms show ?th2 by (simp add: G_def field_simps) assume j: "j > 0" from *[OF Polygamma_approx_ge_2_real, of "j - 1"] assms j show ?th3 by (simp add: G_def stirling_sum'_def power_add power_diff field_simps) qed subsection \Asymptotics of the complex Gamma function\ text \ The \m\-th order remainder of Stirling's formula for $\log\Gamma$ is $O(s^{-m})$ uniformly over any complex cone $\text{Arg}(z) \leq \alpha$, $z\neq 0$ for any angle $\alpha\in(0, \pi)$. This means that there is bounded by $c z^{-m}$ for some constant $c$ for all $z$ in this cone. \ context fixes F and \ assumes \: "\ \ {0<.. principal (complex_cone' \ - {0})" begin lemma stirling_integral_bigo: fixes m :: nat assumes m: "m > 0" shows "stirling_integral m \ O[F](\s. 1 / s ^ m)" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF \m > 0\ \] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] \ by simp finally have "c \ 0" . have "eventually (\s. s \ complex_cone' \ - {0}) F" unfolding F_def by (auto simp: eventually_principal) hence "eventually (\s. norm (stirling_integral m s) \ c * norm (1 / s ^ m)) F" by eventually_elim (use c in \simp add: norm_divide norm_power\) thus "stirling_integral m \ O[F](\s. 1 / s ^ m)" by (intro bigoI[of _ c]) auto qed end text \ The following is a more explicit statement of this: \ theorem ln_Gamma_complex_asymptotics_explicit: fixes m :: nat and \ :: real assumes "m > 0" and "\ \ {0<.. complex" where "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1/2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ norm (R s) \ C / norm s ^ m" proof - obtain c where c: "\s. s \ complex_cone' \ - {0} \ norm (stirling_integral m s) \ c / norm s ^ m" using stirling_integral_bound'[OF assms] by blast have "0 \ norm (stirling_integral m 1 :: complex)" by simp also have "\ \ c" using c[of 1] assms by simp finally have "c \ 0" . define R where "R = (\s::complex. stirling_integral m s / of_nat m)" show ?thesis proof (rule that) from ln_Gamma_stirling_complex[of _ m] assms show "\s::complex. s \ \\<^sub>\\<^sub>0 \ ln_Gamma s = (s - 1 / 2) * ln s - s + ln (2 * pi) / 2 + (\k=1..s. s \ 0 \ \Arg s\ \ \ \ cmod (R s) \ c / real m / cmod s ^ m" proof (safe, goal_cases) case (1 s) show ?case using 1 c[of s] assms by (auto simp: complex_cone_altdef abs_le_iff R_def norm_divide field_simps) qed qed qed text \ Lastly, we can also derive the asymptotics of $\Gamma$ itself: \[\Gamma(z) \sim \sqrt{2\pi / z} \left(\frac{z}{e}\right)^z\] uniformly for $|z|\to\infty$ within the cone $\text{Arg}(z) \leq \alpha$ for $\alpha\in(0,\pi)$: \ context fixes F and \ assumes \: "\ \ {0<.. inf at_infinity (principal (complex_cone' \))" begin lemma Gamma_complex_asymp_equiv: "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2))" proof - define I :: "complex \ complex" where "I = stirling_integral 1" have "eventually (\s. s \ complex_cone' \) F" by (auto simp: eventually_inf_principal F_def) moreover have "eventually (\s. s \ 0) F" unfolding F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto ultimately have "eventually (\s. Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)) F" proof eventually_elim case (elim s) from elim have s': "s \ \\<^sub>\\<^sub>0" using complex_cone_inter_nonpos_Reals[of "-\" \] \ by auto from elim have [simp]: "s \ 0" by auto from s' have "Gamma s = exp (ln_Gamma s)" unfolding Gamma_complex_altdef using nonpos_Ints_subset_nonpos_Reals by auto also from s' have "ln_Gamma s = (s-1/2) * Ln s - s + complex_of_real (ln (2 * pi) / 2) - I s" by (subst ln_Gamma_stirling_complex[of _ 1]) (simp_all add: exp_add exp_diff I_def) also have "exp \ = exp ((s - 1 / 2) * Ln s) / exp s * exp (complex_of_real (ln (2 * pi) / 2)) / exp (I s)" unfolding exp_diff exp_add by (simp add: exp_diff exp_add) also have "exp ((s - 1 / 2) * Ln s) = s powr (s - 1 / 2)" by (simp add: powr_def) also have "exp (complex_of_real (ln (2 * pi) / 2)) = sqrt (2 * pi)" by (subst exp_of_real) (auto simp: powr_def simp flip: powr_half_sqrt) also have "exp s = exp 1 powr s" by (simp add: powr_def) also have "s powr (s - 1 / 2) / exp 1 powr s = (s powr s / exp 1 powr s) / s powr (1/2)" by (subst powr_diff) auto also have *: "Ln (s / exp 1) = Ln s - 1" using Ln_divide_of_real[of "exp 1" s] by (simp flip: exp_of_real) hence "s powr s / exp 1 powr s = (s / exp 1) powr s" unfolding powr_def by (subst *) (auto simp: exp_diff field_simps) finally show "Gamma s = sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s)" by (simp add: algebra_simps) qed hence "Gamma \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / exp (I s))" by (rule asymp_equiv_refl_ev) also have "\ \[F] (\s. sqrt (2 * pi) * (s / exp 1) powr s / s powr (1 / 2) / 1)" proof (intro asymp_equiv_intros) have "F \ principal (complex_cone' \ - {0})" unfolding le_principal F_def eventually_inf_principal using eventually_not_equal_at_infinity by eventually_elim auto moreover have "I \ O[principal (complex_cone' \ - {0})](\s. 1 / s)" using stirling_integral_bigo[of \ 1] \ unfolding F_def by (simp add: I_def) ultimately have "I \ O[F](\s. 1 / s)" by (rule landau_o.big.filter_mono) also have "(\s. 1 / s) \ o[F](\s. 1)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" hence "eventually (\z::complex. norm z \ 1 / c) at_infinity" by (auto simp: eventually_at_infinity) moreover have "eventually (\z::complex. z \ 0) at_infinity" by (rule eventually_not_equal_at_infinity) ultimately show "eventually (\z::complex. norm (1 / z) \ c * norm (1 :: complex)) F" unfolding F_def eventually_inf_principal by eventually_elim (use \c > 0\ in \auto simp: norm_divide field_simps\) qed finally have "I \ o[F](\s. 1)" . from smalloD_tendsto[OF this] have [tendsto_intros]: "(I \ 0) F" by simp show "(\x. exp (I x)) \[F] (\x. 1)" by (rule asymp_equivI' tendsto_eq_intros refl | simp)+ qed finally show ?thesis by simp qed end end