diff --git a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy --- a/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy +++ b/thys/Irrationality_J_Hancl/Irrationality_J_Hancl.thy @@ -1,1090 +1,1087 @@ (* File: Irrationality_J_Hancl.thy Authors: Angeliki Koutsoukou-Argyraki and Wenda Li, Computer Laboratory, University of Cambridge, UK. Email: ak2110@cam.ac.uk, wl302@cam.ac.uk *) section \Irrational Rapidly Convergent Series\ theory Irrationality_J_Hancl imports "HOL-Analysis.Analysis" "HOL-Decision_Procs.Approximation" begin text \This is the formalisation of a proof by J. Hancl, in particular of the proof of his Theorem 3 in the paper: Irrational Rapidly Convergent Series, Rend. Sem. Mat. Univ. Padova, Vol 107 (2002). The statement asserts the irrationality of the sum of a series consisting of rational numbers defined using sequences that fulfill certain properties. Even though the statement is number-theoretic, the proof uses only arguments from introductory Analysis.\ text \We show the central result (theorem Hancl3) by proof by contradiction, by making use of some of the auxiliary lemmas. To this end, assuming that the sum is a rational number, for a quantity $\textrm{ALPHA}(n)$ we show that $\textrm{ALPHA}(n) \geq 1$ for all $n \in \mathbb{N}$. After that we show that we can find an $n \in \mathbb{N}$ for which $\textrm{ALPHA}(n) < 1$ which yields a contradiction and we thus conclude that the sum of the series is a rational number. We finally give an immediate application of theorem Hancl3 for a specific series (corollary Hancl3corollary, requiring lemma summable\_ln\_plus) which corresponds to Corollary 2 in the original paper by J. Hancl. \ hide_const floatarith.Max subsection \Misc\ lemma filterlim_sequentially_iff: "filterlim f F1 sequentially \ filterlim (\x. f (x+k)) F1 sequentially" unfolding filterlim_iff apply (subst eventually_sequentially_seg) by auto lemma filterlim_realpow_sequentially_at_top: "(x::real) > 1 \ filterlim ((^) x) at_top sequentially" apply (rule LIMSEQ_divide_realpow_zero[THEN filterlim_inverse_at_top,of _ 1,simplified]) by auto lemma filterlim_at_top_powr_real: fixes g::"'b \ real" assumes "filterlim f at_top F" "(g \ g') F" "g'>1" shows "filterlim (\x. g x powr f x) at_top F" proof - have "filterlim (\x. ((g'+1)/2) powr f x) at_top F" proof (subst filterlim_at_top_gt[of _ _ 1],rule+) fix Z assume "Z>(1::real)" have "\\<^sub>F x in F. ln Z \ ln (((g' + 1) / 2) powr f x)" using assms(1)[unfolded filterlim_at_top,rule_format,of "ln Z / ln ((g' + 1) / 2)"] apply (eventually_elim) using \g'>1\ by (auto simp:ln_powr divide_simps) then show "\\<^sub>F x in F. Z \ ((g' + 1) / 2) powr f x" apply (eventually_elim) apply (subst (asm) ln_le_cancel_iff) using \Z>1\ \g'>1\ by auto qed moreover have "\\<^sub>F x in F. ((g'+1)/2) powr f x \ g x powr f x" proof - have "\\<^sub>F x in F. g x > (g'+1)/2" apply (rule order_tendstoD[OF assms(2)]) using \g'>1\ by auto moreover have "\\<^sub>F x in F. f x>0" using assms(1)[unfolded filterlim_at_top_dense,rule_format,of 0] . ultimately show ?thesis apply eventually_elim using \g'>1\ by (auto intro!: powr_mono2) qed ultimately show ?thesis using filterlim_at_top_mono by fast qed lemma powrfinitesum: fixes a::real and s::nat assumes "s \ n" shows " (\j=s..(n::nat).(a powr (2^j))) = a powr (\j=s..(n::nat).(2^j)) " using \s \ n\ proof(induct n) case 0 then show ?case by auto next case (Suc n) have ?case when "s\n" using Suc.hyps by (metis Suc.prems add.commute linorder_not_le powr_add prod.nat_ivl_Suc' sum.cl_ivl_Suc that) moreover have ?case when "s=Suc n" proof- have "(\j = s..Suc n. a powr 2 ^ j) =(a powr 2 ^(Suc n)) " using \s=Suc n\ by simp also have "(a powr 2 ^(Suc n) ) = a powr sum ((^) 2) {s..Suc n} " using that by auto ultimately show " (\j = s..Suc n. a powr 2 ^ j) = a powr sum ((^) 2) {s..Suc n}" using \s\Suc n\ by linarith qed ultimately show ?case using \s\Suc n\ by linarith qed lemma summable_ratio_test_tendsto: fixes f :: "nat \ 'a::banach" assumes "c < 1" and "\n. f n\0" and "(\n. norm (f (Suc n)) / norm (f n)) \ c" shows "summable f" proof - obtain N where N_dist:"\n\N. dist (norm (f (Suc n)) / norm (f n)) c < (1-c)/2" using assms unfolding tendsto_iff eventually_sequentially by (meson diff_gt_0_iff_gt zero_less_divide_iff zero_less_numeral) have "norm (f (Suc n)) / norm (f n) \ (1+c)/2" when "n\N" for n using N_dist[rule_format,OF that] \c<1\ apply (auto simp add:field_simps dist_norm) by argo then have "norm (f (Suc n)) \ (1+c)/2 * norm (f n)" when "n\N" for n using that assms(2)[rule_format, of n] by (auto simp add:divide_simps) moreover have "(1+c)/2 < 1" using \c<1\ by auto ultimately show ?thesis using summable_ratio_test[of _ N f] by blast qed lemma summable_ln_plus: fixes f::"nat \ real" assumes "summable f" "\n. f n>0" shows "summable (\n. ln (1+f n))" proof (rule summable_comparison_test_ev[OF _ assms(1)]) have "ln (1 + f n) > 0" for n by (simp add: assms(2) ln_gt_zero) moreover have "ln (1 + f n) \ f n" for n apply (rule ln_add_one_self_le_self2) using assms(2)[rule_format,of n] by auto ultimately show "\\<^sub>F n in sequentially. norm (ln (1 + f n)) \ f n" by (auto intro!: eventuallyI simp add:less_imp_le) qed lemma suminf_real_offset_le: fixes f :: "nat \ real" assumes f: "\i. 0 \ f i" and "summable f" shows "(\i. f (i + k)) \ suminf f" proof - have "(\n. \i (\i. f (i + k))" using summable_sums[OF \summable f\] by (simp add: assms(2) summable_LIMSEQ summable_ignore_initial_segment) moreover have "(\n. \i (\i. f i)" using summable_sums[OF \summable f\] by (simp add: sums_def atLeast0LessThan f) then have "(\n. \i (\i. f i)" by (rule LIMSEQ_ignore_initial_segment) ultimately show ?thesis proof (rule LIMSEQ_le, safe intro!: exI[of _ k]) fix n assume "k \ n" have "(\ii (\i. i + k)) i)" by simp also have "\ = (\i\(\i. i + k) ` {.. \ sum f {..i sum f {.. n" shows " (\i=s..n. 2^i) < (2^(n+1) :: real) " using assms proof (induct n) case 0 show ?case by auto next case (Suc n) have ?case when "s=n+1" using that by auto moreover have ?case when "s \ n+1" proof - have"(\i=s..(n+1). 2^i ) = (\i=s..n. 2^i ) + (2::real)^(n+1) " using sum.cl_ivl_Suc \s \ Suc n \ by (auto simp add:add.commute) also have "... < (2) ^ (n +1) + (2)^(n+1)" proof - have "s \n" using that \s \ Suc n \ by auto then show ?thesis using Suc.hyps \s \ n\ by linarith qed also have "... = 2^(n+2)" by simp finally show "(\i=s..(Suc n). 2^i )< (2::real)^(Suc n+1)" by auto qed ultimately show ?case by blast qed subsection \Auxiliary lemmas and the main proof\ lemma showpre7: fixes a b ::"nat\int " and q p::int assumes "q>0" and "p>0"and a: "\n. a n>0" and "\n. b n>0" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*(suminf (\(j::nat). (b (j+n+1)/ a (j+n+1 )))) = ((\j=1..n. of_int( a j)))*(p -q* (\j=1..n. b j / a j)) " proof - define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "(\j. ff (j+n)) = (\n. ff n) - sum ff {..j=1..n. ff (j-1))" proof - have "sum ff {..j=1..n. ff (j-1))" apply (subst sum_bounds_lt_plus1[symmetric]) by simp then show ?thesis unfolding ff_def by auto qed finally have "(\j. ff (j + n)) = p / q - (\j = 1..n. ff (j - 1))" . then have "q*(\j. ff (j + n)) = p - q*(\j = 1..n. ff (j - 1))" using \q>0\ by (auto simp add:field_simps) then have "aa*q*(\j. ff (j + n)) = aa*(p - q*(\j = 1..n. ff (j - 1)))" by auto then show ?thesis unfolding aa_def ff_def by auto qed lemma show7: fixes d::"nat\real" and a b::"nat\int " and q p::int assumes "q \1" and "p \ 1" and a: "\n. a n \ 1" and "\n. b n \ 1" and assumerational:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" shows "q*((\j=1..n. of_int( a j)))*( suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1 )))) \ 1 " (is "?L \ _") proof- define LL where "LL=?L" define aa where "aa=(\j = 1..n. real_of_int (a j))" define ff where "ff=(\i. real_of_int (b (i+1)) / real_of_int (a (i+1)))" have "?L > 0" proof - have "aa > 0" unfolding aa_def using a apply (induct n,auto) by (simp add: int_one_le_iff_zero_less prod_pos) moreover have "(\j. ff (j + n)) > 0" proof (rule suminf_pos) have "summable ff" unfolding ff_def using assumerational using summable_def by blast then show " summable (\j. ff (j + n))" using summable_iff_shift[of ff n] by auto show " \i. 0 < ff (i + n)" unfolding ff_def using a assms(4) int_one_le_iff_zero_less by auto qed ultimately show ?thesis unfolding aa_def ff_def using \q\1\ by auto qed moreover have "?L \ \" proof - have "?L = aa *(p -q* ( \j=1..n. b j / a j))" unfolding aa_def apply (rule showpre7) using assms int_one_le_iff_zero_less by auto also have "... = aa * p - q * (\j=1..n. aa * b j / a j)" by (auto simp add:algebra_simps sum_distrib_left) also have "... = prod a {1..n} * p - q * (\j = 1..n. b j * prod a ({1..n} - {j}))" proof - have "(\j=1..n. aa * b j / a j) = (\j=1..n. b j * prod a ({1..n} - {j}))" unfolding of_int_sum proof (rule sum.cong) fix x assume "x \ {1..n}" then have "(\i = 1..n. real_of_int (a i)) = a x * (\i\{1..n} - {x}. real_of_int (a i))" apply (rule_tac prod.remove) by auto then have "aa / real_of_int (a x) = prod a ({1..n} - {x})" unfolding aa_def using a[rule_format,of x] by (auto simp add:field_simps) then show "aa * b x / a x = b x * prod a ({1..n} - {x})" by (metis mult.commute of_int_mult times_divide_eq_right) qed simp moreover have "aa * p = (\j = 1..n. (a j)) * p" unfolding aa_def by auto ultimately show ?thesis by force qed also have "... \ \" using Ints_of_int by blast finally show ?thesis . qed ultimately show ?thesis apply (fold LL_def) by (metis Ints_cases int_one_le_iff_zero_less not_less of_int_0_less_iff of_int_less_1_iff) qed lemma show8: fixes d ::"nat\real " and a :: "nat\int" and s k::nat assumes "A > 1" and d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "convergent_prod d" and assu2: "\n \ s. ( A/ (of_int (a n)) powr(1/of_int (2^n)))> prodinf (\j. d(n +j))" shows "\n\s. prodinf (\j. d(j+n)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n}))" proof(rule,rule) fix n assume "s \ n" define sp where "sp = (\n. prodinf (\j. d(j+n)))" define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have "sp i \ sp n" when "i\n" for i proof - have "(\j. d (j + i)) = (\ia. d (ia + (n - i) + i)) * (\iaconvergent_prod d\ convergent_prod_iff_shift[of d i] by simp subgoal for j using d[rule_format,of "j+i"] by auto done then have "sp i = sp n * (\jn\i\ by (auto simp:algebra_simps) moreover have "sp i>1" "sp n>1" unfolding sp_def using convergent_prod_iff_shift \convergent_prod d\ d by (auto intro!:less_1_prodinf) moreover have "(\j1" apply (rule prod_ge_1) using d less_imp_le by auto ultimately show ?thesis by auto qed moreover have "\j\s. A / ff j > sp j" unfolding ff_def sp_def using assu2 by (auto simp:algebra_simps) ultimately have "\j. s\j \ j\n \ A / ff j > sp n" by force then show "sp n< A / Max (ff ` {s..n})" by (metis (mono_tags, hide_lams) Max_in \n\s\ atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty order_refl) qed lemma auxiliary1_9: fixes d ::"nat\real" and a::"nat\int " and s m::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and "n \ m" and " m \ s" and auxifalse_assu: "\n\m. (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (d (n+1))* (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} ))" shows "(of_int (a (n+1))) powr(1 /of_int (2^(n+1))) < (\j=(m+1)..(n+1). d j) * (Max ((\ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..m}))" proof- define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" have [simp]:"ff j > 0" for j unfolding ff_def by (metis a less_numeral_extra(3) of_int_0_less_iff powr_gt_zero) have ff_asm:"ff (n+1) < d (n+1) * Max (ff ` {s..n})" when "n\m" for n using auxifalse_assu that unfolding ff_def by simp from \n\m\ have Q: "(Max( ff ` {s..n} )) \ (\j=(m+1)..n. d j)* (Max (ff ` {s..m}))" proof(induct n) case 0 then show ?case using \m\s\ by simp next case (Suc n) have ?case when "m=Suc n" using that by auto moreover have ?case when "m\Suc n" proof - have "m \ n" using that Suc(2) by simp then have IH:"Max (ff ` {s..n}) \ prod d {m + 1..n} * Max (ff ` {s..m})" using Suc(1) by linarith have "Max (ff ` {s..Suc n}) = Max (ff ` {s..n} \ {ff (Suc n)})" using Suc.prems assms(5) atLeastAtMostSuc_conv by auto also have "... = max (Max (ff ` {s..n})) (ff (Suc n))" using Suc.prems assms(5) max_def sup_assoc that by auto also have "... \ max (Max (ff ` {s..n})) (d (n+1) * Max (ff ` {s..n}))" apply (rule max.mono) using ff_asm[of n] \ m \ Suc n\ that \s\m\ by auto also have "... \ Max (ff ` {s..n}) * max 1 (d (n+1))" proof - have "Max (ff ` {s..n}) \0" by (metis (mono_tags, hide_lams) Max_in \\j. 0 < ff j\ \m \ n\ assms(5) atLeastAtMost_iff empty_iff finite_atLeastAtMost finite_imageI imageE image_is_empty less_eq_real_def) then show ?thesis using max_mult_distrib_right by (simp add: max_mult_distrib_right mult.commute) qed also have "... = Max (ff ` {s..n}) * d (n+1)" using d[rule_format, of "n+1"] by auto also have "... \ prod d {m + 1..n} * Max (ff ` {s..m}) * d (n+1)" using IH d[rule_format, of "n+1"] by auto also have "... = prod d {m + 1..n+1} * Max (ff ` {s..m})" using \n\m\ by (simp add:prod.nat_ivl_Suc' algebra_simps) finally show ?case by simp qed ultimately show ?case by blast qed then have "d (n+1) * Max (ff ` {s..n} ) \ (\j=(m+1)..(n+1). d j)* (Max (ff ` {s..m}))" using \m\n\ d[rule_format,of "Suc n"] by (simp add:prod.nat_ivl_Suc') then show ?thesis using ff_asm[of n] \s\m\ \m\n\ unfolding ff_def by auto qed lemma show9: fixes d ::"nat\real " and a :: "nat\int " and s ::nat and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and "convergent_prod d" and 8: "\n\s. prodinf (\j. d( n+j)) < A/(Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n})) " shows "\m \s. \n\m. ( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" proof (rule ccontr) define ff where "ff = (\(j::nat). (real_of_int (a j)) powr(1 /of_int (2^j)))" assume assumptioncontra: " \ (\m \s. \n\m. ( (ff (n+1)) \ (d (n+1))* (Max ( ff ` {s..n}))))" then obtain t where "t\s" and ttt: " \n\t. ( (ff (n+1)) < (d (n+1))* (Max ( ff ` {s..n} ) ))" by fastforce define B where "B=prodinf (\j. d(t+1+j))" have "B>0" unfolding B_def apply (rule less_0_prodinf) subgoal using convergent_prod_iff_shift[of d "t+1"] \convergent_prod d\ by (auto simp:algebra_simps) subgoal using d le_less_trans zero_le_one by blast done have "A \ B * Max ( ff ` {s..t})" proof (rule tendsto_le[of sequentially "\n. (\j=(t+1)..(n+1). d j) * Max ( ff ` {s..t})" _ "\n. ff (n+1)"]) show "(\n. ff (n + 1)) \ A" using assu1[folded ff_def] LIMSEQ_ignore_initial_segment by blast have "(\n. prod d {t + 1..n + 1}) \ B" proof - have "(\n. \i\n. d (t + 1 + i)) \ B" apply (rule convergent_prod_LIMSEQ[of "(\j. d(t+1+j))",folded B_def]) using \convergent_prod d\ convergent_prod_iff_shift[of d "t+1"] by (simp add:algebra_simps) then have "(\n. \i\{0..n}. d (i+(t + 1))) \ B" using atLeast0AtMost by (auto simp:algebra_simps) then have "(\n. prod d {(t + 1)..n + (t + 1)}) \ B" apply (subst (asm) prod.shift_bounds_cl_nat_ivl[symmetric]) by simp from seq_offset_neg[OF this,of "t"] show "(\n. prod d {t + 1..n+1}) \ B" apply (elim Lim_transform) apply (rule LIMSEQ_I) apply (rule exI[where x="t+1"]) by auto qed then show "(\n. prod d {t + 1..n + 1} * Max (ff ` {s..t})) \ B * Max (ff ` {s..t})" by (auto intro:tendsto_eq_intros) have "\\<^sub>F n in sequentially. (ff (n+1)) < (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" unfolding eventually_sequentially ff_def using auxiliary1_9[OF d a \s>0\ _ \t\s\ ttt[unfolded ff_def]] by blast then show "\\<^sub>F n in sequentially. (ff (n+1)) \ (\j=(t+1)..(n+1). d j) * (Max ( ff ` {s..t}))" by (eventually_elim,simp) qed simp also have "... \ B * Max ( ff ` {s..t+1})" proof - have "Max (ff ` {s..t}) \ Max (ff ` {s..t + 1})" apply (rule Max_mono) using \t\s\ by auto then show ?thesis using \B>0\ by auto qed finally have "A \ B * Max (ff ` {s..t + 1})" unfolding B_def . moreover have "B < A / Max (ff ` {s..t + 1})" using 8[rule_format, of "t+1",folded ff_def B_def] \s\t\ by auto moreover have "Max (ff ` {s..t+1})>0" using \A \ B * Max (ff ` {s..t + 1})\ \B>0\ \A>1\ zero_less_mult_pos [of B "Max (ff ` {s..Suc t})"] by simp ultimately show False by (auto simp add:field_simps) qed lemma show10: fixes d ::"nat\real " and a ::"nat\int " and s::nat assumes d: "\n. d n> 1" and a: "\n. a n>0" and " s>0" and 9: "\m \s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" shows "\m\s. \n\m. (((d (n+1))powr(2^(n+1))) * (\j=1..n. of_int( a j)) * (1/ (\j=1..s-1. (of_int( a j) )))) \ (a (n+1)) " proof (rule,rule) fix m assume "s \ m" from 9[rule_format,OF this] obtain n where "n\m" and asm_9:"( (of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ( ( \ (j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" by auto with \s\m\ have "s\n" by auto have prod:"(\j=1..n. real_of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) ))) = (\j=s..n. of_int( a j))" proof - define f where "f= (\j. real_of_int( a j))" have "{Suc 0..n}= {Suc 0..s - Suc 0} \ {s..n}" using \n\s\ \s >0\ by auto then have "(\j=1..n. f j) = (\j=1..s-1. f j) * (\j=s..n. f j)" apply (subst prod.union_disjoint[symmetric]) by auto moreover have "(\j=1..s-1. f j) > 0 " apply (rule linordered_semidom_class.prod_pos) using a unfolding f_def by auto then have "(\j=1..s-1. f j) \ 0" by argo ultimately show ?thesis unfolding f_def by auto qed then have " (((d (n+1))powr(2^(n+1) )) * (\j=1..n. of_int( a j)) * ( 1/ (\j=1..s-1. (of_int( a j) )))) =(((d (n+1))powr(2^(n+1) )) * (\j=s..n. of_int( a j)))" proof - define f where "f= (\j. real_of_int( a j))" define c where "c = (d (n+1))powr(2^(n+1))" show ?thesis using prod apply (fold f_def c_def) by (metis mult.assoc) qed also have "... \ ((d (n+1))powr(2^(n+1) ) * (\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )) powr(2^i)) )" proof (rule mult_left_mono) show "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto show "(\j = s..n. real_of_int (a j)) \ (\i = s..n. Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i)" proof (rule prod_mono) fix i assume i: "i \ {s..n}" have "real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i" unfolding powr_powr by (simp add: a less_eq_real_def) also have "\ \ (Max(( \ (j::nat). ( real_of_int( a j) powr(1 /real_of_int (2^j)))) ` {s..n}))powr(2^i)" proof (rule powr_mono2) show "real_of_int (a i) powr (1 / real_of_int (2 ^ i)) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n})" apply (rule Max_ge) apply auto using i by blast qed simp_all finally have "real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" . then show "0 \ real_of_int (a i) \ real_of_int (a i) \ Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) powr 2 ^ i" using a i by (metis \real_of_int (a i) = (real_of_int (a i) powr (1 / real_of_int (2 ^ i))) powr 2 ^ i\ powr_ge_pzero) qed qed also have "... = ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat). ( of_int( a j) powr (1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof- have "((d (n+1))powr(2^(n+1) ))\1 " by (metis Transcendental.log_one d le_powr_iff zero_le_numeral zero_le_power zero_less_one) moreover have "(\i=s..n. (Max(( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } ) ) powr(2^i)) = (Max(( \ (j::nat). ( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n } )) powr (\i=s..n. 2^i ) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" show ?thesis apply (fold ff_def) using \s\n\ powrfinitesum by auto qed ultimately show ?thesis by auto qed also have "... \ ((d (n+1))powr(2^(n+1) )) * (Max(( \ (j::nat).( of_int( a j) powr(1 /of_int (2^j)) )) ` {s..n })) powr(2^(n+1)) " proof - define ff where "ff = Max (( \ (j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n } )" have " sum ((^) 2) {s..n} < (2::real) ^ (n + 1)" using factt \s\n\ by auto moreover have "1 \ ff" proof - define S where "S=(\(j::nat). ( of_int( a j) powr(1 /real_of_int (2^j)) )) ` {s..n }" have "finite S" unfolding S_def by auto moreover have "S\{}" unfolding S_def using \s\n\ by auto moreover have "\x\S. x\1" proof- have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ 1 " apply (rule ge_one_powr_ge_zero) apply auto by (simp add: a int_one_le_iff_zero_less) moreover have " ( of_int( a s) powr(1 /real_of_int (2^s)) ) \ S" unfolding S_def apply (rule rev_image_eqI[where x=s]) using \s\n\ by auto ultimately show ?thesis by auto qed ultimately show ?thesis using Max_ge_iff[of S 1] unfolding S_def ff_def by blast qed moreover have "0 \ (d (n + 1)) powr 2 ^ (n + 1)" by auto ultimately show ?thesis apply(fold ff_def) apply (rule mult_left_mono) apply (rule powr_mono) by auto qed also have "... = ((d (n+1)) * (Max((\j. (of_int( a j) powr(1 /of_int (2^j)))) ` {s..n}))) powr(2^(n+1))" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo moreover have "Max ss \0" proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed ultimately show ?thesis apply (fold ss_def) using powr_mult by auto qed also have "... \ (real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))) powr 2 ^ (n + 1)" proof - define ss where "ss = (\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}" show ?thesis proof (fold ss_def,rule powr_mono2) have "Max ss \0" \ \NOTE: we are repeating the same proof, so it may be a good idea to put this conclusion in an outer block so that it can be reused (without reproving).\ proof - have "(a s) powr (1 / (2 ^ s)) \ 0" by auto moreover have "(a s) powr (1 / (2 ^ s)) \ ss" unfolding ss_def apply (rule_tac x=s in rev_image_eqI) using \s\n\ by auto moreover have "finite ss" "ss \ {}" unfolding ss_def using \s\n\ by auto ultimately show ?thesis using Max_ge_iff[of ss 0] by blast qed moreover have "d (n + 1) \0" using d[rule_format,of "n+1"] by argo ultimately show "0 \ (d (n + 1)) * Max ss" by auto show "(d (n + 1)) * Max ss \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" using asm_9[folded ss_def] . qed simp qed also have "... = (of_int (a (n+1)))" by (simp add: a less_eq_real_def pos_add_strict powr_powr) finally show "\n\m. d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / (\j = 1..s - 1. real_of_int (a j))) \ real_of_int (a (n + 1))" using \n\m\ \m\s\ apply (rule_tac x=n in exI) by auto qed lemma lasttoshow: fixes d ::"nat\real " and a b ::"nat\int " and q::int and s::nat assumes d: "\n. d n> 1" and a:"\n. a n>0" and "s>0" and "q>0" and "A > 1" and b:"\n. b n>0" and 9: "\m\s. \n\m. ((of_int (a (n+1))) powr(1 /of_int (2^(n+1))) \ (d (n+1))* (Max ((\(j::nat). (of_int (a j)) powr(1 /of_int (2^j))) ` {s..n} )))" and assu3: "filterlim( \ n. (d n)^(2^n)/ b n) at_top sequentially " and 5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" shows "\n. q*((\j=1..n. real_of_int(a j))) * suminf (\(j::nat). (b (j+n+1)/ a (j+n+1)))<1" proof- define as where "as=(\j = 1..s - 1. real_of_int (a j))" obtain n where "n\s" and n_def1:"real_of_int q * as * 2 * real_of_int (b (n + 1)) / d (n + 1) powr 2 ^ (n + 1) < 1" and n_def2:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. real_of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" and n_def3:"(\j. (b (n+1 + j)) / (a (n+1 + j))) \ 2 * b (n+1) / a (n+1)" proof - have *:"(\n. real_of_int (b n) / d n ^ 2 ^ n) \ 0" using tendsto_inverse_0_at_top[OF assu3] by auto then have "(\n. real_of_int (b n) / d n powr 2 ^ n) \ 0" proof - have "d n ^ 2 ^ n = d n powr (of_nat (2 ^ n))" for n apply (subst powr_realpow) using d[rule_format, of n] by auto then show ?thesis using * by auto qed from tendsto_mult_right_zero[OF this,of "q * as * 2"] have "(\n. q * as * 2 * b n / d n powr 2 ^ n) \ 0" by auto then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1" by (elim order_tendstoD) simp then have "\\<^sub>F n in sequentially. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" using 5 by eventually_elim auto then obtain N where N_def:"\n\N. q * as * 2 * b n / d n powr 2 ^ n < 1 \ (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding eventually_sequentially by auto obtain n where "n\s" and "n\N" and n_def:"d (n + 1) powr 2 ^ (n + 1) * (\j = 1..n. of_int (a j)) * (1 / as) \ real_of_int (a (n + 1))" using show10[OF d a \s>0\ 9,folded as_def,rule_format,of "max s N"] by auto with N_def[rule_format, of "n+1"] that[of n] show ?thesis by auto qed define pa where "pa = (\j = 1..n. real_of_int (a j))" define dn where "dn = d (n + 1) powr 2 ^ (n + 1)" have [simp]:"dn >0" "as > 0" subgoal unfolding dn_def by (metis d not_le numeral_One powr_gt_zero zero_le_numeral) subgoal unfolding as_def by (simp add: a prod_pos) done have [simp]:"pa>0" unfolding pa_def using a by (simp add: prod_pos) have K: "q* (\j=1..n. real_of_int (a j)) * suminf (\ (j::nat). (b (j+n+1)/ a (j+n+1))) \q* (\j=1..n. real_of_int (a j)) *2* (b (n+1))/(a( n+1))" apply (fold pa_def) using mult_left_mono[OF n_def3,of "real_of_int q * pa"] \n\s\ \pa>0\ \q>0\ by (auto simp add:algebra_simps) also have KK:"... \ 2*q* (\j=1..n. real_of_int (a j)) * (b(n+1))* ((\j=1..s-1. real_of_int( a j))/((d (n+1))powr(2^(n+1)) * (\j=1..n. real_of_int ( a j))))" proof - have " dn * pa * (1 / as) \ real_of_int (a (n + 1))" using n_def2 unfolding dn_def pa_def . then show ?thesis apply (fold pa_def dn_def as_def) using \pa>0\ \q>0\ a[rule_format, of "Suc n"] b[rule_format, of "Suc n"] by (auto simp add:divide_simps algebra_simps) qed also have KKK: "... = (q* ((\j=1..(s-1). real_of_int( a j)))*2 * (b (n+1)))/ ((d (n+1))powr(2^(n+1)))" apply (fold as_def pa_def dn_def) apply simp using \0 < pa\ by blast also have KKKK: "... < 1" using n_def1 unfolding as_def by simp finally show ?thesis by auto qed lemma fixes d ::"nat\real " and a b ::"nat\int " and A::real assumes "A > 1" and d: "\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu3: "filterlim ( \ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows issummable: "summable (\j. b j / a j)" and show5: "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" proof- define c where "c = (\j. b j / a j)" have c_pos:"c j>0" for j using a b unfolding c_def by simp have c_ratio_tendsto:"(\n. c (n+1) / c n ) \ 0" proof - define nn where "nn=(\n. (2::int)^ (Suc n))" define ff where "ff=(\ n. (a n / a (Suc n)) powr(1 /nn n)*(d(Suc n)))" have nn_pos:"nn n>0" and ff_pos:"ff n>0" for n subgoal unfolding nn_def by simp subgoal unfolding ff_def using d[rule_format, of "Suc n"] a[rule_format, of n] a[rule_format, of "Suc n"] by auto done have ff_tendsto:"ff \ 1/ sqrt A" proof - have "(of_int (a n)) powr(1 / (nn n)) = sqrt(of_int (a n) powr(1 /of_int (2^n)))" for n unfolding nn_def using a by (simp add: powr_half_sqrt [symmetric] powr_powr ac_simps) moreover have "(( \ n. sqrt(of_int (a n) powr(1 /of_int (2^n))))\ sqrt A) sequentially " using assu1 tendsto_real_sqrt by blast ultimately have "(( \ n. (of_int (a n)) powr(1 /of_int (nn n)))\ sqrt A) sequentially " by auto from tendsto_divide[OF this assu1[THEN LIMSEQ_ignore_initial_segment[where k=1]]] have "(\n. (a n / a (Suc n)) powr(1 /nn n)) \ 1/sqrt A" using \A>1\ a unfolding nn_def - by (auto simp add:powr_divide inverse_eq_divide sqrt_divide_self_eq) + by (auto simp add:powr_divide less_imp_le inverse_eq_divide sqrt_divide_self_eq) moreover have "(\n. d (Suc n))\ 1" apply (rule convergent_prod_imp_LIMSEQ) using convergent_prod_iff_shift[of d 1] \convergent_prod d\ by auto ultimately show ?thesis unfolding ff_def by (auto intro:tendsto_eq_intros) qed have "(\n. (ff n) powr nn n) \ 0" proof - define aa where "aa=(1+1/sqrt A) / 2" have "eventually (\n. ff nA>1\ by (auto simp add:field_simps) moreover have "(\n. aa powr nn n) \ 0" proof - have "(\y. aa ^ (nat \ nn) y) \ 0" apply (rule tendsto_power_zero) subgoal unfolding nn_def comp_def apply (rule filterlim_subseq) by (auto intro:strict_monoI) subgoal unfolding aa_def using \A>1\ by auto done then show ?thesis proof (elim filterlim_mono_eventually) have "aa>0" unfolding aa_def using \A>1\ by (auto simp add:field_simps pos_add_strict) then show " \\<^sub>F x in sequentially. aa ^ (nat \ nn) x = aa powr real_of_int (nn x)" by (auto simp: powr_int order.strict_implies_order[OF nn_pos]) qed auto qed ultimately show ?thesis apply (elim metric_tendsto_imp_tendsto) apply (auto intro!:powr_mono2 elim!:eventually_mono) using nn_pos ff_pos by (meson le_cases not_le)+ qed then have "(\n. (d (Suc n))^(nat (nn n)) * (a n / a (Suc n))) \ 0" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in sequentially. ff x powr (nn x) = d (Suc x) ^ nat (nn x) * (a x / a (Suc x))" apply (rule eventuallyI) subgoal for x unfolding ff_def using a[rule_format,of x] a[rule_format,of "Suc x"] d[rule_format,of "Suc x"] nn_pos[of x] apply (auto simp add:field_simps powr_divide powr_powr powr_mult ) by (simp add: powr_int) done qed auto moreover have "(\n. b (Suc n) / (d (Suc n))^(nat (nn n))) \ 0" using tendsto_inverse_0_at_top[OF assu3,THEN LIMSEQ_ignore_initial_segment[where k=1]] unfolding nn_def by (auto simp add:field_simps nat_mult_distrib nat_power_eq) ultimately have "(\n. b (Suc n) * (a n / a (Suc n))) \ 0" apply - subgoal premises asm using tendsto_mult[OF asm,simplified] apply (elim filterlim_mono_eventually) using d by (auto simp add:algebra_simps,metis (mono_tags, lifting) always_eventually not_one_less_zero) done then have "(\n. (b (Suc n) / b n) * (a n / a (Suc n))) \ 0" apply (elim Lim_transform_bound[rotated]) apply (rule eventuallyI) subgoal for x using a[rule_format, of x] a[rule_format, of "Suc x"] b[rule_format, of x] b[rule_format, of "Suc x"] by (auto simp add:field_simps) done then show ?thesis unfolding c_def by (auto simp add:algebra_simps) qed from c_ratio_tendsto have "(\n. norm (b (Suc n) / a (Suc n)) / norm (b n / a n)) \ 0" unfolding c_def using a b by (force simp add:divide_simps abs_of_pos intro: Lim_transform_eventually) from summable_ratio_test_tendsto[OF _ _ this] a b show "summable c" unfolding c_def apply auto by (metis less_irrefl) have "\\<^sub>F n in sequentially. (\j. c (n + j)) \ 2 * c n" proof - obtain N where N_ratio:"\n\N. c (n + 1) / c n < 1 / 2" proof - have "eventually (\n. c (n+1) / c n < 1/2) sequentially" using c_ratio_tendsto[unfolded tendsto_iff,rule_format, of "1/2",simplified] apply eventually_elim subgoal for n using c_pos[of n] c_pos[of "Suc n"] by auto done then show ?thesis using that unfolding eventually_sequentially by auto qed have "(\j. c (j + n)) \ 2 * c n" when "n\N" for n proof - have "(\j 2*c n * (1 - 1 / 2^(m+1))" for m proof (induct m) case 0 then show ?case using c_pos[of n] by simp next case (Suc m) have "(\ji c n + (\i c (i + n) / 2" for i using N_ratio[rule_format,of "i+n"] \n\N\ c_pos[of "i+n"] by simp then show ?thesis by (auto intro:sum_mono) qed also have "... = c n + (\i c n + c n * (1 - 1 / 2 ^ (m + 1))" using Suc by auto also have "... = 2 * c n * (1 - 1 / 2 ^ (Suc m + 1))" by (auto simp add:field_simps) finally show ?case . qed then have "(\j 2*c n" for m using c_pos[of n] by (smt divide_le_eq_1_pos divide_pos_pos nonzero_mult_div_cancel_left zero_less_power) moreover have "summable (\j. c (j + n))" using \summable c\ by (simp add: summable_iff_shift) ultimately show ?thesis using suminf_le_const[of "\j. c (j+n)" "2*c n"] by auto qed then show ?thesis unfolding eventually_sequentially by (auto simp add:algebra_simps) qed then show "\\<^sub>F n in sequentially. (\j. (b (n + j)) / (a (n + j))) \ 2 * b n / a n" unfolding c_def by simp qed theorem Hancl3: fixes d ::"nat\real " and a b ::"nat\int " assumes "A > 1" and d:"\n. d n> 1" and a: "\n. a n>0" and b:"\n. b n>0" and "s>0" and assu1: "(( \ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and assu2: "\n \ s. (A/ (of_int (a n)) powr(1 /of_int (2^n)))> prodinf (\j. d(n +j))" and assu3: "filterlim (\ n. (d n)^(2^n)/ b n) at_top sequentially" and "convergent_prod d" shows "suminf(\ n. b n / a n ) \ Rats" proof(rule ccontr) assume asm:"\ (suminf(\ n. b n / a n ) \ Rats)" have ab_sum:"summable (\j. b j / a j)" using issummable[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] . obtain p q ::int where "q>0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p/q)" proof - from asm have "suminf(\ n. b n / a n ) \ Rats" by auto then have "suminf(\ n. b (n+1) / a (n+1)) \ Rats" apply (subst suminf_minus_initial_segment[OF ab_sum,of 1]) by auto then obtain p' q' ::int where "q'\0" and pq_def:"(\ n. b (n+1) / a (n+1) ) sums (p'/q')" unfolding Rats_eq_int_div_int using summable_ignore_initial_segment[OF ab_sum,of 1,THEN summable_sums] by force define p q where "p=(if q'<0 then - p' else p')" and "q=(if q'<0 then - q' else q')" have "p'/q'=p/q" "q>0" using \q'\0\ unfolding p_def q_def by auto then show ?thesis using that[of q p] pq_def by auto qed define ALPHA where "ALPHA = (\n. (of_int q)*((\j=1..n. of_int(a j)))*(suminf (\ (j::nat). (b (j+n+1)/a (j+n+1)))))" have "ALPHA n \ 1" for n proof - have "suminf(\ n. b (n+1) / a (n+1)) > 0" apply (rule suminf_pos) subgoal using summable_ignore_initial_segment[OF ab_sum,of 1] by auto subgoal using a b by simp done then have "p/q > 0" using sums_unique[OF pq_def,symmetric] by auto then have "q\1" "p\1" using \q>0\ by (auto simp add:divide_simps) moreover have "\n. 1 \ a n" "\n. 1 \ b n" using a b by (auto simp add: int_one_le_iff_zero_less) ultimately show ?thesis unfolding ALPHA_def using show7[OF _ _ _ _ pq_def] by auto qed moreover have "\n. ALPHA n < 1" unfolding ALPHA_def proof (rule lasttoshow[OF d a \s>0\ \q>0\ \A>1\ b _ assu3]) show "\\<^sub>F n in sequentially. (\j. real_of_int (b (n + j)) / real_of_int (a (n + j))) \ real_of_int (2 * b n) / real_of_int (a n)" using show5[OF \A>1\ d a b assu1 assu3 \convergent_prod d\] by simp show "\m\s. \n\m. d (n + 1) * Max ((\j. real_of_int (a j) powr (1 / real_of_int (2 ^ j))) ` {s..n}) \ real_of_int (a (n + 1)) powr (1 / real_of_int (2 ^ (n + 1)))" apply (rule show9[OF \A>1\ d a \s>0\ assu1 \convergent_prod d\]) using show8[OF \A>1\ d a \s>0\ \convergent_prod d\ assu2] by (simp add:algebra_simps) qed ultimately show False using not_le by blast qed corollary Hancl3corollary: fixes A::real and a b ::"nat\int " assumes "A > 1" and a: "\n. a n>0" and b:"\n. b n>0" and assu1: "((\ n. (of_int (a n)) powr(1 /of_int (2^n)))\ A) sequentially " and asscor2: "\n \6. (of_int (a n)) powr(1 /of_int (2^n ))*(1+ 4*(2/3)^n) \ A \ (b n \2 powr((4/3)^(n-1)) ) " shows "suminf(\ n. b n / a n ) \ Rats" proof- define d::"nat\real" where "d= (\ n. 1+(2/3)^(n+1))" have dgt1:"\n. 1 < d n " unfolding d_def by auto moreover have "convergent_prod d" unfolding d_def apply (rule abs_convergent_prod_imp_convergent_prod) apply (rule summable_imp_abs_convergent_prod) using summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1] by simp moreover have "\n\6. (\j. d (n + j)) < A / real_of_int (a n) powr (1 / real_of_int (2 ^ n))" proof rule+ fix n::nat assume "6 \ n" have d_sum:"summable (\j. ln (d j))" unfolding d_def apply (rule summable_ln_plus) apply (rule summable_ignore_initial_segment[OF complete_algebra_summable_geometric [of "2/3::real",simplified],of 1]) by simp have "(\j. ln (d (n + j))) < ln (1+4 * (2 / 3) ^ n)" proof - define c::real where "c=(2 / 3) ^ n" have "0n\6\ apply (subst power_add[symmetric]) by auto also have "... \ (2 / 3)^6" by (auto intro:power_le_one) also have "... < 1/8" by (auto simp add:field_simps) finally show "c<1/8" . qed (simp add:c_def) have "(\j. ln (d (n + j))) \ (\j. (2 / 3) ^ (n + j + 1))" apply (rule suminf_le) subgoal unfolding d_def apply (intro allI ln_add_one_self_le_self2 ) apply (rule order.strict_trans[of _ 0]) by auto subgoal using summable_ignore_initial_segment[OF d_sum,of n] by (auto simp add:algebra_simps) subgoal using summable_geometric[THEN summable_ignore_initial_segment,of "2/3" "n+1"] by (auto simp add:algebra_simps) done also have "... = (\j. (2 / 3)^(n+1)*(2 / 3) ^ j)" by (auto simp add:algebra_simps power_add) also have "... = (2 / 3)^(n+1) * (\j. (2 / 3) ^ j)" apply (rule suminf_mult) by (auto intro:summable_geometric) also have "... = 2 * c" unfolding c_def apply (subst suminf_geometric) by auto also have "... < 4 * c - (4 * c)\<^sup>2" using \0 \c<1/8\ by (sos "((((A<0 * A<1) * R<1) + ((A<=0 * R<1) * (R<1/16 * [1]^2))))") also have "... \ ln (1 + 4 * c)" apply (rule ln_one_plus_pos_lower_bound) using \0 \c<1/8\ by auto finally show ?thesis unfolding c_def by simp qed then have "exp (\j. ln (d (n + j))) < 1 + 4 * (2 / 3) ^ n" by (metis Groups.mult_ac(2) add.right_neutral add_mono_thms_linordered_field(5) divide_inverse divide_less_eq_numeral1(1) divide_pos_pos exp_gt_zero less_eq_real_def ln_exp ln_less_cancel_iff mult_zero_left rel_simps(27) rel_simps(76) zero_less_one zero_less_power) moreover have "exp (\j. ln (d (n + j))) = (\j. d (n + j))" proof (subst exp_suminf_prodinf_real [symmetric]) show "\k. 0 \ ln (d (n + k))" using dgt1 by (simp add: less_imp_le) show "abs_convergent_prod (\na. exp (ln (d (n + na))))" apply (subst exp_ln) subgoal for j using dgt1[rule_format,of "n+j"] by auto subgoal unfolding abs_convergent_prod_def real_norm_def apply (subst abs_of_nonneg) using convergent_prod_iff_shift[of d n] \convergent_prod d\ by (auto simp add: dgt1 less_imp_le algebra_simps) done show "(\na. exp (ln (d (n + na)))) = (\j. d (n + j))" apply (subst exp_ln) subgoal using dgt1 le_less_trans zero_le_one by blast subgoal by simp done qed ultimately have "(\j. d (n + j)) < 1 + 4 * (2 / 3) ^ n" by simp also have "... \ A / (a n) powr (1 / of_int (2 ^ n))" proof - have "a n powr (1 / real_of_int (2 ^ n)) > 0" using a[rule_format,of n] by auto then show ?thesis using asscor2[rule_format,OF \6\n\] by (auto simp add:field_simps) qed finally show "(\j. d (n + j)) < A / real_of_int (a n) powr (1 / of_int (2 ^ n))" . qed moreover have "LIM n sequentially. d n ^ 2 ^ n / real_of_int (b n) :> at_top" proof - have "LIM n sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) :> at_top" proof - define n1 where "n1=(\n. (2::real)* (3/2)^(n-1))" define n2 where "n2=(\n. ((4::real)/3)^(n-1))" have "LIM n sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) :> at_top" proof (rule filterlim_at_top_powr_real[where g'="exp (8/9) / (2::real)"]) define e1 where "e1=exp (8/9) / (2::real)" show "e1>1" unfolding e1_def by (approximation 4) show "(\n. ((1+(8/9)/(n1 n)) powr (n1 n))/2) \ e1" proof - have "(\n. (1+(8/9)/(n1 n)) powr (n1 n)) \ exp (8/9)" apply (rule filterlim_compose[OF tendsto_exp_limit_at_top]) unfolding n1_def by (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_realpow_sequentially_at_top simp:filterlim_sequentially_iff[of "\x. (3 / 2) ^ (x - Suc 0)" _ 1]) then show ?thesis unfolding e1_def by (intro tendsto_intros,auto) qed show "filterlim n2 at_top sequentially" unfolding n2_def apply (subst filterlim_sequentially_iff[of "\n. (4 / 3) ^ (n - 1)" _ 1]) by (auto intro:filterlim_realpow_sequentially_at_top) qed moreover have "\\<^sub>F n in sequentially. (((1+(8/9)/(n1 n)) powr (n1 n))/2) powr (n2 n) = d n ^ 2 ^ n / 2 powr((4/3)^(n-1))" proof (rule eventually_sequentiallyI[of 1]) fix x assume "x\(1::nat)" have " ((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = (((1 + 8 / 9 / n1 x) powr n1 x) powr n2 x) / 2 powr (4 / 3) ^ (x - 1)" - apply (subst powr_divide) - apply (simp_all add:n1_def n2_def) - by (smt divide_nonneg_nonneg zero_le_power) + by (simp add:n1_def n2_def powr_divide) also have "... = (1 + 8 / 9 / n1 x) powr (n1 x * n2 x) / 2 powr (4 / 3) ^ (x - 1)" - apply (subst powr_powr) - by simp + by (simp add: powr_powr) also have "... = (1 + 8 / 9 / n1 x) powr (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" proof - have "n1 x * n2 x = 2 ^ x" unfolding n1_def n2_def apply (subst mult.assoc) apply (subst power_mult_distrib[symmetric]) - using \x\1\ by (auto simp add:power_Suc[symmetric] simp del:power_Suc) + using \x\1\ by (auto simp flip:power_Suc) then show ?thesis by simp qed also have "... = (1 + 8 / 9 / n1 x) ^ (2 ^ x) / 2 powr (4 / 3) ^ (x - 1)" apply (subst (3) powr_realpow[symmetric]) apply (simp_all add: n1_def) by (smt zero_le_divide_iff zero_le_power) also have "... = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1)" proof - define x1 where "x1=x-1" have *:"x=x1+1" unfolding x1_def using \x\1\ by simp have **: "8 / 9 / n1 x = (2 / 3) ^ (x + 1)" unfolding n1_def using \x\1\ apply (fold x1_def *[symmetric]) by (auto simp add:divide_simps) then show ?thesis unfolding d_def apply (subst **) by auto qed finally show "((1 + 8 / 9 / n1 x) powr n1 x / 2) powr n2 x = d x ^ 2 ^ x / 2 powr (4 / 3) ^ (x - 1) " . qed ultimately show ?thesis using filterlim_cong by fast qed moreover have "\\<^sub>F n in sequentially. d n ^ 2 ^ n / 2 powr((4/3)^(n-1)) \ d n ^ 2 ^ n / real_of_int (b n)" apply (rule eventually_sequentiallyI[of 6]) apply (rule divide_left_mono) subgoal for x using asscor2[rule_format,of x] by auto subgoal for x using \\n. 1 < d n\[rule_format, of x] by auto subgoal for x using b by auto done ultimately show ?thesis by (auto elim: filterlim_at_top_mono) qed ultimately show ?thesis using Hancl3[OF \A>1\ _ a b _ assu1,of d 6] by force qed end diff --git a/thys/QHLProver/Matrix_Limit.thy b/thys/QHLProver/Matrix_Limit.thy --- a/thys/QHLProver/Matrix_Limit.thy +++ b/thys/QHLProver/Matrix_Limit.thy @@ -1,1769 +1,1769 @@ section \Matrix limits\ theory Matrix_Limit imports Complex_Matrix begin subsection \Definition of limit of matrices\ definition limit_mat :: "(nat \ complex mat) \ complex mat \ nat \ bool" where "limit_mat X A m \ (\ n. X n \ carrier_mat m m \ A \ carrier_mat m m \ (\ i < m. \ j < m. (\ n. (X n) $$ (i, j)) \ (A $$ (i, j))))" lemma limit_mat_unique: assumes limA: "limit_mat X A m" and limB: "limit_mat X B m" shows "A = B" proof - have dim: "A \ carrier_mat m m" "B \ carrier_mat m m" using limA limB limit_mat_def by auto { fix i j assume i: "i < m" and j: "j < m" have "(\ n. (X n) $$ (i, j)) \ (A $$ (i, j))" using limit_mat_def limA i j by auto moreover have "(\ n. (X n) $$ (i, j)) \ (B $$ (i, j))" using limit_mat_def limB i j by auto ultimately have "(A $$ (i, j)) = (B $$ (i, j))" using LIMSEQ_unique by auto } then show "A = B" using mat_eq_iff dim by auto qed lemma limit_mat_const: fixes A :: "complex mat" assumes "A \ carrier_mat m m" shows "limit_mat (\k. A) A m" unfolding limit_mat_def using assms by auto lemma limit_mat_scale: fixes X :: "nat \ complex mat" and A :: "complex mat" assumes limX: "limit_mat X A m" shows "limit_mat (\n. c \\<^sub>m X n) (c \\<^sub>m A) m" proof - have dimA: "A \ carrier_mat m m" using limX limit_mat_def by auto have dimX: "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have "\i j. i < m \ j < m \ (\n. (c \\<^sub>m X n) $$ (i, j)) \ (c \\<^sub>m A) $$ (i, j)" proof - fix i j assume i: "i < m" and j: "j < m" have "(\n. (X n) $$ (i, j)) \ A$$(i, j)" using limX limit_mat_def i j by auto moreover have "(\n. c) \ c" by auto ultimately have "(\n. c * (X n) $$ (i, j)) \ c * A$$(i, j)" using tendsto_mult[of "\n. c" c] limX limit_mat_def by auto moreover have "(c \\<^sub>m X n) $$ (i, j) = c * (X n) $$ (i, j)" for n using index_smult_mat(1)[of i "X n" j c] i j dimX[of n] by auto moreover have "(c \\<^sub>m A) $$ (i, j) = c * A $$ (i, j)" using index_smult_mat(1)[of i "A" j c] i j dimA by auto ultimately show "(\n. (c \\<^sub>m X n) $$ (i, j)) \ (c \\<^sub>m A) $$ (i, j)" by auto qed then show ?thesis unfolding limit_mat_def using dimA dimX by auto qed lemma limit_mat_add: fixes X :: "nat \ complex mat" and Y :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes limX: "limit_mat X A m" and limY: "limit_mat Y B m" shows "limit_mat (\k. X k + Y k) (A + B) m" proof - have dimA: "A \ carrier_mat m m" using limX limit_mat_def by auto have dimB: "B \ carrier_mat m m" using limY limit_mat_def by auto have dimX: "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have dimY: "\n. Y n \ carrier_mat m m" using limY unfolding limit_mat_def by auto then have dimXAB: "\n. X n + Y n \ carrier_mat m m \ A + B \ carrier_mat m m" using dimA dimB dimX dimY by (simp) have "(\i j. i < m \ j < m \ (\n. (X n + Y n) $$ (i, j)) \ (A + B) $$ (i, j))" proof - fix i j assume i: "i < m" and j: "j < m" have "(\n. (X n) $$ (i, j)) \ A$$(i, j)" using limX limit_mat_def i j by auto moreover have "(\n. (Y n) $$ (i, j)) \ B$$(i, j)" using limY limit_mat_def i j by auto ultimately have "(\n. (X n)$$(i, j) + (Y n) $$ (i, j)) \ (A$$(i, j) + B$$(i, j))" using tendsto_add[of "\n. (X n) $$ (i, j)" "A $$ (i, j)"] by auto moreover have "(X n + Y n) $$ (i, j) = (X n)$$(i, j) + (Y n) $$ (i, j)" for n using i j dimX dimY index_add_mat(1)[of i "Y n" j "X n"] by fastforce moreover have "(A + B) $$ (i, j) = A$$(i, j) + B$$(i, j)" using i j dimA dimB by fastforce ultimately show "(\n. (X n + Y n) $$ (i, j)) \ (A + B) $$ (i, j)" by auto qed then show ?thesis unfolding limit_mat_def using dimXAB by auto qed lemma limit_mat_minus: fixes X :: "nat \ complex mat" and Y :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes limX: "limit_mat X A m" and limY: "limit_mat Y B m" shows "limit_mat (\k. X k - Y k) (A - B) m" proof - have dimA: "A \ carrier_mat m m" using limX limit_mat_def by auto have dimB: "B \ carrier_mat m m" using limY limit_mat_def by auto have dimX: "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have dimY: "\n. Y n \ carrier_mat m m" using limY unfolding limit_mat_def by auto have "-1 \\<^sub>m Y n = - Y n" for n using dimY by auto moreover have "-1 \\<^sub>m B = - B" using dimB by auto ultimately have "limit_mat (\n. - Y n) (- B) m" using limit_mat_scale[OF limY, of "-1"] by auto then have "limit_mat (\n. X n + (- Y n)) (A + (- B)) m" using limit_mat_add limX by auto moreover have "X n + (- Y n) = X n - Y n" for n using dimX dimY by auto moreover have "A + (- B) = A - B" by auto ultimately show ?thesis by auto qed lemma limit_mat_mult: fixes X :: "nat \ complex mat" and Y :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes limX: "limit_mat X A m" and limY: "limit_mat Y B m" shows "limit_mat (\k. X k * Y k) (A * B) m" proof - have dimA: "A \ carrier_mat m m" using limX limit_mat_def by auto have dimB: "B \ carrier_mat m m" using limY limit_mat_def by auto have dimX: "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have dimY: "\n. Y n \ carrier_mat m m" using limY unfolding limit_mat_def by auto then have dimXAB: "\n. X n * Y n \ carrier_mat m m \ A * B \ carrier_mat m m" using dimA dimB dimX dimY by fastforce have "(\i j. i < m \ j < m \ (\n. (X n * Y n) $$ (i, j)) \ (A * B) $$ (i, j))" proof - fix i j assume i: "i < m" and j: "j < m" have eqn: "(X n * Y n) $$ (i, j) = (\k=0..k=0..n. (X n) $$ (i, k)) \ A$$(i, k)" if "k < m" for k using limX limit_mat_def that i by auto moreover have "(\n. (Y n) $$ (k, j)) \ B$$(k, j)" if "k < m" for k using limY limit_mat_def that j by auto ultimately have "(\n. (X n)$$(i, k) * (Y n)$$(k,j)) \ A$$(i, k) * B$$(k, j)" if "k < m" for k using tendsto_mult[of "\n. (X n) $$ (i, k)" "A$$(i, k)" _ "\n. (Y n)$$(k, j)" "B$$(k, j)"] that by auto then have "(\n. (\k=0.. (\k=0..k n. (X n)$$(i,k) * (Y n)$$(k,j)" "\k. A$$(i, k) * B$$(k, j)"] by auto then show "(\n. (X n * Y n) $$ (i, j)) \ (A * B) $$ (i, j)" using eqn eq by auto qed then show ?thesis unfolding limit_mat_def using dimXAB by fastforce qed text \Adding matrix A to the sequence X\ definition mat_add_seq :: "complex mat \ (nat \ complex mat) \ nat \ complex mat" where "mat_add_seq A X = (\n. A + X n)" lemma mat_add_limit: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes dimB: "B \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (mat_add_seq B X) (B + A) m" unfolding mat_add_seq_def using limit_mat_add limit_mat_const[OF dimB] limX by auto lemma mat_minus_limit: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes dimB: "B \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (\n. B - X n) (B - A) m" using limit_mat_minus limit_mat_const[OF dimB] limX by auto text \Multiply matrix A by the sequence X\ definition mat_mult_seq :: "complex mat \ (nat \ complex mat) \ nat \ complex mat" where "mat_mult_seq A X = (\n. A * X n)" lemma mat_mult_limit: fixes X :: "nat \ complex mat" and A B :: "complex mat" and m :: nat assumes dimB: "B \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (mat_mult_seq B X) (B * A) m" unfolding mat_mult_seq_def using limit_mat_mult limit_mat_const[OF dimB] limX by auto lemma mult_mat_limit: fixes X :: "nat \ complex mat" and A B :: "complex mat" and m :: nat assumes dimB: "B \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (\k. X k * B) (A * B) m" unfolding mat_mult_seq_def using limit_mat_mult limit_mat_const[OF dimB] limX by auto lemma quadratic_form_mat: fixes A :: "complex mat" and v :: "complex vec" and m :: nat assumes dimv: "dim_vec v = m" and dimA: "A \ carrier_mat m m" shows "inner_prod v (A *\<^sub>v v) = (\i=0..j=0..v v) = (\i=0..j=0.. nat \'a::ab_group_add" shows "(\x\A. \y\B. h x y - g x y) = (\x\A. \y\B. h x y) - (\x\A. \y\B. g x y)" proof - have "\ x \ A. (\y\B. h x y - g x y) = (\y\B. h x y) - (\y\B. g x y)" proof - { fix x assume x: "x \ A" have "(\y\B. h x y - g x y) = (\y\B. h x y) - (\y\B. g x y)" using sum_subtractf by auto } then show ?thesis using sum_subtractf by blast qed then have "(\x\A.\y\B. h x y - g x y) = (\x\A. ((\y\B. h x y) - (\y\B. g x y)))" by auto also have "\ = (\x\A. \y\B. h x y) - (\x\A. \y\B. g x y)" by (simp add: sum_subtractf) finally have " (\x\A. \y\B. h x y - g x y) = (\x\A. sum (h x) B) - (\x\A. sum (g x) B)" by auto then show ?thesis by auto qed lemma sum_abs_complex: fixes h :: "nat \ nat \ complex" shows "cmod (\x\A.\y\B. h x y) \ (\x\A. \y\B. cmod(h x y))" proof - have B: "\ x \ A. cmod( \y\B .h x y) \ (\y\B. cmod(h x y))" using sum_abs norm_sum by blast have "cmod (\x\A.\y\B. h x y) \ (\x\A. cmod( \y\B .h x y))" using sum_abs norm_sum by blast also have "\ \ (\x\A. \y\B. cmod(h x y))" using sum_abs norm_sum B by (simp add: sum_mono) finally have "cmod (\x\A. \y\B. h x y) \ (\x\A. \y\B. cmod (h x y))" by auto then show ?thesis by auto qed lemma hermitian_mat_lim_is_hermitian: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat assumes limX: "limit_mat X A m" and herX: "\ n. hermitian (X n)" shows "hermitian A" proof - have dimX: "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have dimA : "A \ carrier_mat m m" using limX unfolding limit_mat_def by auto from herX have herXn: "\ n. adjoint (X n) = (X n)" unfolding hermitian_def by auto from limX have limXn: "\ijn. X n $$ (i, j)) \ A $$ (i, j)" unfolding limit_mat_def by auto have "\ijn. X n $$ (i, j)) \ A $$ (i, j)" using limXn i j by auto have ji: "(\n. X n $$ (j, i)) \ A $$ (j, i)" using limXn i j by auto then have "\r>0. \no. \n\no. dist (conjugate (X n $$ (j, i))) (conjugate (A $$ (j, i))) < r" proof - { fix r :: real assume r : "r > 0" have "\no. \n\no. cmod (X n $$ (j, i) - A $$ (j, i)) < r" using ji r unfolding LIMSEQ_def dist_norm by auto then obtain no where Xji: "\n\no. cmod (X n $$ (j, i) - A $$ (j, i)) < r" by auto then have "\n\no. cmod (conjugate (X n $$ (j, i) - A $$ (j, i))) < r" using complex_mod_cnj conjugate_complex_def by presburger then have "\n\no. dist (conjugate (X n $$ (j, i))) (conjugate (A $$ (j, i))) < r" unfolding dist_norm by auto then have "\no. \n\no. dist (conjugate (X n $$ (j, i))) (conjugate (A $$ (j, i))) < r" by auto } then show ?thesis by auto qed then have conjX: "(\n. conjugate (X n $$ (j, i))) \ conjugate (A $$ (j, i))" unfolding LIMSEQ_def by auto from herXn have "\ n. conjugate (X n $$ (j,i)) = X n$$ (i, j)" using adjoint_eval i j dimX by (metis adjoint_dim_col carrier_matD(1)) then have "(\n. X n $$ (i, j)) \ conjugate (A $$ (j, i))" using conjX by auto then have "conjugate (A $$ (j,i)) = A$$ (i, j)" using ij by (simp add: LIMSEQ_unique) then have "(adjoint A)$$ (i, j) = A$$ (i, j)" using adjoint_eval i j by (simp add:aij) } then show ?thesis by auto qed then have "hermitian A" using hermitian_def dimA by (metis adjoint_dim carrier_matD(1) carrier_matD(2) eq_matI) then show ?thesis by auto qed lemma quantifier_change_order_once: fixes P :: "nat \ nat \ bool" and m :: nat shows "\jno. \n\no. P n j \ \no. \jn\no. P n j" proof (induct m) case 0 then show ?case by auto next case (Suc m) then show ?case proof - have mm: "\no. \jn\no. P n j" using Suc by auto then obtain M where MM: "\jn\M. P n j" by auto have sucm: "\no. \n\no. P n m" using Suc(2) by auto then obtain N where NN: "\n\N. P n m" by auto let ?N = "max M N" from MM NN have "\jn\?N. P n j" by (metis less_antisym max.boundedE) then have "\no. \jn\no. P n j" by blast then show ?thesis by auto qed qed lemma quantifier_change_order_twice: fixes P :: "nat \ nat \ nat \ bool" and m n :: nat shows "\ij no. \n\no. P n i j \ \no. \ijn\no. P n i j" proof - assume fact: "\ij no. \n\no. P n i j" have one: "\ino.\jn\no. P n i j" using fact quantifier_change_order_once by auto have two: "\ino.\jn\no. P n i j \ \no. \ijn\no. P n i j" proof (induct m) case 0 then show ?case by auto next case (Suc m) then show ?case proof - obtain M where MM: "\ijn\M. P n i j" using Suc by auto obtain N where NN: "\jn\N. P n m j" using Suc(2) by blast let ?N = "max M N" from MM NN have "\ijn\?N. P n i j" by (metis less_antisym max.boundedE) then have "\no. \ijn\no. P n i j" by blast then show ?thesis by auto qed qed with fact show ?thesis using one by auto qed lemma pos_mat_lim_is_pos: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat assumes limX: "limit_mat X A m" and posX: "\n. positive (X n)" shows "positive A" proof (rule ccontr) have dimX : "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto have dimA : "A \ carrier_mat m m" using limX unfolding limit_mat_def by auto have herX : "\ n. hermitian (X n)" using posX positive_is_hermitian by auto then have herA : "hermitian A" using hermitian_mat_lim_is_hermitian limX by auto then have herprod: "\ v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) \ Reals" using hermitian_inner_prod_real dimA by auto assume npA: " \ positive A" from npA have "\ (A \ carrier_mat (dim_col A) (dim_col A)) \ \ (\v. dim_vec v = dim_col A \ 0 \ inner_prod v (A *\<^sub>v v))" unfolding positive_def by blast then have evA: "\ v. dim_vec v = dim_col A \ \ inner_prod v (A *\<^sub>v v) \ 0" using dimA by blast then have "\ v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) < 0" proof - obtain v where vA: "dim_vec v = dim_col A \ \ inner_prod v (A *\<^sub>v v) \ 0" using evA by auto from vA herprod have "\ 0 \ inner_prod v (A *\<^sub>v v) \ inner_prod v (A *\<^sub>v v) \ Reals" by auto then have "inner_prod v (A *\<^sub>v v) < 0" using complex_is_Real_iff by auto then have "\ v. dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) < 0" using vA by auto then show ?thesis by auto qed then obtain v where neg: "dim_vec v = dim_col A \ inner_prod v (A *\<^sub>v v) < 0" by auto have nzero: "v \ 0\<^sub>v m" proof (rule ccontr) assume nega: " \ v \ 0\<^sub>v m" have zero: "v = 0\<^sub>v m" using nega by auto have "(A *\<^sub>v v) = 0\<^sub>v m" unfolding mult_mat_vec_def using zero using dimA by auto then have zerov: "inner_prod v (A *\<^sub>v v) = 0" by (simp add: zero) from neg zerov have "\ v \ 0\<^sub>v m \ False" using dimA by auto with nega show False by auto qed have invgeq: "inner_prod v v > 0" proof - have "inner_prod v v = vec_norm v * vec_norm v" unfolding vec_norm_def by (metis carrier_matD(2) carrier_vec_dim_vec dimA mult_cancel_left1 neg normalized_cscalar_prod normalized_vec_norm nzero vec_norm_def) moreover have "vec_norm v > 0" using nzero vec_norm_ge_0 neg dimA by (metis carrier_matD(2) carrier_vec_dim_vec) ultimately have "inner_prod v v > 0" by auto then show ?thesis by auto qed have invv: "inner_prod v v = (\i = 0.. i < m. conjugate (v $ i) * (v $ i) \ 0" using conjugate_square_smaller_0 by simp then have vi: "\ i < m. conjugate (v $ i) * (v $ i) = cmod (conjugate (v $ i) * (v $ i))" using cmod_eq_Re by (simp add: complex.expand) have "inner_prod v v= (\i = 0.. = (\i = 0.. = (\i = 0..i = 0..v v) = (\i=0..j=0..ijn. X n $$ (i, j)) \ A $$ (i, j)" unfolding limit_mat_def by auto then have limXv: "(\ n. inner_prod v ((X n) *\<^sub>v v)) \ inner_prod v (A *\<^sub>v v)" proof - have XAless: "cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) \ (\i = 0..j = 0.. i < m. \ j < m. conjugate (v$i) * X n $$(i, j) * v$j - conjugate (v$i) * A$$(i, j) * v$j = conjugate (v$i) * (X n $$(i, j)-A$$(i, j)) * v$j" by (simp add: mult.commute right_diff_distrib) then have ele: "\ i < m.(\j=0..j=0.. i < m. \ j < m. cmod(conjugate (v $ i) * (X n $$ (i, j) - A $$ (i, j)) * v $ j) = cmod(conjugate (v $ i)) * cmod (X n $$ (i, j) - A $$ (i, j)) * cmod(v $ j)" by (simp add: norm_mult) then have less: "\ i < m.(\j = 0..j = 0..v v) - inner_prod v (A *\<^sub>v v) = (\i=0..j=0..i=0..j=0.. = (\i=0..j=0.. i j. conjugate (v $ i) * X n $$ (i, j) * v $ j" "\ i j. conjugate (v $ i) * A $$ (i, j) * v $ j" "{0.. = (\i=0..j=0..v v) - inner_prod v (A *\<^sub>v v) = (\i = 0..j = 0..v v) - inner_prod v (A *\<^sub>v v)) = cmod (\i = 0..j = 0.. \ (\i = 0..j = 0.. = (\i = 0..j = 0..ijr>0. \no. \n\no. cmod (X n $$ (i, j) - A $$ (i, j)) < r" unfolding LIMSEQ_def dist_norm by auto from limX have mg: "m > 0" using limit_mat_def by (metis carrier_matD(1) carrier_matD(2) mat_eq_iff neq0_conv not_less0 npA posX) have cmoda: "\no. \n\no. (\i = 0..j = 0.. 0" for r proof - let ?u = "(\i = 0..j = 0.. 0" proof - have ur: "?u = (\i = 0..j = 0..j = 0.. cmod (v $ i)" if i: "i < m" for i using member_le_sum[of i "{0.. j. cmod (v$j)"] cmod_def i by simp then have "\ i < m. (cmod (conjugate (v $ i)) * (\j = 0.. (cmod (conjugate (v $ i)) * cmod (v $ i))" by (simp add: mult_left_mono) then have "?u \ (\i = 0.. i. cmod (conjugate (v $ i)) * cmod (v $ i)" "\ i. cmod (conjugate (v $ i)) * (\j = 0..i = 0..i = 0..i = 0.. inner_prod v v" by (metis (no_types, lifting) Im_complex_of_real Re_complex_of_real invv less_eq_complex_def norm_mult sum.cong) then have "?u > 0" using invgeq by auto then show ?thesis by auto qed let ?s = "r / (2 * ?u)" have sgz: "?s > 0" using ug rl by (smt divide_pos_pos dual_order.strict_iff_order linordered_semiring_strict_class.mult_pos_pos zero_less_norm_iff r) from limijm have sij: "\no. \n\no. cmod (X n $$ (i, j) - A $$ (i, j)) < ?s" if i: "i < m" and j: "j < m" for i j proof - obtain N where Ns: "\n\N. cmod (X n $$ (i, j) - A $$ (i, j)) < ?s" using sgz limijm i j by blast then show ?thesis by auto qed then have "\no. \ijn\no. cmod (X n $$ (i, j) - A $$ (i, j)) < ?s" using quantifier_change_order_twice[of m m "\ n i j. (cmod (X n $$ (i, j) - A $$ (i, j))ijn\N. cmod (X n $$ (i, j) - A $$ (i, j)) < ?s" by auto then have mmN: "cmod (conjugate (v $ i)) * cmod (X n $$ (i, j) - A $$ (i, j)) * cmod (v $ j) \ ?s * (cmod (conjugate (v $ i)) * cmod (v $ j))" if i: "i < m" and j: "j < m" and n: "n \ N" for i j n proof - have geq: "cmod (conjugate (v $ i)) \ 0 \ cmod (v $ j)\0" by simp then have "cmod (conjugate (v $ i)) * cmod (X n $$ (i, j) - A $$ (i, j)) \cmod (conjugate (v $ i)) * ?s" using Nno i j n by (smt mult_left_mono) then have "cmod (conjugate (v $ i)) * cmod (X n $$ (i, j) - A $$ (i, j)) * cmod (v $ j) \ cmod (conjugate (v $ i)) *?s * cmod (v $ j)" using geq mult_right_mono by blast also have "\ = ?s * (cmod (conjugate (v $ i)) * cmod (v $ j))" by simp finally show ?thesis by auto qed then have "(\i = 0..j = 0.. N" for n proof - have mmX: "\ij ?s * (cmod (conjugate (v $ i)) * cmod (v $ j))" using n mmN by blast have "(\j = 0.. (\j = 0..j ?s * (cmod (conjugate (v $ i)) * cmod (v $ j))" using mmX i by auto then show ?thesis using sum_mono[of "{0.. j. cmod (conjugate (v $ i)) * cmod (X n $$ (i, j) - A $$ (i, j)) * cmod (v $ j)" "\ j. (?s * (cmod (conjugate (v $ i)) * cmod (v $ j)))"] atLeastLessThan_iff by blast qed then have "(\i = 0..j = 0.. (\i = 0..j = 0.. = ?s * (\i = 0..j = 0.. = r / 2" using nonzero_mult_divide_mult_cancel_right sgz by fastforce finally show ?thesis using r by auto qed then show ?thesis by auto qed then have XnAv:"\no. \n\no. cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < r" if r: "r > 0" for r proof - obtain no where nno: "\n\no. (\i = 0..j = 0..n\no. cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < r" using XAless neg by smt then show ?thesis by auto qed then have "(\n. inner_prod v (X n *\<^sub>v v)) \ inner_prod v (A *\<^sub>v v)" unfolding LIMSEQ_def dist_norm by auto then show ?thesis by auto qed from limXv have "\r>0. \no. \n\no. cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < r" unfolding LIMSEQ_def dist_norm by auto then have "\no. \n\no. cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < -?r" using rl by auto then obtain N where Ng: "\n\N. cmod (inner_prod v (X n *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < -?r" by auto then have XN: "cmod (inner_prod v (X N *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) < -?r" by auto from posX have "positive (X N)" by auto then have XNv:"inner_prod v (X N *\<^sub>v v) \ 0" by (metis Complex_Matrix.positive_def carrier_matD(2) dimA dimX neg) from rl XNv have XX: "cmod (inner_prod v (X N *\<^sub>v v) - inner_prod v (A *\<^sub>v v)) = cmod(inner_prod v (X N *\<^sub>v v)) - cmod(inner_prod v (A *\<^sub>v v))" using XN cmod_eq_Re by auto then have YY: "cmod(inner_prod v (X N *\<^sub>v v)) - cmod(inner_prod v (A *\<^sub>v v)) < -?r" using XN by auto then have "cmod(inner_prod v (X N *\<^sub>v v)) - cmod(inner_prod v (A *\<^sub>v v)) < cmod(inner_prod v (A *\<^sub>v v))" using rl cmod_eq_Re by auto then have "cmod(inner_prod v (X N *\<^sub>v v)) < 0" using XNv XX YY cmod_eq_Re by auto then have "False" using XNv by simp with npA show False by auto qed lemma limit_mat_ignore_initial_segment: "limit_mat g A d \ limit_mat (\n. g (n + k)) A d" proof - assume asm: "limit_mat g A d" then have lim: "\ i < d. \ j < d. (\ n. (g n) $$ (i, j)) \ (A $$ (i, j))" using limit_mat_def by auto then have limk: "\ i < d. \ j < d. (\ n. (g (n + k)) $$ (i, j)) \ (A $$ (i, j))" proof - { fix i j assume dims: "i < d" "j < d" then have "(\ n. (g n) $$ (i, j)) \ (A $$ (i, j))" using lim by auto then have "(\ n. (g (n + k)) $$ (i, j)) \ (A $$ (i, j))" using LIMSEQ_ignore_initial_segment by auto } then show "\ i < d. \ j < d. (\ n. (g (n + k)) $$ (i, j)) \ (A $$ (i, j))" by auto qed have "\ n. g n \ carrier_mat d d" using asm unfolding limit_mat_def by auto then have "\ n. g (n + k) \ carrier_mat d d" by auto moreover have "A \ carrier_mat d d" using asm limit_mat_def by auto ultimately show "limit_mat (\n. g (n + k)) A d" using limit_mat_def limk by auto qed lemma mat_trace_limit: "limit_mat g A d \ (\n. trace (g n)) \ trace A" proof - assume lim: "limit_mat g A d" then have dgn: "g n \ carrier_mat d d" for n using limit_mat_def by auto from lim have dA: "A \ carrier_mat d d" using limit_mat_def by auto have trg: "trace (g n) = (\k=0..k < d. (\n. (g n)$$(k, k)) \ A$$(k, k)" using limit_mat_def lim by auto then have "(\n. (\k=0.. (\k=0..n. trace (g n)) \ trace A" unfolding trace_def using trg carrier_matD[OF dgn] carrier_matD[OF dA] by auto qed subsection \Existence of least upper bound for the L\"{o}wner order\ definition lowner_is_lub :: "(nat \ complex mat) \ complex mat \ bool" where "lowner_is_lub f M \ (\n. f n \\<^sub>L M) \ (\M'. (\n. f n \\<^sub>L M') \ M \\<^sub>L M')" locale matrix_seq = fixes dim :: nat and f :: "nat \ complex mat" assumes dim: "\n. f n \ carrier_mat dim dim" and pdo: "\n. partial_density_operator (f n)" and inc: "\n. lowner_le (f n) (f (Suc n))" begin definition lowner_is_lub :: "complex mat \ bool" where "lowner_is_lub M \ (\n. f n \\<^sub>L M) \ (\M'. (\n. f n \\<^sub>L M') \ M \\<^sub>L M')" lemma lowner_is_lub_dim: assumes "lowner_is_lub M" shows "M \ carrier_mat dim dim" proof - have "f 0 \\<^sub>L M" using assms lowner_is_lub_def by auto then have 1: "dim_row (f 0) = dim_row M \ dim_col (f 0) = dim_col M" using lowner_le_def by auto moreover have 2: "f 0 \ carrier_mat dim dim" using dim by auto ultimately show ?thesis by auto qed lemma trace_adjoint_eq_u: fixes A :: "complex mat" shows "trace (A * adjoint A) = (\ i \ {0 ..< dim_row A}. \ j \ {0 ..< dim_col A}. (norm(A $$ (i,j)))\<^sup>2)" proof - have "trace (A * adjoint A) = (\ i \ {0 ..< dim_row A}. row A i \ conjugate (row A i))" by (simp add: trace_def cmod_def adjoint_def scalar_prod_def) also have "\ = (\ i \ {0 ..< dim_row A}. \ j \ {0 ..< dim_col A}. (norm(A $$ (i,j)))\<^sup>2)" proof (simp add: scalar_prod_def cmod_def) have cnjmul: "\ i ia. A $$ (i, ia) * cnj (A $$ (i, ia)) = ((complex_of_real (Re (A $$ (i, ia))))\<^sup>2 + (complex_of_real (Im (A $$ (i, ia))))\<^sup>2)" by (simp add: complex_mult_cnj) then have "\ i. (\ia = 0..ia = 0..2 + (complex_of_real (Im (A $$ (i, ia))))\<^sup>2))" by auto then show"(\i = 0..ia = 0..x = 0..xa = 0..2) + (\x = 0..xa = 0..2)" by auto qed finally show ?thesis . qed lemma trace_adjoint_element_ineq: fixes A :: "complex mat" assumes rindex: "i \ {0 ..< dim_row A}" and cindex: "j \ {0 ..< dim_col A}" shows "(norm(A $$ (i,j)))\<^sup>2 \ trace (A * adjoint A)" proof (simp add: trace_adjoint_eq_u) have ineqi: "(cmod (A $$ (i, j)))\<^sup>2 \ (\xa = 0..2)" using cindex member_le_sum[of j " {0 ..< dim_col A}" "\ x. (cmod (A $$ (i, x)))\<^sup>2"] by auto also have ineqj: "\ \ (\x = 0..xa = 0..2)" using rindex member_le_sum[of i " {0 ..< dim_row A}" "\ x. \xa = 0..2"] by (simp add: sum_nonneg) then show "(cmod (A $$ (i, j)))\<^sup>2 \ (\x = 0..xa = 0..2)" using ineqi by linarith qed lemma positive_is_normal: fixes A :: "complex mat" assumes pos: "positive A" shows "A * adjoint A = adjoint A * A" proof - have hA: "hermitian A" using positive_is_hermitian pos by auto then show ?thesis by (simp add: hA hermitian_is_normal) qed lemma diag_mat_mul_diag_diag: fixes A B :: "complex mat" assumes dimA: "A \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and dA: "diagonal_mat A" and dB: "diagonal_mat B" shows "diagonal_mat (A * B)" proof - have AB: "A * B = mat n n (\(i,j). (if (i = j) then (A$$(i, i)) * (B$$(i, i)) else 0))" using diag_mat_mult_diag_mat[of A n B] dimA dimB dA dB by auto then have dAB: "\ij j \ (A*B) $$ (i,j) = 0" proof - { fix i j assume i: "i < n" and j: "j < n" and ij: "i \ j" have "(A*B) $$ (i,j) = 0" using AB i j ij by auto } then show ?thesis by auto qed then show ?thesis using diagonal_mat_def dAB dimA dimB by (metis carrier_matD(1) carrier_matD(2) index_mult_mat(2) index_mult_mat(3)) qed lemma diag_mat_mul_diag_ele: fixes A B :: "complex mat" assumes dimA: "A \ carrier_mat n n" and dimB: "B \ carrier_mat n n" and dA: "diagonal_mat A" and dB: "diagonal_mat B" shows "\i(i,j). if i = j then (A$$(i, i)) * (B$$(i, i)) else 0)" using diag_mat_mult_diag_mat[of A n B] dimA dimB dA dB by auto then show ?thesis using AB by auto qed lemma trace_square_less_square_trace: fixes B :: "complex mat" assumes dimB: "B \ carrier_mat n n" and dB: "diagonal_mat B" and pB: "\i. i < n \ B$$(i, i) \ 0" shows "trace (B*B) \ (trace B)\<^sup>2" proof - have tB: "trace B = (\ i \ {0 ..2 = (\ i \ {0 .. j \ {0 ..i. i < n \ (B*B) $$ (i,i) = (B$$(i, i))\<^sup>2" using diag_mat_mul_diag_ele[of B n B] dimB dB by (metis numeral_1_eq_Suc_0 power_Suc0_right power_add_numeral semiring_norm(2)) have tBB: "trace (B*B) = (\ i \ {0 .. = (\ i \ {0 ..2)" using BB by auto finally have BBt: " trace (B * B) = (\i = 0..2)" by auto have lesseq: "\i \ {0 ..2 \ (\ j \ {0 ..j = 0..2 + sum (\ j. (B $$ (i, i) * B $$ (j, j))) ({0 .. j. (B $$ (i, i) * B $$ (j, j))) ({0 .. 0" proof (cases "{0.. {}") case True then show ?thesis using pB i sum_nonneg[of "{0.. j. (B $$ (i, i) * B $$ (j, j))"] by auto next case False have "(\j\{0..j = 0.. (B $$ (i, i))\<^sup>2" by auto } then show ?thesis by auto qed from tBtB BBt lesseq have "trace (B*B) \ (trace B)\<^sup>2" using sum_mono[of "{0.. i. (B $$ (i, i))\<^sup>2" "\ i. (\j = 0.. (trace A)\<^sup>2" proof - from assms have normal: "A * adjoint A = adjoint A * A" by (rule positive_is_normal) moreover from assms positive_dim_eq obtain n where cA: "A \ carrier_mat n n" by auto moreover from assms complex_mat_char_poly_factorizable cA obtain es where charpo: " char_poly A = (\ a \ es. [:- a, 1:]) \ length es = n" by auto moreover obtain B P Q where B: "unitary_schur_decomposition A es = (B,P,Q)" by (cases "unitary_schur_decomposition A es", auto) ultimately have smw: "similar_mat_wit A B P (adjoint P)" and ut: "diagonal_mat B" and uP: "unitary P" and dB: "diag_mat B = es" and QaP: "Q = adjoint P" using normal_complex_mat_has_spectral_decomposition[of A n es B P Q] unitary_schur_decomposition by auto from smw cA QaP uP have cB: "B \ carrier_mat n n" and cP: "P \ carrier_mat n n" and cQ: "Q \ carrier_mat n n" unfolding similar_mat_wit_def Let_def unitary_def by auto then have caP: "adjoint P \ carrier_mat n n" using adjoint_dim[of P n] by auto from smw QaP cA have A: "A = P * B * adjoint P" and traceA: "trace A = trace (P * B * Q)" and PB: "P * Q = 1\<^sub>m n \ Q * P = 1\<^sub>m n" unfolding similar_mat_wit_def by auto have traceAB: "trace (P * B * Q) = trace ((Q*P)*B)" using cQ cP cB by (mat_assoc n) also have traceelim: "\ = trace B" using traceAB PB cA cB cP cQ left_mult_one_mat[of "P*Q" n n] using similar_mat_wit_sym by auto finally have traceAB: "trace A = trace B" using traceA by auto from A cB cP have aAa: "adjoint A = adjoint((P * B) * adjoint P)" by auto have aA: "adjoint A = P * adjoint B * adjoint P" unfolding aAa using cP cB by (mat_assoc n) have hA: "hermitian A" using pos positive_is_hermitian by auto then have AaA: "A = adjoint A" using hA hermitian_def[of A] by auto then have PBaP: "P * B * adjoint P = P * adjoint B * adjoint P" using A aA by auto then have BaB: "B = adjoint B" using unitary_elim[of B n "adjoint B" P] uP cP cB adjoint_dim[of B n] by auto have aPP: "adjoint P * P = 1\<^sub>m n" using uP PB QaP by blast have "A * A = P * B * (adjoint P * P) * B * adjoint P" unfolding A using cP cB by (mat_assoc n) also have "\ = P * B * B * adjoint P" unfolding aPP using cP cB by (mat_assoc n) finally have AA: "A * A = P * B * B * adjoint P" by auto then have tAA: "trace (A*A) = trace (P * B * B * adjoint P)" by auto also have tBB: "\ = trace (adjoint P * P * B * B)" using cP cB by (mat_assoc n) also have "\ = trace (B * B)" using uP unitary_def[of P] inverts_mat_def[of P "adjoint P"] using PB QaP cB by auto finally have traceAABB: "trace (A * A) = trace (B * B)" by auto have BP: "\i. i < n \ B$$(i, i) \ 0" proof - { fix i assume i: "i < n" then have "B$$(i, i) \ 0" using positive_eigenvalue_positive[of A n es B P Q i] cA pos charpo B by auto then show "B$$(i, i) \ 0" by auto } qed have Brel: "trace (B*B) \ (trace B)\<^sup>2" using trace_square_less_square_trace[of B n] cB ut BP by auto from AaA traceAABB traceAB Brel have "trace (A*adjoint A) \ (trace A)\<^sup>2" by auto then show ?thesis by auto qed lemma lowner_le_transitive: fixes m n :: nat assumes re: "n \ m" shows "positive (f n - f m)" proof - from re show "positive (f n - f m)" proof (induct n) case 0 then show ?case using positive_zero by (metis dim le_0_eq minus_r_inv_mat) next case (Suc n) then show ?case proof (cases "Suc n = m") case True then show ?thesis using positive_zero by (metis dim minus_r_inv_mat) next case False then show ?thesis proof - from False Suc have nm: "n \ m" by linarith from Suc nm have pnm: "positive (f n - f m)" by auto from inc have "positive (f (Suc n) - f n)" unfolding lowner_le_def by auto then have pf: "positive ((f (Suc n) - f n) + (f n - f m))" using positive_add dim pnm by (meson minus_carrier_mat) have "(f (Suc n) - f n) + (f n - f m) = f (Suc n) + ((- f n) + f n) + (- f m)" using local.dim by (mat_assoc dim, auto) also have "\ = f (Suc n) + 0\<^sub>m dim dim + (- f m)" using local.dim by (subst uminus_l_inv_mat[where nc=dim and nr=dim], auto) also have "\ = f (Suc n) - f m" using local.dim by (mat_assoc dim, auto) finally have re: "f (Suc n) - f n + (f n - f m) = f (Suc n) - f m" . from pf re have "positive (f (Suc n) - f m)" by auto then show ?thesis by auto qed qed qed qed text \The sequence of matrices converges pointwise.\ lemma inc_partial_density_operator_converge: assumes i: "i \ {0 .. {0 ..n. f n $$ (i, j))" proof- have tracefn: "trace (f n) \ 0 \ trace (f n) \ 1" for n proof - from pdo show ?thesis unfolding partial_density_operator_def using positive_trace[of "f n"] using dim by blast qed from tracefn have normf: "norm(trace (f n)) \ norm(trace (f (Suc n))) \ norm(trace (f n)) \ 1" for n proof - have trless: "trace (f n) \ trace (f (Suc n))" using pdo inc dim positive_trace[of "f(Suc n) - f n"] trace_minus_linear[of "f (Suc n)" dim "f n"] unfolding partial_density_operator_def lowner_le_def using Complex_Matrix.positive_def by force moreover from trless tracefn have "norm(trace (f n)) \ norm(trace (f (Suc n)))" unfolding cmod_def by simp moreover from trless tracefn have "norm(trace (f n)) \ 1" using pdo partial_density_operator_def cmod_def by simp ultimately show ?thesis by auto qed then have inctrace: "incseq (\ n. norm(trace (f n)))" by (simp add: incseq_SucI) then have tr_sup: "(\ n. norm(trace (f n))) \ (SUP i. norm (trace (f i)))" using LIMSEQ_incseq_SUP[of "\ n. norm(trace (f n))"] pdo partial_density_operator_def normf by (meson bdd_aboveI2) then have tr_cauchy: "Cauchy (\ n. norm(trace (f n)))" using Cauchy_convergent_iff convergent_def by blast then have tr_cauchy_def: "\e>0. \M. \m\M. \n\M. dist(norm(trace (f n))) (norm(trace (f m))) < e" unfolding Cauchy_def by blast moreover have "\m n. dist(norm(trace (f m))) (norm(trace (f n))) = norm(trace (f m) - trace (f n))" using tracefn cmod_eq_Re dist_real_def by auto ultimately have norm_trace: "\e>0.\M. \m\M. \n\M. norm((trace (f n)) - (trace (f m))) < e" by auto have eq_minus: "\ m n. trace (f m) - trace (f n) = trace (f m - f n)" using trace_minus_linear dim by metis from eq_minus norm_trace have norm_trace_cauchy: "\e>0.\M. \m\M. \n\M. norm((trace (f n - f m))) < e" by auto then have norm_trace_cauchy_iff: "\e>0.\M. \m\M. \n\m. norm((trace (f n - f m))) < e" by (meson order_trans_rules(23)) then have norm_square: "\e>0.\M. \m\M. \n\m. (norm((trace (f n - f m))))\<^sup>2 < e\<^sup>2" by (metis abs_of_nonneg norm_ge_zero order_less_le real_sqrt_abs real_sqrt_less_iff) have tr_re: "\ m. \ n \ m. trace ((f n - f m) * adjoint (f n - f m)) \ ((trace (f n- f m)))\<^sup>2" using trace_positive_eq lowner_le_transitive by auto have tr_re_g: "\ m. \ n \ m. trace ((f n - f m) * adjoint (f n - f m)) \ 0" using lowner_le_transitive positive_trace trace_adjoint_positive by auto have norm_trace_fmn: "norm(trace ((f n - f m) * adjoint (f n - f m))) \ (norm(trace (f n - f m)))\<^sup>2" if nm: "n \ m" for m n proof - have mnA: "trace ((f n - f m) * adjoint (f n - f m)) \ (trace (f n - f m))\<^sup>2" using tr_re nm by auto have mnB: "trace ((f n - f m) * adjoint (f n - f m)) \ 0" using tr_re_g nm by auto from mnA mnB show ?thesis by (smt cmod_eq_Re less_eq_complex_def norm_power zero_complex.sel(1) zero_complex.sel(2)) qed then have cauchy_adj: "\M. \m\M. \n\m. norm(trace ((f n- f m) * adjoint (f n - f m))) < e\<^sup>2" if e: "e > 0" for e proof - have "\M. \m\M. \n\m. (cmod (trace (f n - f m)))\<^sup>2 < e\<^sup>2" using norm_square e by auto then obtain M where " \m\M. \n\m. (cmod (trace (f n - f m)))\<^sup>2 < e\<^sup>2" by auto then have "\m\M. \n\m. norm(trace ((f n- f m) * adjoint (f n - f m))) < e\<^sup>2" using norm_trace_fmn by fastforce then show ?thesis by auto qed have norm_minus: "\ m. \ n \ m. (norm ((f n - f m) $$ (i, j)))\<^sup>2 \ trace ((f n - f m) * adjoint (f n - f m))" using trace_adjoint_element_ineq i j by (smt adjoint_dim_row carrier_matD(1) index_minus_mat(2) index_mult_mat(2) lowner_le_transitive matrix_seq_axioms matrix_seq_def positive_is_normal) then have norm_minus_le: "(norm ((f n - f m) $$ (i, j)))\<^sup>2 \ norm (trace ((f n - f m) * adjoint (f n - f m)))" if nm: "n \ m" for n m proof - have "(norm ((f n - f m) $$ (i, j)))\<^sup>2 \ (trace ((f n - f m) * adjoint (f n - f m)))" using norm_minus nm by auto also have "\ = norm (trace ((f n - f m) * adjoint (f n - f m)))" using tr_re_g nm by (smt Re_complex_of_real less_eq_complex_def matrix_seq.trace_adjoint_eq_u matrix_seq_axioms mult_cancel_left2 norm_one norm_scaleR of_real_def of_real_hom.hom_zero) finally show ?thesis by auto qed from norm_minus_le cauchy_adj have cauchy_ij: "\M. \m\M. \n\m. (norm ((f n - f m) $$ (i, j)))\<^sup>2 < e\<^sup>2" if e: "e > 0" for e proof - have "\M. \m\M. \n\m. norm(trace ((f n- f m) * adjoint (f n - f m))) < e\<^sup>2" using cauchy_adj e by auto then obtain M where " \m\M. \n\m. norm(trace ((f n - f m) * adjoint (f n - f m))) < e\<^sup>2" by auto then have "\m\M. \n\m. (norm ((f n - f m) $$ (i, j)))\<^sup>2 < e\<^sup>2" using norm_minus_le by fastforce then show ?thesis by auto qed then have cauchy_ij_norm: "\M. \m\M. \n\m. (norm ((f n - f m) $$ (i, j))) < e" if e: "e > 0" for e proof - have "\M. \m\M. \n\m. (norm ((f n - f m) $$ (i, j)))\<^sup>2 < e\<^sup>2" using cauchy_ij e by auto then obtain M where mn: "\m\M. \n\m. (norm ((f n - f m) $$ (i, j)))\<^sup>2 < e\<^sup>2" by auto have "(norm ((f n - f m) $$ (i, j))) < e" if m: "m \ M" and n: "n \ m" for m n :: nat proof - from m n mn have "(norm ((f n- f m) $$ (i, j)))\<^sup>2 < e\<^sup>2" by auto then show ?thesis using e power_less_imp_less_base by fastforce qed then show ?thesis by auto qed have cauchy_final: "\M. \m\M. \n\M. norm ((f m) $$ (i, j) - (f n) $$ (i, j)) < e" if e: "e > 0" for e proof - obtain M where mnm: "\m\M. \n\m. norm ((f n - f m) $$ (i, j)) < e" using cauchy_ij_norm e by auto have "norm ((f m) $$ (i, j) - (f n) $$ (i, j)) < e" if m: "m \ M" and n: "n \ M" for m n proof (cases "n \ m") case True then show ?thesis proof - from mnm m True have "norm ((f n) $$ (i, j) - (f m) $$ (i, j)) < e" by (metis atLeastLessThan_iff carrier_matD(1) carrier_matD(2) dim i index_minus_mat(1) j) then have "norm ((f m) $$ (i, j) - (f n) $$ (i, j)) < e" by (simp add: norm_minus_commute) then show ?thesis by auto qed next case False then show ?thesis proof - from False n mnm have norm: "norm ((f m - f n) $$ (i, j)) < e" by auto have minus: "(f m - f n) $$ (i, j) = f m $$ (i, j) -f n $$ (i, j)" by (metis atLeastLessThan_iff carrier_matD(1) carrier_matD(2) dim i index_minus_mat(1) j) also have "\ = - (f n - f m) $$ (i, j)" using dim by (metis atLeastLessThan_iff carrier_matD(1) carrier_matD(2) i index_minus_mat(1) j minus_diff_eq) finally have fmn: "(f m - f n) $$ (i, j) = - (f n - f m) $$ (i, j)" by auto then have "norm ((- (f n - f m)) $$ (i, j)) < e" using norm by (metis (no_types, lifting) atLeastLessThan_iff carrier_matD(1) carrier_matD(2) i index_minus_mat(2) index_minus_mat(3) index_uminus_mat(1) j matrix_seq_axioms matrix_seq_def) then have "norm (((f n - f m)) $$ (i, j)) < e" using fmn norm by auto then have "norm (f n $$ (i, j) - f m $$ (i, j)) < e" by (metis minus norm norm_minus_commute) then have "norm (f m $$ (i, j) - f n $$ (i, j)) < e" by (simp add: norm_minus_commute) then show ?thesis by auto qed qed then show ?thesis by auto qed from cauchy_final have "Cauchy (\ n. f n $$ (i, j))" by (simp add: Cauchy_def dist_norm) then show ?thesis by (simp add: Cauchy_convergent_iff) qed definition mat_seq_minus :: "(nat \ complex mat) \ complex mat \ nat \ complex mat" where "mat_seq_minus X A = (\n. X n - A)" definition minus_mat_seq :: "complex mat \ (nat \ complex mat) \ nat \ complex mat" where "minus_mat_seq A X = (\n. A - X n)" lemma pos_mat_lim_is_pos_aux: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat assumes limX: "limit_mat X A m" and posX: "\k. \n\k. positive (X n)" shows "positive A" proof - from posX obtain k where posk: "\ n\k. positive (X n)" by auto let ?Y = "\n. X (n + k)" have posY: "\n. positive (?Y n)" using posk by auto from limX have dimXA: "\n. X (n + k) \ carrier_mat m m \ A \ carrier_mat m m" unfolding limit_mat_def by auto have "(\n. X (n + k) $$ (i, j)) \ A $$ (i, j)" if i: "i < m" and j: "j < m" for i j proof - have "(\n. X n $$ (i, j)) \ A $$ (i, j)" using limX limit_mat_def i j by auto then have limseqX: "\r>0. \no. \n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" unfolding LIMSEQ_def by auto then have "\no. \n\no. dist (X (n + k) $$ (i, j)) (A $$ (i, j)) < r" if r: "r > 0" for r proof - obtain no where "\n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" using limseqX r by auto then have "\n\no. dist (X (n + k) $$ (i, j)) (A $$ (i, j)) < r" by auto then show ?thesis by auto qed then show ?thesis unfolding LIMSEQ_def by auto qed then have limXA: "limit_mat (\n. X (n + k)) A m" unfolding limit_mat_def using dimXA by auto from posY limXA have "positive A" using pos_mat_lim_is_pos[of ?Y A m] by auto then show ?thesis by auto qed lemma minus_mat_limit: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes dimB: "B \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (mat_seq_minus X B) (A - B) m" proof - have dimXAB: "\n. X n - B \ carrier_mat m m \ A - B \ carrier_mat m m" using index_minus_mat dimB by auto have "(\n. (X n - B) $$ (i, j)) \ (A - B) $$ (i, j)" if i: "i < m" and j: "j < m" for i j proof - from limX i j have "(\n. (X n) $$ (i, j)) \ (A) $$ (i, j)" unfolding limit_mat_def by auto then have X: "\r>0. \no. \n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" unfolding LIMSEQ_def by auto then have XB: "\no. \n\no. dist ((X n - B) $$ (i, j)) ((A - B) $$ (i, j)) < r" if r: "r > 0" for r proof - obtain no where "\n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" using r X by auto then have dist: "\n\no. norm (X n $$ (i, j) - A $$ (i, j)) < r" unfolding dist_norm by auto then have "norm ((X n - B) $$ (i, j) - (A - B) $$ (i, j)) < r" if n: "n \ no" for n proof - have "(X n - B) $$ (i, j) - (A - B) $$ (i, j) = (X n) $$ (i, j) - A $$ (i, j)" using dimB i j by auto then have "norm ((X n - B) $$ (i, j) - (A - B) $$ (i, j)) = norm ((X n) $$ (i, j) - A $$ (i, j))" by auto then show ?thesis using dist n by auto qed then show ?thesis using dist_norm by metis qed then show ?thesis unfolding LIMSEQ_def by auto qed then show ?thesis unfolding limit_mat_def mat_seq_minus_def using dimXAB by auto qed lemma mat_minus_limit: fixes X :: "nat \ complex mat" and A :: "complex mat" and m :: nat and B :: "complex mat" assumes dimA: "A \ carrier_mat m m" and limX: "limit_mat X A m" shows "limit_mat (minus_mat_seq B X) (B - A) m" proof- have dimX : "\n. X n \ carrier_mat m m" using limX unfolding limit_mat_def by auto then have dimXAB: "\n. B - X n \ carrier_mat m m \ B - A \ carrier_mat m m" using index_minus_mat dimA by (simp add: minus_carrier_mat) have "(\n. (B - X n) $$ (i, j)) \ (B - A) $$ (i, j)" if i: "i < m" and j: "j < m" for i j proof - from limX i j have "(\n. (X n) $$ (i, j)) \ (A) $$ (i, j)" unfolding limit_mat_def by auto then have X: "\r>0. \no. \n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" unfolding LIMSEQ_def by auto then have XB: "\no. \n\no. dist ((B - X n) $$ (i, j)) ((B - A) $$ (i, j)) < r" if r: "r > 0" for r proof - obtain no where "\n\no. dist (X n $$ (i, j)) (A $$ (i, j)) < r" using r X by auto then have dist: "\n\no. norm (X n $$ (i, j) - A $$ (i, j)) < r" unfolding dist_norm by auto then have "norm ((B - X n) $$ (i, j) - (B - A) $$ (i, j)) < r" if n: "n \ no" for n proof - have "(B - X n) $$ (i, j) - (B - A) $$ (i, j) = - ((X n) $$ (i, j) - A $$ (i, j))" using dimA i j by (smt cancel_ab_semigroup_add_class.diff_right_commute cancel_comm_monoid_add_class.diff_cancel carrier_matD(1) carrier_matD(2) diff_add_cancel dimX index_minus_mat(1) minus_diff_eq) then have "norm ((B - X n) $$ (i, j) - (B - A) $$ (i, j)) = norm ((X n) $$ (i, j) - A $$ (i, j))" by (metis norm_minus_cancel) then show ?thesis using dist n by auto qed then show ?thesis using dist_norm by metis qed then show ?thesis unfolding LIMSEQ_def by auto qed then have "limit_mat (minus_mat_seq B X) (B - A) m" unfolding limit_mat_def minus_mat_seq_def using dimXAB by auto then show ?thesis by auto qed lemma lowner_lub_form: "lowner_is_lub (mat dim dim (\ (i, j). (lim (\ n. (f n) $$ (i, j)))))" proof - from inc_partial_density_operator_converge have conf: "\ i \ {0 .. j \ {0 .. n. f n $$ (i, j))" by auto let ?A = "mat dim dim (\ (i, j). (lim (\ n. (f n) $$ (i, j))))" have dim_A: "?A \ carrier_mat dim dim" by auto have lim_A: "(\n. f n $$ (i, j)) \ mat dim dim (\(i, j). lim (\n. f n $$ (i, j))) $$ (i, j)" if i: "i < dim" and j: "j < dim" for i j proof - from i j have ij: "mat dim dim (\(i, j). lim (\n. f n $$ (i, j))) $$ (i, j) = lim (\n. f n $$ (i, j))" by (metis case_prod_conv index_mat(1)) have "convergent (\n. f n $$ (i, j))" using conf i j by auto then have "(\n. f n $$ (i, j)) \ lim (\n. f n $$ (i, j)) " using convergent_LIMSEQ_iff by auto then show ?thesis using ij by auto qed from dim dim_A lim_A have lim_mat_A: "limit_mat f ?A dim" unfolding limit_mat_def by auto have is_ub: "f n \\<^sub>L ?A" for n proof - have "\ m \ n. positive (f m - f n)" using lowner_le_transitive by auto then have le: "\ m \ n. f n \\<^sub>L f m " unfolding lowner_le_def using dim by (metis carrier_matD(1) carrier_matD(2)) have dimn: "f n \ carrier_mat dim dim" using dim by auto then have limAf: "limit_mat (mat_seq_minus f (f n)) (?A - f n) dim" using minus_mat_limit lim_mat_A by auto have " \m\n. positive (f m - f n)" using lowner_le_transitive by auto then have "\k. \m\k. positive (f m - f n)" by auto then have posAf: "\ k. \ m \ k. positive ((mat_seq_minus f (f n)) m)" unfolding mat_seq_minus_def by auto from limAf posAf have "positive (?A - f n)" using pos_mat_lim_is_pos_aux by auto then have "f n \\<^sub>L mat dim dim (\(i, j). lim (\n. f n $$ (i, j)))" unfolding lowner_le_def using dim by auto then show ?thesis by auto qed have is_lub: "?A \\<^sub>L M'" if ub: "\n. f n \\<^sub>L M'" for M' proof - have dim_M: "M' \ carrier_mat dim dim" using ub unfolding lowner_le_def using dim by (metis carrier_matD(1) carrier_matD(2) carrier_mat_triv) from ub have posAf: "\ n. positive (minus_mat_seq M' f n)" unfolding minus_mat_seq_def lowner_le_def by auto have limAf: "limit_mat (minus_mat_seq M' f) (M' - ?A) dim" using mat_minus_limit dim_A lim_mat_A by auto from posAf limAf have "positive (M' - ?A)" using pos_mat_lim_is_pos_aux by auto then have "?A \\<^sub>L M'" unfolding lowner_le_def using dim dim_A dim_M by auto then show ?thesis by auto qed from is_ub is_lub show ?thesis unfolding lowner_is_lub_def by auto qed text \Lowner partial order is a complete partial order.\ lemma lowner_lub_exists: "\M. lowner_is_lub M" using lowner_lub_form by auto lemma lowner_lub_unique: "\!M. lowner_is_lub M" proof (rule HOL.ex_ex1I) show "\M. lowner_is_lub M" by (rule lowner_lub_exists) next fix M N assume M: "lowner_is_lub M" and N: "lowner_is_lub N" have Md: "M \ carrier_mat dim dim" using M by (rule lowner_is_lub_dim) have Nd: "N \ carrier_mat dim dim" using N by (rule lowner_is_lub_dim) have MN: "M \\<^sub>L N" using M N by (simp add: lowner_is_lub_def) have NM: "N \\<^sub>L M" using M N by (simp add: lowner_is_lub_def) show "M = N" using MN NM by (auto intro: lowner_le_antisym[OF Md Nd]) qed definition lowner_lub :: "complex mat" where "lowner_lub = (THE M. lowner_is_lub M)" lemma lowner_lub_prop: "lowner_is_lub lowner_lub" unfolding lowner_lub_def apply (rule HOL.theI') by (rule lowner_lub_unique) lemma lowner_lub_is_limit: "limit_mat f lowner_lub dim" proof - define A where "A = lowner_lub" then have "A = (THE M. lowner_is_lub M)" using lowner_lub_def by auto then have Af: "A = (mat dim dim (\ (i, j). (lim (\ n. (f n) $$ (i, j)))))" using lowner_lub_form lowner_lub_unique by auto show "limit_mat f A dim" unfolding Af limit_mat_def apply (auto simp add: dim) proof - fix i j assume dims: "i < dim" "j < dim" then have "convergent (\n. f n $$ (i, j))" using inc_partial_density_operator_converge by auto then show "(\n. f n $$ (i, j)) \ lim (\n. f n $$ (i, j))" using convergent_LIMSEQ_iff by auto qed qed lemma lowner_lub_trace: assumes "\ n. trace (f n) \ x" shows "trace lowner_lub \ x" proof - have "\ n. trace (f n) \ 0" using positive_trace pdo unfolding partial_density_operator_def using dim by blast then have Re: "\ n. Re (trace (f n)) \ 0 \ Im (trace (f n)) = 0" by auto then have lex: "\ n. Re (trace (f n)) \ Re x \ Im x = 0" using assms by auto have "limit_mat f lowner_lub dim" using lowner_lub_is_limit by auto then have conv: "(\n. trace (f n)) \ trace lowner_lub" using mat_trace_limit by auto then have "(\n. Re (trace (f n))) \ Re (trace lowner_lub)" by (simp add: tendsto_Re) then have Rell: "Re (trace lowner_lub) \ Re x" using lex Lim_bounded[of "(\n. Re (trace (f n)))" "Re (trace lowner_lub)" 0 "Re x"] by simp from conv have "(\n. Im (trace (f n))) \ Im (trace lowner_lub)" by (simp add: tendsto_Im) then have Imll: "Im (trace lowner_lub) = 0" using Re by (simp add: Lim_bounded Lim_bounded2 dual_order.antisym) from Rell Imll lex show ?thesis by simp qed lemma lowner_lub_is_positive: shows "positive lowner_lub" using lowner_lub_is_limit pos_mat_lim_is_pos pdo unfolding partial_density_operator_def by auto end subsection \Finite sum of matrices\ text \Add f in the interval [0, n)\ fun matrix_sum :: "nat \ (nat \ 'b::semiring_1 mat) \ nat \ 'b mat" where "matrix_sum d f 0 = 0\<^sub>m d d" | "matrix_sum d f (Suc n) = f n + matrix_sum d f n" definition matrix_inf_sum :: "nat \ (nat \ complex mat) \ complex mat" where "matrix_inf_sum d f = matrix_seq.lowner_lub (\n. matrix_sum d f n)" lemma matrix_sum_dim: fixes f :: "nat \ 'b::semiring_1 mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ matrix_sum d f n \ carrier_mat d d" proof (induct n) case 0 show ?case by auto next case (Suc n) then have "f n \ carrier_mat d d" by auto then show ?case using Suc by auto qed lemma matrix_sum_cong: fixes f :: "nat \ 'b::semiring_1 mat" shows "(\k. k < n \ f k = f' k) \ matrix_sum d f n = matrix_sum d f' n" proof (induct n) case 0 show ?case by auto next case (Suc n) then show ?case unfolding matrix_sum.simps by auto qed lemma matrix_sum_add: fixes f :: "nat \ 'b::semiring_1 mat" and g :: "nat \ 'b::semiring_1 mat" and h :: "nat \ 'b::semiring_1 mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ (\k. k < n \ g k \ carrier_mat d d) \ (\k. k < n \ h k \ carrier_mat d d) \ (\k. k < n \ f k = g k + h k) \ matrix_sum d f n = matrix_sum d g n + matrix_sum d h n" proof (induct n) case 0 then show ?case by auto next case (Suc n) then show ?case proof - have gh: "matrix_sum d g n \ carrier_mat d d \ matrix_sum d h n \ carrier_mat d d" using matrix_sum_dim Suc(3, 4) by (simp add: matrix_sum_dim) have nSuc: "n < Suc n" by auto have sumf: "matrix_sum d f n = matrix_sum d g n + matrix_sum d h n" using Suc by auto have "matrix_sum d f (Suc n) = matrix_sum d g (Suc n) + matrix_sum d h (Suc n)" unfolding matrix_sum.simps Suc(5)[OF nSuc] sumf apply (mat_assoc d) using gh Suc by auto then show ?thesis by auto qed qed lemma matrix_sum_smult: fixes f :: "nat \ 'b::semiring_1 mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ matrix_sum d (\ k. c \\<^sub>m f k) n = c \\<^sub>m matrix_sum d f n" proof (induct n) case 0 then show ?case by auto next case (Suc n) then show ?case apply auto using add_smult_distrib_left_mat Suc matrix_sum_dim by (metis lessI less_SucI) qed lemma matrix_sum_remove: fixes f :: "nat \ 'b::semiring_1 mat" assumes j: "j < n" and df: "(\k. k < n \ f k \ carrier_mat d d)" and f': "(\k. f' k = (if k = j then 0\<^sub>m d d else f k))" shows "matrix_sum d f n = f j + matrix_sum d f' n" proof - have df': "\k. k < n \ f' k \ carrier_mat d d" using f' df by auto have dsf: "k < n \ matrix_sum d f k \ carrier_mat d d" for k using matrix_sum_dim[OF df] by auto have dsf': "k < n \ matrix_sum d f' k \ carrier_mat d d" for k using matrix_sum_dim[OF df'] by auto have flj: "\k. k < j \ f' k = f k" using j f' by auto then have "matrix_sum d f j = matrix_sum d f' j" using matrix_sum_cong[of j f' f, OF flj] df df' j by auto then have eqj: "matrix_sum d f (Suc j) = f j + matrix_sum d f' (Suc j)" unfolding matrix_sum.simps by (subst (1) f', simp add: df dsf' j) have lm: "(j + 1) + l \ n \ matrix_sum d f ((j + 1) + l) = f j + matrix_sum d f' ((j + 1) + l)" for l proof (induct l) case 0 show ?case using j eqj by auto next case (Suc l) then have eq: "matrix_sum d f ((j + 1) + l) = f j + matrix_sum d f' ((j + 1) + l)" by auto have s: "((j + 1) + Suc l) = Suc ((j + 1) + l)" by simp have eqf': "f' (j + 1 + l) = f (j + 1 + l)" using f' Suc by auto have dims: "f (j + 1 + l) \ carrier_mat d d" "f j \ carrier_mat d d" "matrix_sum d f' (j + 1 + l) \ carrier_mat d d" using df df' dsf' Suc by auto show ?case apply (subst (1 2) s) unfolding matrix_sum.simps apply (subst eq, subst eqf') apply (mat_assoc d) using dims by auto qed have p: "(j + 1) + (n - j - 1) \ n" using j by auto show ?thesis using lm[OF p] j by auto qed lemma matrix_sum_Suc_remove_head: fixes f :: "nat \ complex mat" shows "(\k. k < n + 1 \ f k \ carrier_mat d d) \ matrix_sum d f (n + 1) = f 0 + matrix_sum d (\k. f (k + 1)) n" proof (induct n) case 0 then show ?case by auto next case (Suc n) then have dSS: "\k. k < Suc (Suc n) \ f k \ carrier_mat d d" by auto have ds: "matrix_sum d (\k. f (k + 1)) n \ carrier_mat d d" using matrix_sum_dim[OF dSS, of "n" "\k. k + 1"] by auto have "matrix_sum d f (Suc n + 1) = f (n + 1) + matrix_sum d f (n + 1)" by auto also have "\ = f (n + 1) + (f 0 + matrix_sum d (\k. f (k + 1)) n)" using Suc by auto also have "\ = f 0 + (f (n + 1) + matrix_sum d (\k. f (k + 1)) n)" using ds apply (mat_assoc d) using dSS by auto finally show ?case by auto qed lemma matrix_sum_positive: fixes f :: "nat \ complex mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ (\k. k < n \ positive (f k)) \ positive (matrix_sum d f n)" proof (induct n) case 0 show ?case using positive_zero by auto next case (Suc n) then have dfn: "f n \ carrier_mat d d" and psn: "positive (matrix_sum d f n)" and pn: "positive (f n)" and d: "k < n \ f k \ carrier_mat d d" for k by auto then have dsn: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim by auto show ?case unfolding matrix_sum.simps using positive_add[OF pn psn dfn dsn] by auto qed lemma matrix_sum_mult_right: shows "(\k. k < n \ f k \ carrier_mat d d) \ A \ carrier_mat d d \ matrix_sum d (\k. (f k) * A) n = matrix_sum d (\k. f k) n * A" proof (induct n) case 0 then show ?case by auto next case (Suc n) then have "k < n \ f k \ carrier_mat d d" and dfn: "f n \ carrier_mat d d" for k by auto then have dsfn: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim by auto have "(f n + matrix_sum d f n) * A = f n * A + matrix_sum d f n * A" apply (mat_assoc d) using Suc dsfn by auto also have "\ = f n * A + matrix_sum d (\k. f k * A) n" using Suc by auto finally show ?case by auto qed lemma matrix_sum_add_distrib: shows "(\k. k < n \ f k \ carrier_mat d d) \ (\k. k < n \ g k \ carrier_mat d d) \ matrix_sum d (\k. (f k) + (g k)) n = matrix_sum d f n + matrix_sum d g n" proof (induct n) case 0 then show ?case by auto next case (Suc n) then have dfn: "f n \ carrier_mat d d" and dgn: "g n \ carrier_mat d d" and dfk: "k < n \ f k \ carrier_mat d d" and dgk: "k < n \ g k \ carrier_mat d d" and eq: "matrix_sum d (\k. f k + g k) n = matrix_sum d f n + matrix_sum d g n" for k by auto have dsf: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim dfk by auto have dsg: "matrix_sum d g n \ carrier_mat d d" using matrix_sum_dim dgk by auto show ?case unfolding matrix_sum.simps eq using dfn dgn dsf dsg by (mat_assoc d) qed lemma matrix_sum_minus_distrib: fixes f g :: "nat \ complex mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ (\k. k < n \ g k \ carrier_mat d d) \ matrix_sum d (\k. (f k) - (g k)) n = matrix_sum d f n - matrix_sum d g n" proof - have eq: "-1 \\<^sub>m g k = - g k" for k by auto assume dfk: "\k. k < n \ f k \ carrier_mat d d" and dgk: "\k. k < n \ (g k) \ carrier_mat d d" then have "k < n \ (f k) - (g k) = (f k) + (- (g k))" for k by auto then have "matrix_sum d (\k. (f k) - (g k)) n = matrix_sum d (\k. (f k) + (- (g k))) n" using matrix_sum_cong[of n "\k. (f k) - (g k)"] dfk dgk by auto also have "\ = matrix_sum d f n + matrix_sum d (\k. - (g k)) n" using matrix_sum_add_distrib[of n "f"] dfk dgk by auto also have "\ = matrix_sum d f n - matrix_sum d g n" apply (subgoal_tac "matrix_sum d (\k. - (g k)) n = - matrix_sum d g n", auto) apply (subgoal_tac "- 1 \\<^sub>m matrix_sum d g n = - matrix_sum d g n") by (simp add: matrix_sum_smult[of n g d "-1", OF dgk, simplified eq, simplified], auto) finally show ?thesis . qed lemma matrix_sum_shift_Suc: shows "(\k. k < (Suc n) \ f k \ carrier_mat d d) \ matrix_sum d f (Suc n) = f 0 + matrix_sum d (\k. f (Suc k)) n" proof (induct n) case 0 then show ?case by auto next case (Suc n) have dfk: "k < Suc (Suc n) \ f k \ carrier_mat d d" for k using Suc by auto have dsSk: "k < Suc n \ matrix_sum d (\k. f (Suc k)) n \ carrier_mat d d" for k using matrix_sum_dim[of _ "\k. f (Suc k)"] dfk by fastforce have "matrix_sum d f (Suc (Suc n)) = f (Suc n) + matrix_sum d f (Suc n)" by auto also have "\ = f (Suc n) + f 0 + matrix_sum d (\k. f (Suc k)) n" using Suc dsSk assoc_add_mat[of "f (Suc n)" d d "f 0"] by fastforce also have "\ = f 0 + (f (Suc n) + matrix_sum d (\k. f (Suc k)) n)" apply (mat_assoc d) using dsSk dfk by auto also have "\ = f 0 + matrix_sum d (\k. f (Suc k)) (Suc n)" by auto finally show ?case . qed lemma lowner_le_matrix_sum: fixes f g :: "nat \ complex mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ (\k. k < n \ g k \ carrier_mat d d) \ (\k. k < n \ f k \\<^sub>L g k) \ matrix_sum d f n \\<^sub>L matrix_sum d g n" proof (induct n) case 0 show ?case unfolding matrix_sum.simps using lowner_le_refl[of "0\<^sub>m d d" d] by auto next case (Suc n) then have dfn: "f n \ carrier_mat d d" and dgn: "g n \ carrier_mat d d" and le1: "f n \\<^sub>L g n" by auto then have le2: "matrix_sum d f n \\<^sub>L matrix_sum d g n" using Suc by auto have "k < n \ f k \ carrier_mat d d" for k using Suc by auto then have dsf: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim by auto have "k < n \ g k \ carrier_mat d d" for k using Suc by auto then have dsg: "matrix_sum d g n \ carrier_mat d d" using matrix_sum_dim by auto show ?case unfolding matrix_sum.simps using lowner_le_add dfn dsf dgn dsg le1 le2 by auto qed lemma lowner_lub_add: assumes "matrix_seq d f" "matrix_seq d g" "\ n. trace (f n + g n) \ 1" shows "matrix_seq.lowner_lub (\n. f n + g n) = matrix_seq.lowner_lub f + matrix_seq.lowner_lub g" proof - have msf: "matrix_seq.lowner_is_lub f (matrix_seq.lowner_lub f)" using assms(1) matrix_seq.lowner_lub_prop by auto then have "limit_mat f (matrix_seq.lowner_lub f) d" using matrix_seq.lowner_lub_is_limit assms by auto then have lim1: "\ijn. f n $$ (i, j)) \ (matrix_seq.lowner_lub f) $$ (i, j)" using limit_mat_def assms by auto have msg: "matrix_seq.lowner_is_lub g (matrix_seq.lowner_lub g)" using assms(2) matrix_seq.lowner_lub_prop by auto then have "limit_mat g (matrix_seq.lowner_lub g) d" using matrix_seq.lowner_lub_is_limit assms by auto then have lim2: "\ijn. g n $$ (i, j)) \ (matrix_seq.lowner_lub g) $$ (i, j)" using limit_mat_def assms by auto have "\n. f n + g n \ carrier_mat d d" using assms unfolding matrix_seq_def by fastforce moreover have "\n. partial_density_operator (f n + g n)" using assms unfolding matrix_seq_def partial_density_operator_def using positive_add by blast moreover have "(f n + g n) \\<^sub>L (f (Suc n) + g (Suc n))" for n using assms unfolding matrix_seq_def using lowner_le_add[of "f n" d "f (Suc n)" "g n" "g (Suc n)"] by auto ultimately have msfg: "matrix_seq d (\n. f n + g n)" using assms unfolding matrix_seq_def by auto then have mslfg: "matrix_seq.lowner_is_lub (\n. f n + g n) (matrix_seq.lowner_lub (\n. f n + g n))" using matrix_seq.lowner_lub_prop by auto then have "limit_mat (\n. f n + g n) (matrix_seq.lowner_lub (\n. f n + g n)) d" using matrix_seq.lowner_lub_is_limit msfg by auto then have lim3: "\ijn. (f n + g n) $$ (i, j)) \ (matrix_seq.lowner_lub (\n. f n + g n)) $$ (i, j)" using limit_mat_def assms by auto have "\ i j n. (f n + g n) $$ (i, j) = f n $$ (i, j) + g n $$ (i, j)" using assms unfolding matrix_seq_def by (metis carrier_matD(1) carrier_matD(2) index_add_mat(1)) then have add: "\ijn. f n $$ (i, j) + g n $$ (i, j)) \ (matrix_seq.lowner_lub (\n. f n + g n)) $$ (i, j)" using lim3 by auto have "matrix_seq.lowner_lub f $$ (i, j) + matrix_seq.lowner_lub g $$ (i, j) = matrix_seq.lowner_lub (\n. f n + g n) $$ (i, j)" if i: "i < d" and j: "j < d" for i j proof - have "(\n. f n $$ (i, j)) \ matrix_seq.lowner_lub f $$ (i, j)" using lim1 i j by auto moreover have "(\n. g n $$ (i, j)) \ matrix_seq.lowner_lub g $$ (i, j)" using lim2 i j by auto ultimately have "(\n. f n $$ (i, j) + g n $$ (i, j)) \ matrix_seq.lowner_lub f $$ (i, j) + matrix_seq.lowner_lub g $$ (i, j)" - using tendsto_intros(12)[of "\n. f n $$ (i, j)" "matrix_seq.lowner_lub f $$ (i, j)" sequentially "\n. g n $$ (i, j)" "matrix_seq.lowner_lub g $$ (i, j)"] by auto + using tendsto_add[of "\n. f n $$ (i, j)" "matrix_seq.lowner_lub f $$ (i, j)" sequentially "\n. g n $$ (i, j)" "matrix_seq.lowner_lub g $$ (i, j)"] by auto moreover have "(\n. f n $$ (i, j) + g n $$ (i, j)) \ matrix_seq.lowner_lub (\n. f n + g n) $$ (i, j)" using add i j by auto ultimately show ?thesis using LIMSEQ_unique by auto qed moreover have "matrix_seq.lowner_lub f \ carrier_mat d d" using matrix_seq.lowner_is_lub_dim assms(1) msf unfolding matrix_seq_def by auto moreover have "matrix_seq.lowner_lub g \ carrier_mat d d" using matrix_seq.lowner_is_lub_dim assms(2) msg unfolding matrix_seq_def by auto moreover have "matrix_seq.lowner_lub (\n. f n + g n) \ carrier_mat d d" using matrix_seq.lowner_is_lub_dim msfg mslfg unfolding matrix_seq_def by auto ultimately show ?thesis unfolding matrix_seq_def using mat_eq_iff by auto qed lemma lowner_lub_scale: fixes c :: real assumes "matrix_seq d f" "\ n. trace (c \\<^sub>m f n) \ 1" "c\0" shows "matrix_seq.lowner_lub (\n. c \\<^sub>m f n) = c \\<^sub>m matrix_seq.lowner_lub f" proof - have msf: "matrix_seq.lowner_is_lub f (matrix_seq.lowner_lub f)" using assms(1) matrix_seq.lowner_lub_prop by auto then have "limit_mat f (matrix_seq.lowner_lub f) d" using matrix_seq.lowner_lub_is_limit assms by auto then have lim1: "\ijn. f n $$ (i, j)) \ (matrix_seq.lowner_lub f) $$ (i, j)" using limit_mat_def assms by auto have dimcf: "\n. c \\<^sub>m f n \ carrier_mat d d" using assms unfolding matrix_seq_def by fastforce moreover have "\n. partial_density_operator (c \\<^sub>m f n)" using assms unfolding matrix_seq_def partial_density_operator_def using positive_scale by blast moreover have "\n. c \\<^sub>m f n \\<^sub>L c \\<^sub>m f (Suc n)" using lowner_le_smult assms(1,3) unfolding matrix_seq_def partial_density_operator_def by blast ultimately have mscf: "matrix_seq d (\n. c \\<^sub>m f n)" unfolding matrix_seq_def by auto then have mslfg: "matrix_seq.lowner_is_lub (\n. c \\<^sub>m f n) (matrix_seq.lowner_lub (\n. c \\<^sub>m f n))" using matrix_seq.lowner_lub_prop by auto then have "limit_mat (\n. c \\<^sub>m f n) (matrix_seq.lowner_lub (\n. c \\<^sub>m f n)) d" using matrix_seq.lowner_lub_is_limit mscf by auto then have lim3: "\ijn. (c \\<^sub>m f n) $$ (i, j)) \ (matrix_seq.lowner_lub (\n. c \\<^sub>m f n)) $$ (i, j)" using limit_mat_def assms by auto from mslfg mscf have dleft: "matrix_seq.lowner_lub (\n. c \\<^sub>m f n) \ carrier_mat d d" using matrix_seq.lowner_is_lub_dim by auto have dllf: "matrix_seq.lowner_lub f \ carrier_mat d d" using matrix_seq.lowner_is_lub_dim assms(1) msf unfolding matrix_seq_def by auto then have dright: "c \\<^sub>m matrix_seq.lowner_lub f \ carrier_mat d d" using index_smult_mat(2,3) by auto have "\ i j n. (c \\<^sub>m f n) $$ (i, j) = c * f n $$ (i, j)" using assms(1) unfolding matrix_seq_def using index_smult_mat(1) by (metis carrier_matD(1-2)) then have smult: "\ijn. c * f n $$ (i, j)) \ (matrix_seq.lowner_lub (\n. c \\<^sub>m f n)) $$ (i, j)" using lim3 by auto have ij: "(c \\<^sub>m matrix_seq.lowner_lub f) $$ (i, j) = (matrix_seq.lowner_lub (\n. c \\<^sub>m f n)) $$ (i, j)" if i: "i < d" and j: "j < d" for i j proof - have "(\n. f n $$ (i, j)) \ matrix_seq.lowner_lub f $$ (i, j)" using lim1 i j by auto moreover have "\ij\<^sub>m matrix_seq.lowner_lub f) $$ (i, j) = c * matrix_seq.lowner_lub f $$ (i, j)" using index_smult_mat dllf by fastforce ultimately have "\ijn. c * f n $$ (i, j)) \(c \\<^sub>m matrix_seq.lowner_lub f) $$ (i, j)" using tendsto_intros(18)[of "\n. c" "c" sequentially "\n. f n $$ (i, j)" "matrix_seq.lowner_lub f $$ (i, j)"] i j by (simp add: lim1 tendsto_mult_left) then show ?thesis using smult i j LIMSEQ_unique by metis qed from dleft dright ij show ?thesis using mat_eq_iff[of "matrix_seq.lowner_lub (\n. c \\<^sub>m f n)" "c \\<^sub>m matrix_seq.lowner_lub f"] by (metis (mono_tags) carrier_matD(1) carrier_matD(2)) qed lemma trace_matrix_sum_linear: fixes f :: "nat \ complex mat" shows "(\k. k < n \ f k \ carrier_mat d d) \ trace (matrix_sum d f n) = sum (\k. trace (f k)) {0..k. k < n \ f k \ carrier_mat d d" by auto then have ds: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim by auto have "trace (matrix_sum d f (Suc n)) = trace (f n) + trace (matrix_sum d f n)" unfolding matrix_sum.simps apply (mat_assoc d) using ds Suc by auto also have "\ = sum (trace \ f) {0.. f) n" using Suc by auto also have "\ = sum (trace \ f) {0.. complex mat" shows "P \ carrier_mat d d \ (\k. k < n \ f k \ carrier_mat d d) \ matrix_sum d (\k. P * (f k)) n = P * (matrix_sum d f n)" proof (induct n) case 0 show ?case unfolding matrix_sum.simps using 0 by auto next case (Suc n) then have "\k. k < n \ f k \ carrier_mat d d" by auto then have ds: "matrix_sum d f n \ carrier_mat d d" using matrix_sum_dim by auto then have dPf: "\k. k < n \ P * f k \ carrier_mat d d" using Suc by auto then have "matrix_sum d (\k. P * f k) n \ carrier_mat d d" using matrix_sum_dim[OF dPf] by auto have "matrix_sum d (\k. P * f k) (Suc n) = P * f n + matrix_sum d (\k. P * f k) n " unfolding matrix_sum.simps using Suc(2) by auto also have "\ = P * f n + P * matrix_sum d f n" using Suc by auto also have "\ = P * (f n + matrix_sum d f n)" apply (mat_assoc d) using ds dPf Suc by auto finally show "matrix_sum d (\k. P * f k) (Suc n) = P * (matrix_sum d f (Suc n))" by auto qed subsection \Measurement\ definition measurement :: "nat \ nat \ (nat \ complex mat) \ bool" where "measurement d n M \ (\j < n. M j \ carrier_mat d d) \ matrix_sum d (\j. (adjoint (M j)) * M j) n = 1\<^sub>m d" lemma measurement_dim: assumes "measurement d n M" shows "\k. k < n \ (M k) \ carrier_mat d d" using assms unfolding measurement_def by auto lemma measurement_id2: assumes "measurement d 2 M" shows "adjoint (M 0) * M 0 + adjoint (M 1) * M 1 = 1\<^sub>m d" proof - have ssz: "(Suc (Suc 0)) = 2" by auto have "M 0 \ carrier_mat d d" "M 1 \ carrier_mat d d" using assms measurement_def by auto then have "adjoint (M 0) * M 0 + adjoint (M 1) * M 1 = matrix_sum d (\j. (adjoint (M j)) * M j) (Suc (Suc 0)) " by auto also have "\ = matrix_sum d (\j. (adjoint (M j)) * M j) (2::nat)" by (subst ssz, auto) also have "\ = 1\<^sub>m d" using measurement_def[of d 2 M] assms by auto finally show ?thesis by auto qed text \Result of measurement on $\rho$ by matrix M\ definition measurement_res :: "complex mat \ complex mat \ complex mat" where "measurement_res M \ = M * \ * adjoint M" lemma add_positive_le_reduce1: assumes dA: "A \ carrier_mat n n" and dB: "B \ carrier_mat n n" and dC: "C \ carrier_mat n n" and pB: "positive B" and le: "A + B \\<^sub>L C" shows "A \\<^sub>L C" unfolding lowner_le_def positive_def proof (auto simp add: carrier_matD[OF dA] carrier_matD[OF dC] simp del: less_eq_complex_def) have eq: "C - (A + B) = (C - A + (-B))" using dA dB dC by auto have "positive (C - (A + B))" using le lowner_le_def dA dB dC by auto with eq have p: "positive (C - A + (-B))" by auto fix v :: "complex vec" assume " n = dim_vec v" then have dv: "v \ carrier_vec n" by auto have ge: "inner_prod v (B *\<^sub>v v) \ 0" using pB dv dB positive_def by auto have "0 \ inner_prod v ((C - A + (-B)) *\<^sub>v v) " using p positive_def dv dA dB dC by auto also have "\ = inner_prod v ((C - A)*\<^sub>v v + (-B) *\<^sub>v v) " using dv dA dB dC add_mult_distrib_mat_vec[OF minus_carrier_mat[OF dA]] by auto also have "\ = inner_prod v ((C - A) *\<^sub>v v) + inner_prod v ((-B) *\<^sub>v v)" apply (subst inner_prod_distrib_right) by (rule dv, auto simp add: mult_mat_vec_carrier[OF minus_carrier_mat[OF dA]] mult_mat_vec_carrier[OF uminus_carrier_mat[OF dB]] dv) also have "\ = inner_prod v ((C - A) *\<^sub>v v) - inner_prod v (B *\<^sub>v v)" using dB dv by auto also have "\ \ inner_prod v ((C - A) *\<^sub>v v)" using ge by auto finally show "0 \ inner_prod v ((C - A) *\<^sub>v v)". qed lemma add_positive_le_reduce2: assumes dA: "A \ carrier_mat n n" and dB: "B \ carrier_mat n n" and dC: "C \ carrier_mat n n" and pB: "positive B" and le: "B + A \\<^sub>L C" shows "A \\<^sub>L C" apply (subgoal_tac "B + A = A + B") using add_positive_le_reduce1[of A n B C] assms by auto lemma measurement_le_one_mat: assumes "measurement d n f" shows "\j. j < n \ adjoint (f j) * f j \\<^sub>L 1\<^sub>m d" proof - fix j assume j: "j < n" define M where "M = adjoint (f j) * f j" have df: "k < n \ f k \ carrier_mat d d" for k using assms measurement_dim by auto have daf: "k < n \ adjoint (f k) * f k \ carrier_mat d d" for k proof - assume "k < n" then have "f k \ carrier_mat d d" "adjoint (f k) \ carrier_mat d d" using df adjoint_dim by auto then show "adjoint (f k) * f k \ carrier_mat d d" by auto qed have pafj: "k < n \ positive (adjoint (f k) * (f k)) " for k apply (subst (2) adjoint_adjoint[of "f k", symmetric]) by (metis adjoint_adjoint daf positive_if_decomp) define f' where "\k. f' k = (if k = j then 0\<^sub>m d d else adjoint (f k) * f k)" have pf': "k < n \ positive (f' k)" for k unfolding f'_def using positive_zero pafj j by auto have df': "k < n \ f' k \ carrier_mat d d" for k using daf j zero_carrier_mat f'_def by auto then have dsf': "matrix_sum d f' n \ carrier_mat d d" using matrix_sum_dim[of n f' d] by auto have psf': "positive (matrix_sum d f' n)" using matrix_sum_positive pafj df' pf' by auto have "M + matrix_sum d f' n = matrix_sum d (\k. adjoint (f k) * f k) n" using matrix_sum_remove[OF j , of "(\k. adjoint (f k) * f k)", OF daf, of f'] f'_def unfolding M_def by auto also have "\ = 1\<^sub>m d" using measurement_def assms by auto finally have "M + matrix_sum d f' n = 1\<^sub>m d". moreover have "1\<^sub>m d \\<^sub>L 1\<^sub>m d" using lowner_le_refl[of _ d] by auto ultimately have "(M + matrix_sum d f' n) \\<^sub>L 1\<^sub>m d" by auto then show "M \\<^sub>L 1\<^sub>m d" unfolding M_def using add_positive_le_reduce1[OF _ dsf' one_carrier_mat psf'] daf j by auto qed lemma pdo_close_under_measurement: fixes M \ :: "complex mat" assumes dM: "M \ carrier_mat n n" and dr: "\ \ carrier_mat n n" and pdor: "partial_density_operator \" and le: "adjoint M * M \\<^sub>L 1\<^sub>m n" shows "partial_density_operator (M * \ * adjoint M)" unfolding partial_density_operator_def proof show "positive (M * \ * adjoint M)" using positive_close_under_left_right_mult_adjoint[OF dM dr] pdor partial_density_operator_def by auto next have daM: "adjoint M \ carrier_mat n n" using dM by auto then have daMM: "adjoint M * M \ carrier_mat n n" using dM by auto have "trace (M * \ * adjoint M) = trace (adjoint M * M * \)" using dM dr by (mat_assoc n) also have "\ \ trace (1\<^sub>m n * \)" using lowner_le_trace[where ?B = "1\<^sub>m n" and ?A = "adjoint M * M", OF daMM one_carrier_mat] le dr pdor by auto also have "\ = trace \" using dr by auto also have "\ \ 1" using pdor partial_density_operator_def by auto finally show "trace (M * \ * adjoint M) \ 1" by auto qed lemma trace_measurement: assumes m: "measurement d n M" and dA: "A \ carrier_mat d d" shows "trace (matrix_sum d (\k. (M k) * A * adjoint (M k)) n) = trace A" proof - have dMk: "k < n \ (M k) \ carrier_mat d d" for k using m unfolding measurement_def by auto then have daMk: "k < n \ adjoint (M k) \ carrier_mat d d" for k using m adjoint_dim unfolding measurement_def by auto have d1: "k < n \ M k * A * adjoint (M k) \ carrier_mat d d"for k using dMk daMk dA by fastforce then have ds1: "k < n \ matrix_sum d (\k. M k * A * adjoint (M k)) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. M k * A * adjoint (M k)" d] by auto have d2: "k < n \ adjoint (M k) *M k * A \ carrier_mat d d" for k using daMk dMk dA by fastforce then have ds2: "k < n \ matrix_sum d (\k. adjoint (M k) *M k * A) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. adjoint (M k) *M k * A" d] by auto have daMMk: "k < n \ adjoint (M k) * M k \ carrier_mat d d" for k using dMk by fastforce have "k \ n \ trace (matrix_sum d (\k. (M k) * A * adjoint (M k)) k) = trace (matrix_sum d (\k. adjoint (M k) * (M k) * A) k)" for k proof (induct k) case 0 then show ?case by auto next case (Suc k) then have k: "k < n" by auto have "trace (M k * A * adjoint (M k)) = trace (adjoint (M k) * M k * A)" using dA apply (mat_assoc d) using dMk k by auto then show ?case unfolding matrix_sum.simps using ds1 ds2 d1 d2 k Suc daMk dMk dA by (subst trace_add_linear[of _ d], auto)+ qed then have "trace (matrix_sum d (\k. (M k) * A * adjoint (M k)) n) = trace (matrix_sum d (\k. adjoint (M k) * (M k) * A) n)" by auto also have "\ = trace (matrix_sum d (\k. adjoint (M k) * (M k)) n * A)" using matrix_sum_mult_right[OF daMMk, of n id A] dA by auto also have "\ = trace A" using m dA unfolding measurement_def by auto finally show ?thesis by auto qed lemma mat_inc_seq_positive_transform: assumes dfn: "\n. f n \ carrier_mat d d" and inc: "\n. f n \\<^sub>L f (Suc n)" shows "\n. f n - f 0 \ carrier_mat d d" and "\n. (f n - f 0) \\<^sub>L (f (Suc n) - f 0)" proof - show "\n. f n - f 0 \ carrier_mat d d" using dfn by fastforce have "f 0 \\<^sub>L f 0" using lowner_le_refl[of "f 0" d] dfn by auto then show "(f n - f 0) \\<^sub>L (f (Suc n) - f 0)" for n using lowner_le_minus[of "f n" d "f (Suc n)" "f 0" "f 0"] dfn inc by fastforce qed lemma mat_inc_seq_lub: assumes dfn: "\n. f n \ carrier_mat d d" and inc: "\n. f n \\<^sub>L f (Suc n)" and ub: "\n. f n \\<^sub>L A" shows "\B. lowner_is_lub f B \ limit_mat f B d" proof - have dmfn0: "\n. f n - f 0 \ carrier_mat d d" and incm0: "\n. (f n - f 0) \\<^sub>L (f (Suc n) - f 0)" using mat_inc_seq_positive_transform[OF dfn, of id] assms by auto define c where "c = 1 / (trace (A - f 0) + 1)" have "f 0 \\<^sub>L A" using ub by auto then have dA: "A \ carrier_mat d d" using ub unfolding lowner_le_def using dfn[of 0] by fastforce then have dAmf0: "A - f 0 \ carrier_mat d d" using dfn[of 0] by auto have "positive (A - f 0)" using ub lowner_le_def by auto then have tgeq0: "trace (A - f 0) \ 0" using positive_trace dAmf0 by auto then have "trace (A - f 0) + 1 > 0" by auto then have gtc: "c > 0" unfolding c_def using complex_is_Real_iff by auto then have gtci: "(1 / c) > 0" using complex_is_Real_iff by auto have "trace (c \\<^sub>m (A - f 0)) = c * trace (A - f 0)" using trace_smult dAmf0 by auto also have "\ = (1 / (trace (A - f 0) + 1)) * trace (A - f 0)" unfolding c_def by auto also have "\ < 1" using tgeq0 by (simp add: complex_is_Real_iff) finally have lt1: "trace (c \\<^sub>m (A - f 0)) < 1". have le0: "- f 0 \\<^sub>L - f 0" using lowner_le_refl[of "- f 0" d] dfn by auto have dmf0: "- f 0 \ carrier_mat d d" using dfn by auto have mf0smcle: "(c \\<^sub>m (X - f 0)) \\<^sub>L (c \\<^sub>m (Y - f 0))" if "X \\<^sub>L Y" and "X \ carrier_mat d d" and "Y \ carrier_mat d d" for X Y proof - have "(X - f 0) \\<^sub>L (Y - f 0)" using lowner_le_minus[of "X" d "Y" "f 0" "f 0"] that dfn lowner_le_refl by auto then show ?thesis using lowner_le_smultc[of c "(X - f 0)" "Y - f 0" d] using that dfn gtc by fastforce qed have "(c \\<^sub>m (f n - f 0)) \\<^sub>L (c \\<^sub>m (A - f 0))" for n using mf0smcle ub dfn dA by auto then have "trace (c \\<^sub>m (f n - f 0)) \ trace (c \\<^sub>m (A - f 0))" for n using lowner_le_imp_trace_le[of "c \\<^sub>m (f n - f 0)" d] dmfn0 dAmf0 by auto then have trlt1: "trace (c \\<^sub>m (f n - f 0)) < 1" for n using lt1 by fastforce have "f 0 \\<^sub>L f n" for n proof (induct n) case 0 then show ?case using dfn lowner_le_refl by auto next case (Suc n) then show ?case using dfn lowner_le_trans[of "f 0" d "f n"] inc by auto qed then have "positive (f n - f 0)" for n using lowner_le_def by auto then have p: "positive (c \\<^sub>m (f n - f 0))" for n by (intro positive_smult, insert gtc dmfn0, auto) have inc': "c \\<^sub>m (f n - f 0) \\<^sub>L c \\<^sub>m (f (Suc n) - f 0)" for n using incm0 lowner_le_smultc[of c "f n - f 0"] gtc dmfn0 by fastforce define g where "g n = c \\<^sub>m (f n - f 0)" for n then have "positive (g n)" and "trace (g n) < 1" and "(g n) \\<^sub>L (g (Suc n))" and dgn: "(g n) \ carrier_mat d d" for n unfolding g_def using p trlt1 inc' dmfn0 by auto then have ms: "matrix_seq d g" unfolding matrix_seq_def partial_density_operator_def by fastforce then have uniM: "\!M. matrix_seq.lowner_is_lub g M" using matrix_seq.lowner_lub_unique by auto then obtain M where M: "matrix_seq.lowner_is_lub g M" by auto then have leg: "g n \\<^sub>L M" and lubg: "\M'. (\n. g n \\<^sub>L M') \ M \\<^sub>L M'" for n unfolding matrix_seq.lowner_is_lub_def[OF ms] by auto have "M = matrix_seq.lowner_lub g" using matrix_seq.lowner_lub_def[OF ms] M uniM theI_unique[of "matrix_seq.lowner_is_lub g"] by auto then have limg: "limit_mat g M d" using M matrix_seq.lowner_lub_is_limit[OF ms] by auto then have dM: "M \ carrier_mat d d" unfolding limit_mat_def by auto define B where "B = f 0 + (1 / c) \\<^sub>m M" have eqinv: "f 0 + (1 / c) \\<^sub>m (c \\<^sub>m (X - f 0)) = X" if "X \ carrier_mat d d" for X proof - have "f 0 + (1 / c) \\<^sub>m (c \\<^sub>m (X - f 0)) = f 0 + (1 / c * c) \\<^sub>m (X - f 0)" apply (subgoal_tac "(1 / c) \\<^sub>m (c \\<^sub>m (X - f 0)) = (1 / c * c) \\<^sub>m (X - f 0)", simp) using smult_smult_mat dfn that by auto also have "\ = f 0 + 1 \\<^sub>m (X - f 0)" using gtc by auto also have "\ = f 0 + (X - f 0)" by auto also have "\ = (- f 0) + f 0 + X" apply (mat_assoc d) using that dfn by auto also have "\ = 0\<^sub>m d d + X" using dfn uminus_l_inv_mat[of "f 0" d d] by fastforce also have "\ = X" using that by auto finally show ?thesis by auto qed have "limit_mat (\n. (1 / c) \\<^sub>m g n) ((1 / c) \\<^sub>m M) d" using limit_mat_scale[OF limg] gtci by auto then have "limit_mat (\n. f 0 + (1 / c) \\<^sub>m g n) (f 0 + (1 / c) \\<^sub>m M ) d" using mat_add_limit[of "f 0"] limg dfn unfolding mat_add_seq_def by auto then have limf: "limit_mat f B d" using eqinv[OF dfn] unfolding B_def g_def by auto have f0acmcile: "(f 0 + (1 / c) \\<^sub>m X) \\<^sub>L (f 0 + (1 / c) \\<^sub>m Y )" if "X \\<^sub>L Y" and "X \ carrier_mat d d" and "Y \ carrier_mat d d" for X Y proof - have "((1 / c) \\<^sub>m X) \\<^sub>L ((1 / c) \\<^sub>m Y)" using lowner_le_smultc[of "1/c"] that gtci by fastforce then show "(f 0 + (1 / c) \\<^sub>m X) \\<^sub>L (f 0 + (1 / c) \\<^sub>m Y)" using lowner_le_add[of _ d _ "(1 / c) \\<^sub>m X" "(1 / c) \\<^sub>m Y"] that gtci dfn lowner_le_refl[of "f 0", OF dfn] by fastforce qed have "(f 0 + (1 / c) \\<^sub>m g n) \\<^sub>L (f 0 + (1 / c) \\<^sub>m M )" for n using f0acmcile[OF leg dgn dM] by auto then have lubf: "f n \\<^sub>L B" for n using eqinv[OF dfn] g_def B_def by auto { fix B' assume asm: "\n. f n \\<^sub>L B'" then have "f 0 \\<^sub>L B'" by auto then have dB': "B' \ carrier_mat d d" unfolding lowner_le_def using dfn[of 0] by auto have "f n \\<^sub>L B'" for n using asm by auto then have "(c \\<^sub>m (f n - f 0)) \\<^sub>L (c \\<^sub>m (B' - f 0))" for n using mf0smcle[of "f n" B'] dfn dB' by auto then have "g n \\<^sub>L (c \\<^sub>m (B' - f 0))" for n using g_def by auto then have "M \\<^sub>L (c \\<^sub>m (B' - f 0))" using lubg by auto then have "(f 0 + (1 / c) \\<^sub>m M) \\<^sub>L (f 0 + (1 / c) \\<^sub>m (c \\<^sub>m (B' - f 0)))" using f0acmcile[of "M" "(c \\<^sub>m (B' - f 0))", OF _ dM] using dB' dfn by fastforce then have "B \\<^sub>L B'" unfolding B_def using eqinv[OF dB'] by auto } with limf lubf have "((\n. f n \\<^sub>L B) \ (\M'. (\n. f n \\<^sub>L M') \ B \\<^sub>L M')) \ limit_mat f B d" by auto then show ?thesis unfolding lowner_is_lub_def by auto qed end diff --git a/thys/QHLProver/Quantum_Hoare.thy b/thys/QHLProver/Quantum_Hoare.thy --- a/thys/QHLProver/Quantum_Hoare.thy +++ b/thys/QHLProver/Quantum_Hoare.thy @@ -1,1417 +1,1417 @@ section \Partial and total correctness\ theory Quantum_Hoare imports Quantum_Program begin context state_sig begin definition density_states :: "state set" where "density_states = {\ \ carrier_mat d d. partial_density_operator \}" lemma denote_density_states: "\ \ density_states \ well_com S \ denote S \ \ density_states" by (simp add: denote_dim_pdo density_states_def) definition is_quantum_predicate :: "complex mat \ bool" where "is_quantum_predicate P \ P \ carrier_mat d d \ positive P \ P \\<^sub>L 1\<^sub>m d" lemma trace_measurement2: assumes m: "measurement n 2 M" and dA: "A \ carrier_mat n n" shows "trace ((M 0) * A * adjoint (M 0)) + trace ((M 1) * A * adjoint (M 1)) = trace A" proof - from m have dM0: "M 0 \ carrier_mat n n" and dM1: "M 1 \ carrier_mat n n" and id: "adjoint (M 0) * (M 0) + adjoint (M 1) * (M 1) = 1\<^sub>m n" using measurement_def measurement_id2 by auto have "trace (M 1 * A * adjoint (M 1)) + trace (M 0 * A * adjoint (M 0)) = trace ((adjoint (M 0) * M 0 + adjoint (M 1) * M 1) * A)" using dM0 dM1 dA by (mat_assoc n) also have "\ = trace (1\<^sub>m n * A)" using id by auto also have "\ = trace A" using dA by auto finally show ?thesis using dA dM0 dM1 local.id state_sig.trace_measure2_id by blast qed lemma qp_close_under_unitary_operator: fixes U P :: "complex mat" assumes dU: "U \ carrier_mat d d" and u: "unitary U" and qp: "is_quantum_predicate P" shows "is_quantum_predicate (adjoint U * P * U)" unfolding is_quantum_predicate_def proof (auto) have dP: "P \ carrier_mat d d" using qp is_quantum_predicate_def by auto show "adjoint U * P * U \ carrier_mat d d" using dU dP by fastforce have "positive P" using qp is_quantum_predicate_def by auto then show "positive (adjoint U * P * U)" using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dU] dP, simplified adjoint_adjoint] by fastforce have "adjoint U * U = 1\<^sub>m d" apply (subgoal_tac "inverts_mat (adjoint U) U") subgoal unfolding inverts_mat_def using dU by auto using u unfolding unitary_def using inverts_mat_symm[OF dU adjoint_dim[OF dU]] by auto then have u': "adjoint U * 1\<^sub>m d * U = 1\<^sub>m d" using dU by auto have le: "P \\<^sub>L 1\<^sub>m d" using qp is_quantum_predicate_def by auto show "adjoint U * P * U \\<^sub>L 1\<^sub>m d" using lowner_le_keep_under_measurement[OF dU dP one_carrier_mat le] u' by auto qed lemma qps_after_measure_is_qp: assumes m: "measurement d n M " and qpk: "\k. k < n \ is_quantum_predicate (P k)" shows "is_quantum_predicate (matrix_sum d (\k. adjoint (M k) * P k * M k) n)" unfolding is_quantum_predicate_def proof (auto) have dMk: "k < n \ M k \ carrier_mat d d" for k using m measurement_def by auto moreover have dPk: "k < n \ P k \ carrier_mat d d" for k using qpk is_quantum_predicate_def by auto ultimately have dk: "k < n \ adjoint (M k) * P k * M k \ carrier_mat d d" for k by fastforce then show d: "matrix_sum d (\k. adjoint (M k) * P k * M k) n \ carrier_mat d d" using matrix_sum_dim[of n "\k. adjoint (M k) * P k * M k"] by auto have "k < n \ positive (P k)" for k using qpk is_quantum_predicate_def by auto then have "k < n \ positive (adjoint (M k) * P k * M k)" for k using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dMk] dPk, simplified adjoint_adjoint] by fastforce then show "positive (matrix_sum d (\k. adjoint (M k) * P k * M k) n)" using matrix_sum_positive dk by auto have "k < n \ P k \\<^sub>L 1\<^sub>m d" for k using qpk is_quantum_predicate_def by auto then have "k < n \ positive (1\<^sub>m d - P k)" for k using lowner_le_def by auto then have p: "k < n \ positive (adjoint (M k) * (1\<^sub>m d - P k) * M k)" for k using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dMk], simplified adjoint_adjoint, of _ "1\<^sub>m d - P k"] dPk by fastforce { fix k assume k: "k < n" have "adjoint (M k) * (1\<^sub>m d - P k) * M k = adjoint (M k) * M k - adjoint (M k) * P k * M k" apply (mat_assoc d) using dMk dPk k by auto } note split = this have dk': "k < n \ adjoint (M k) * M k - adjoint (M k) * P k * M k \ carrier_mat d d" for k using dMk dPk by fastforce have "k < n \ positive (adjoint (M k) * M k - adjoint (M k) * P k * M k)" for k using p split by auto then have p': "positive (matrix_sum d (\k. adjoint (M k) * M k - adjoint (M k) * P k * M k) n)" using matrix_sum_positive[OF dk', of n id, simplified] by auto have daMMk: "k < n \ adjoint (M k) * M k \ carrier_mat d d" for k using dMk by fastforce have daMPMk: "k < n \ adjoint (M k) * P k * M k \ carrier_mat d d" for k using dMk dPk by fastforce have "matrix_sum d (\k. adjoint (M k) * M k - adjoint (M k) * P k * M k) n = matrix_sum d (\k. adjoint (M k) * M k) n - matrix_sum d (\k. adjoint (M k) * P k * M k) n" using matrix_sum_minus_distrib[OF daMMk daMPMk] by auto also have "\ = 1\<^sub>m d - matrix_sum d (\k. adjoint (M k) * P k * M k) n" using m measurement_def by auto finally have "positive (1\<^sub>m d - matrix_sum d (\k. adjoint (M k) * P k * M k) n)" using p' by auto then show "matrix_sum d (\k. adjoint (M k) * P k * M k) n \\<^sub>L 1\<^sub>m d" using lowner_le_def d by auto qed definition hoare_total_correct :: "complex mat \ com \ complex mat \ bool" ("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where "\\<^sub>t {P} S {Q} \ (\\\density_states. trace (P * \) \ trace (Q * denote S \))" definition hoare_partial_correct :: "complex mat \ com \ complex mat \ bool" ("\\<^sub>p {(1_)}/ (_)/ {(1_)}" 50) where "\\<^sub>p {P} S {Q} \ (\\\density_states. trace (P * \) \ trace (Q * denote S \) + (trace \ - trace (denote S \)))" (* Proposition 6.1 (1) *) lemma total_implies_partial: assumes S: "well_com S" and total: "\\<^sub>t {P} S {Q}" shows "\\<^sub>p {P} S {Q}" proof - have "trace (P * \) \ trace (Q * denote S \) + (trace \ - trace (denote S \))" if \: "\ \ density_states" for \ proof - have "trace (P * \) \ trace (Q * denote S \)" using total hoare_total_correct_def \ by auto moreover have "trace (denote S \) \ trace \" using denote_trace[OF S] \ density_states_def by auto ultimately show ?thesis by auto qed then show ?thesis using hoare_partial_correct_def by auto qed lemma predicate_prob_positive: assumes "0\<^sub>m d d \\<^sub>L P" and "\ \ density_states" shows "0 \ trace (P * \)" proof - have "trace (0\<^sub>m d d * \) \ trace (P * \)" apply (rule lowner_le_traceD) using assms unfolding lowner_le_def density_states_def by auto then show ?thesis using assms(2) density_states_def by auto qed (* Proposition 6.1 (2a) *) lemma total_pre_zero: assumes S: "well_com S" and Q: "is_quantum_predicate Q" shows "\\<^sub>t {0\<^sub>m d d} S {Q}" proof - have "trace (0\<^sub>m d d * \) \ trace (Q * denote S \)" if \: "\ \ density_states" for \ proof - have 1: "trace (0\<^sub>m d d * \) = 0" using \ unfolding density_states_def by auto show ?thesis apply (subst 1) apply (rule predicate_prob_positive) subgoal apply (simp add: lowner_le_def, subgoal_tac "Q - 0\<^sub>m d d = Q") using Q is_quantum_predicate_def[of Q] by auto subgoal using denote_density_states \ S by auto done qed then show ?thesis using hoare_total_correct_def by auto qed (* Proposition 6.1 (2b) *) lemma partial_post_identity: assumes S: "well_com S" and P: "is_quantum_predicate P" shows "\\<^sub>p {P} S {1\<^sub>m d}" proof - have "trace (P * \) \ trace (1\<^sub>m d * denote S \) + (trace \ - trace (denote S \))" if \: "\ \ density_states" for \ proof - have "denote S \ \ carrier_mat d d" using S denote_dim \ density_states_def by auto then have "trace (1\<^sub>m d * denote S \) = trace (denote S \)" by auto moreover have "trace (P * \) \ trace (1\<^sub>m d * \)" apply (rule lowner_le_traceD) using \ unfolding density_states_def apply auto using P is_quantum_predicate_def by auto ultimately show ?thesis using density_states_def that by auto qed then show ?thesis using hoare_partial_correct_def by auto qed subsection \Weakest liberal preconditions\ definition is_weakest_liberal_precondition :: "complex mat \ com \ complex mat \ bool" where "is_weakest_liberal_precondition W S P \ is_quantum_predicate W \ \\<^sub>p {W} S {P} \ (\Q. is_quantum_predicate Q \ \\<^sub>p {Q} S {P} \ Q \\<^sub>L W)" definition wlp_measure :: "nat \ (nat \ complex mat) \ ((complex mat \ complex mat) list) \ complex mat \ complex mat" where "wlp_measure n M WS P = matrix_sum d (\k. adjoint (M k) * ((WS!k) P) * (M k)) n" fun wlp_while_n :: "complex mat \ complex mat \ (complex mat \ complex mat) \ nat \ complex mat \ complex mat" where "wlp_while_n M0 M1 WS 0 P = 1\<^sub>m d" | "wlp_while_n M0 M1 WS (Suc n) P = adjoint M0 * P * M0 + adjoint M1 * (WS (wlp_while_n M0 M1 WS n P)) * M1" lemma measurement2_leq_one_mat: assumes dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" and leP: "P \\<^sub>L 1\<^sub>m d" and leQ: "Q \\<^sub>L 1\<^sub>m d" and m: "measurement d 2 M" shows "(adjoint (M 0) * P * (M 0) + adjoint (M 1) * Q * (M 1)) \\<^sub>L 1\<^sub>m d" proof - define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using m M0_def M1_def measurement_def by auto have "adjoint M1 * Q * M1 \\<^sub>L adjoint M1 * 1\<^sub>m d * M1" using lowner_le_keep_under_measurement[OF dM1 dQ _ leQ] by auto then have le1: "adjoint M1 * Q * M1 \\<^sub>L adjoint M1 * M1" using dM1 dQ by fastforce have "adjoint M0 * P * M0 \\<^sub>L adjoint M0 * 1\<^sub>m d * M0" using lowner_le_keep_under_measurement[OF dM0 dP _ leP] by auto then have le0: "adjoint M0 * P * M0 \\<^sub>L adjoint M0 * M0" using dM0 dP by fastforce have "adjoint M0 * P * M0 + adjoint M1 * Q * M1 \\<^sub>L adjoint M0 * M0 + adjoint M1 * M1" apply (rule lowner_le_add[of "adjoint M0 * P * M0" d "adjoint M0 * M0" "adjoint M1 * Q * M1" "adjoint M1 * M1"]) using dM0 dP dM1 dQ le0 le1 by auto also have "\ = 1\<^sub>m d" using m M0_def M1_def measurement_id2 by auto finally show "adjoint M0 * P * M0 + adjoint M1 * Q * M1 \\<^sub>L 1\<^sub>m d". qed lemma wlp_while_n_close: assumes close: "\P. is_quantum_predicate P \ is_quantum_predicate (WS P)" and m: "measurement d 2 M" and qpP: "is_quantum_predicate P" shows "is_quantum_predicate (wlp_while_n (M 0) (M 1) WS k P)" proof (induct k) case 0 then show ?case unfolding wlp_while_n.simps is_quantum_predicate_def using positive_one[of d] lowner_le_refl[of "1\<^sub>m d"] by fastforce next case (Suc k) define M0 where "M0 = M 0" define M1 where "M1 = M 1" define W where "W k = wlp_while_n M0 M1 WS k P" for k show ?case unfolding wlp_while_n.simps is_quantum_predicate_def proof (fold M0_def M1_def, fold W_def, auto) have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using m M0_def M1_def measurement_def by auto have dP: "P \ carrier_mat d d" using qpP is_quantum_predicate_def by auto have qpWk: "is_quantum_predicate (W k)" using Suc M0_def M1_def W_def by auto then have qpWWk: "is_quantum_predicate (WS (W k))" using close by auto from qpWk have dWk: "W k \ carrier_mat d d" using is_quantum_predicate_def by auto from qpWWk have dWWk: "WS (W k) \ carrier_mat d d" using is_quantum_predicate_def by auto show "adjoint M0 * P * M0 + adjoint M1 * WS (W k) * M1 \ carrier_mat d d" using dM0 dP dM1 dWWk by auto have pP: "positive P" using qpP is_quantum_predicate_def by auto then have pM0P: "positive (adjoint M0 * P * M0)" using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dM0]] dM0 dP adjoint_adjoint[of M0] by auto have pWWk: "positive (WS (W k))" using qpWWk is_quantum_predicate_def by auto then have pM1WWk: "positive (adjoint M1 * WS (W k) * M1)" using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dM1]] dM1 dWWk adjoint_adjoint[of M1] by auto then show "positive (adjoint M0 * P * M0 + adjoint M1 * WS (W k) * M1)" using positive_add[OF pM0P pM1WWk] dM0 dP dM1 dWWk by fastforce have leWWk: "WS (W k) \\<^sub>L 1\<^sub>m d" using qpWWk is_quantum_predicate_def by auto have leP: "P \\<^sub>L 1\<^sub>m d" using qpP is_quantum_predicate_def by auto show "(adjoint M0 * P * M0 + adjoint M1 * WS (W k) * M1) \\<^sub>L 1\<^sub>m d " using measurement2_leq_one_mat[OF dP dWWk leP leWWk m] M0_def M1_def by auto qed qed lemma wlp_while_n_mono: assumes "\P. is_quantum_predicate P \ is_quantum_predicate (WS P)" and "\P Q. is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ WS P \\<^sub>L WS Q" and "measurement d 2 M" and "is_quantum_predicate P" and "is_quantum_predicate Q" and "P \\<^sub>L Q" shows "(wlp_while_n (M 0) (M 1) WS k P) \\<^sub>L (wlp_while_n (M 0) (M 1) WS k Q)" proof (induct k) case 0 then show ?case unfolding wlp_while_n.simps using lowner_le_refl[of "1\<^sub>m d"] by fastforce next case (Suc k) define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using assms M0_def M1_def measurement_def by auto define W where "W P k = wlp_while_n M0 M1 WS k P" for k P have dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" using assms is_quantum_predicate_def by auto have eq1: "W P (Suc k) = adjoint M0 * P * M0 + adjoint M1 * (WS (W P k)) * M1" unfolding W_def wlp_while_n.simps by auto have eq2: "W Q (Suc k) = adjoint M0 * Q * M0 + adjoint M1 * (WS (W Q k)) * M1" unfolding W_def wlp_while_n.simps by auto have le1: "adjoint M0 * P * M0 \\<^sub>L adjoint M0 * Q * M0" using lowner_le_keep_under_measurement dM0 dP dQ assms by auto have leWk: "(W P k) \\<^sub>L (W Q k)" unfolding W_def M0_def M1_def using Suc by auto have qpWPk: "is_quantum_predicate (W P k)" unfolding W_def M0_def M1_def using wlp_while_n_close assms by auto then have "is_quantum_predicate (WS (W P k))" unfolding W_def M0_def M1_def using assms by auto then have dWWPk: "(WS (W P k)) \ carrier_mat d d" using is_quantum_predicate_def by auto have qpWQk: "is_quantum_predicate (W Q k)" unfolding W_def M0_def M1_def using wlp_while_n_close assms by auto then have "is_quantum_predicate (WS (W Q k))" unfolding W_def M0_def M1_def using assms by auto then have dWWQk: "(WS (W Q k)) \ carrier_mat d d" using is_quantum_predicate_def by auto have "(WS (W P k)) \\<^sub>L (WS (W Q k))" using qpWPk qpWQk leWk assms by auto then have le2: "adjoint M1 * (WS (W P k)) * M1 \\<^sub>L adjoint M1 * (WS (W Q k)) * M1" using lowner_le_keep_under_measurement dM1 dWWPk dWWQk by auto have "(adjoint M0 * P * M0 + adjoint M1 * (WS (W P k)) * M1) \\<^sub>L (adjoint M0 * Q * M0 + adjoint M1 * (WS (W Q k)) * M1)" using lowner_le_add[OF _ _ _ _ le1 le2] dM0 dP dM1 dQ dWWPk dWWQk le1 le2 by fastforce then have "W P (Suc k) \\<^sub>L W Q (Suc k)" using eq1 eq2 by auto then show ?case unfolding W_def M0_def M1_def by auto qed definition wlp_while :: "complex mat \ complex mat \ (complex mat \ complex mat) \ complex mat \ complex mat" where "wlp_while M0 M1 WS P = (THE Q. limit_mat (\n. wlp_while_n M0 M1 WS n P) Q d)" lemma wlp_while_exists: assumes "\P. is_quantum_predicate P \ is_quantum_predicate (WS P)" and "\P Q. is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ WS P \\<^sub>L WS Q" and m: "measurement d 2 M" and qpP: "is_quantum_predicate P" shows "is_quantum_predicate (wlp_while (M 0) (M 1) WS P) \ (\n. (wlp_while (M 0) (M 1) WS P) \\<^sub>L (wlp_while_n (M 0) (M 1) WS n P)) \ (\W'. (\n. W' \\<^sub>L (wlp_while_n (M 0) (M 1) WS n P)) \ W' \\<^sub>L (wlp_while (M 0) (M 1) WS P)) \ limit_mat (\n. wlp_while_n (M 0) (M 1) WS n P) (wlp_while (M 0) (M 1) WS P) d" proof (auto) define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using assms M0_def M1_def measurement_def by auto define W where "W k = wlp_while_n M0 M1 WS k P" for k have leP: "P \\<^sub>L 1\<^sub>m d" and dP: "P \ carrier_mat d d" and pP: "positive P" using qpP is_quantum_predicate_def by auto have pM0P: "positive (adjoint M0 * P * M0)" using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dM0]] adjoint_adjoint[of "M0"] dP pP by auto have le_qp: "W (Suc k) \\<^sub>L W k \ is_quantum_predicate (W k)" for k proof (induct k) case 0 have "is_quantum_predicate (1\<^sub>m d)" unfolding is_quantum_predicate_def using positive_one lowner_le_refl[of "1\<^sub>m d"] by fastforce then have "is_quantum_predicate (WS (1\<^sub>m d))" using assms by auto then have "(WS (1\<^sub>m d)) \ carrier_mat d d" and "(WS (1\<^sub>m d)) \\<^sub>L 1\<^sub>m d" using is_quantum_predicate_def by auto then have "W 1 \\<^sub>L W 0" unfolding W_def wlp_while_n.simps M0_def M1_def using measurement2_leq_one_mat[OF dP _ leP _ m] by auto moreover have "is_quantum_predicate (W 0)" unfolding W_def wlp_while_n.simps is_quantum_predicate_def using positive_one lowner_le_refl[of "1\<^sub>m d"] by fastforce ultimately show ?case by auto next case (Suc k) then have leWSk: "W (Suc k) \\<^sub>L W k" and qpWk: "is_quantum_predicate (W k)" by auto then have "is_quantum_predicate (WS (W k))" using assms by auto then have dWWk: "WS (W k) \ carrier_mat d d" and leWWk1: "(WS (W k)) \\<^sub>L 1\<^sub>m d" and pWWk: "positive (WS (W k))" using is_quantum_predicate_def by auto then have leWSk1: "W (Suc k) \\<^sub>L 1\<^sub>m d" using measurement2_leq_one_mat[OF dP dWWk leP leWWk1 m] unfolding W_def wlp_while_n.simps M0_def M1_def by auto then have dWSk: "W (Suc k) \ carrier_mat d d" using lowner_le_def by fastforce have pM1WWk: "positive (adjoint M1 * (WS (W k)) * M1)" using positive_close_under_left_right_mult_adjoint[OF adjoint_dim[OF dM1] dWWk pWWk] adjoint_adjoint[of "M1"] by auto have pWSk: "positive (W (Suc k))" unfolding W_def wlp_while_n.simps apply (fold W_def) using positive_add[OF pM0P pM1WWk] dM0 dP dM1 by fastforce have qpWSk:"is_quantum_predicate (W (Suc k))" unfolding is_quantum_predicate_def using dWSk pWSk leWSk1 by auto then have qpWWSk: "is_quantum_predicate (WS (W (Suc k)))" using assms(1) by auto then have dWWSk: "(WS (W (Suc k))) \ carrier_mat d d" using is_quantum_predicate_def by auto have "WS (W (Suc k)) \\<^sub>L WS (W k)" using assms(2)[OF qpWSk qpWk] leWSk by auto then have "adjoint M1 * WS (W (Suc k)) * M1 \\<^sub>L adjoint M1 * WS (W k) * M1" using lowner_le_keep_under_measurement[OF dM1 dWWSk dWWk] by auto then have "(adjoint M0 * P * M0 + adjoint M1 * WS (W (Suc k)) * M1) \\<^sub>L (adjoint M0 * P * M0 + adjoint M1 * WS (W k) * M1)" using lowner_le_add[of _ d _ "adjoint M1 * WS (W (Suc k)) * M1" "adjoint M1 * WS (W k) * M1", OF _ _ _ _ lowner_le_refl[of "adjoint M0 * P * M0"]] dM0 dM1 dP dWWSk dWWk by fastforce then have "W (Suc (Suc k)) \\<^sub>L W (Suc k)" unfolding W_def wlp_while_n.simps by auto with qpWSk show ?case by auto qed then have dWk: "W k \ carrier_mat d d" for k using is_quantum_predicate_def by auto then have dmWk: "- W k \ carrier_mat d d" for k by auto have incmWk: "- (W k) \\<^sub>L - (W (Suc k))" for k using lowner_le_swap[of "W (Suc k)" d "W k"] dWk le_qp by auto have pWk: "positive (W k)" for k using le_qp is_quantum_predicate_def by auto have ubmWk: "- W k \\<^sub>L 0\<^sub>m d d" for k proof - have "0\<^sub>m d d \\<^sub>L W k" for k using zero_lowner_le_positiveI dWk pWk by auto then have "- W k \\<^sub>L - 0\<^sub>m d d" for k using lowner_le_swap[of "0\<^sub>m d d" d "W k"] dWk by auto moreover have "(- 0\<^sub>m d d :: complex mat) = (0\<^sub>m d d)" by auto ultimately show ?thesis by auto qed have "\B. lowner_is_lub (\k. - W k) B \ limit_mat (\k. - W k) B d" using mat_inc_seq_lub[of "\k. - W k" d "0\<^sub>m d d"] dmWk incmWk ubmWk by auto then obtain B where lubB: "lowner_is_lub (\k. - W k) B" and limB: "limit_mat (\k. - W k) B d" by auto then have dB: "B \ carrier_mat d d" using limit_mat_def by auto define A where "A = - B" then have dA: "A \ carrier_mat d d" using dB by auto have "limit_mat (\k. (-1) \\<^sub>m (- W k)) (-1 \\<^sub>m B) d" using limit_mat_scale[OF limB] by auto moreover have "W k = -1 \\<^sub>m (- W k)" for k using dWk by auto moreover have "-1 \\<^sub>m B = - B" by auto ultimately have limA: "limit_mat W A d" using A_def by auto moreover have "(limit_mat W A' d \ A' = A)" for A' using limit_mat_unique[of W A d] limA by auto ultimately have eqA: "(wlp_while (M 0) (M 1) WS P) = A" unfolding wlp_while_def W_def M0_def M1_def using the_equality[of "\X. limit_mat (\n. wlp_while_n (M 0) (M 1) WS n P) X d" A] by fastforce show "limit_mat (\n. wlp_while_n (M 0) (M (Suc 0)) WS n P) (wlp_while (M 0) (M (Suc 0)) WS P) d" using limA eqA unfolding W_def M0_def M1_def by auto have "- W k \\<^sub>L B" for k using lubB lowner_is_lub_def by auto then have glbA: "A \\<^sub>L W k" for k unfolding A_def using lowner_le_swap[of "- W k" d] dB dWk by fastforce then show "\n. wlp_while (M 0) (M (Suc 0)) WS P \\<^sub>L wlp_while_n (M 0) (M (Suc 0)) WS n P" using eqA unfolding W_def M0_def M1_def by auto have "W k \\<^sub>L 1\<^sub>m d" for k using le_qp unfolding is_quantum_predicate_def by auto then have "positive (1\<^sub>m d - W k)" for k using lowner_le_def by auto moreover have "limit_mat (\k. 1\<^sub>m d - W k) (1\<^sub>m d - A) d" using mat_minus_limit limA by auto ultimately have "positive (1\<^sub>m d - A)" using pos_mat_lim_is_pos by auto then have leA1: "A \\<^sub>L 1\<^sub>m d" using dA lowner_le_def by auto have pA: "positive A" using pos_mat_lim_is_pos limA pWk by auto show "is_quantum_predicate (wlp_while (M 0) (M (Suc 0)) WS P)" unfolding is_quantum_predicate_def using pA dA leA1 eqA by auto { fix W' assume asmW': "\k. W' \\<^sub>L W k" then have dW': "W' \ carrier_mat d d" unfolding lowner_le_def using carrier_matD[OF dWk] by auto then have "- W k \\<^sub>L - W'" for k using lowner_le_swap dWk asmW' by auto then have "B \\<^sub>L - W'" using lubB unfolding lowner_is_lub_def by auto then have "W' \\<^sub>L A" unfolding A_def using lowner_le_swap[of "B" d "- W'"] dB dW' by auto then have "W' \\<^sub>L wlp_while (M 0) (M 1) WS P" using eqA by auto } then show "\W'. \n. W' \\<^sub>L wlp_while_n (M 0) (M (Suc 0)) WS n P \ W' \\<^sub>L wlp_while (M 0) (M (Suc 0)) WS P" unfolding W_def M0_def M1_def by auto qed lemma wlp_while_mono: assumes "\P. is_quantum_predicate P \ is_quantum_predicate (WS P)" and "\P Q. is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ WS P \\<^sub>L WS Q" and "measurement d 2 M" and "is_quantum_predicate P" and "is_quantum_predicate Q" and "P \\<^sub>L Q" shows "wlp_while (M 0) (M 1) WS P \\<^sub>L wlp_while (M 0) (M 1) WS Q" proof - define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using assms M0_def M1_def measurement_def by auto define Wn where "Wn P k = wlp_while_n M0 M1 WS k P" for P k define W where "W P = wlp_while M0 M1 WS P" for P have lePQk: "Wn P k \\<^sub>L Wn Q k" for k unfolding Wn_def M0_def M1_def using wlp_while_n_mono assms by auto have "is_quantum_predicate (Wn P k)" for k unfolding Wn_def M0_def M1_def using wlp_while_n_close assms by auto then have dWPk: "Wn P k \ carrier_mat d d" for k using is_quantum_predicate_def by auto have "is_quantum_predicate (Wn Q k)" for k unfolding Wn_def M0_def M1_def using wlp_while_n_close assms by auto then have dWQk:"Wn Q k \ carrier_mat d d" for k using is_quantum_predicate_def by auto have "is_quantum_predicate (W P)" and lePk: "(W P) \\<^sub>L (Wn P k)" and "limit_mat (Wn P) (W P) d" for k unfolding W_def Wn_def M0_def M1_def using wlp_while_exists assms by auto then have dWP: "W P \ carrier_mat d d" using is_quantum_predicate_def by auto have "is_quantum_predicate (W Q)" and "(W Q) \\<^sub>L (Wn Q k)" and glb:"(\k. W' \\<^sub>L (Wn Q k)) \ W' \\<^sub>L (W Q)" and "limit_mat (Wn Q) (W Q) d" for k W' unfolding W_def Wn_def M0_def M1_def using wlp_while_exists assms by auto have "W P \\<^sub>L Wn Q k" for k using lowner_le_trans[of "W P" d "Wn P k" "Wn Q k"] lePk lePQk dWPk dWQk dWP by auto then show "W P \\<^sub>L W Q" using glb by auto qed fun wlp :: "com \ complex mat \ complex mat" where "wlp SKIP P = P" | "wlp (Utrans U) P = adjoint U * P * U" | "wlp (Seq S1 S2) P = wlp S1 (wlp S2 P)" | "wlp (Measure n M S) P = wlp_measure n M (map wlp S) P" | "wlp (While M S) P = wlp_while (M 0) (M 1) (wlp S) P" lemma wlp_measure_expand_m: assumes m: "m \ n" and wc: "well_com (Measure n M S)" shows "wlp (Measure m M S) P = matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * (M k)) m" unfolding wlp.simps wlp_measure_def proof - have "k < m \ map wlp S ! k = wlp (S!k)" for k using wc m by auto then have "k < m \ (map wlp S ! k) P = wlp (S!k) P" for k by auto then show "matrix_sum d (\k. adjoint (M k) * ((map wlp S ! k) P) * (M k)) m = matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * (M k)) m" using matrix_sum_cong[of m "\k. adjoint (M k) * ((map wlp S ! k) P) * (M k)" "\k. adjoint (M k) * (wlp (S!k) P) * (M k)"] by auto qed lemma wlp_measure_expand: assumes wc: "well_com (Measure n M S)" shows "wlp (Measure n M S) P = matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * (M k)) n" using wlp_measure_expand_m[OF Nat.le_refl[of n]] wc by auto lemma wlp_mono_and_close: shows "well_com S \ is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ is_quantum_predicate (wlp S P) \ wlp S P \\<^sub>L wlp S Q" proof (induct S arbitrary: P Q) case SKIP then show ?case by auto next case (Utrans U) then have dU: "U \ carrier_mat d d" and u: "unitary U" and qp: "is_quantum_predicate P" and le: "P \\<^sub>L Q" and dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" using is_quantum_predicate_def by auto then have qp': "is_quantum_predicate (wlp (Utrans U) P)" using qp_close_under_unitary_operator by auto moreover have "adjoint U * P * U \\<^sub>L adjoint U * Q * U" using lowner_le_keep_under_measurement[OF dU dP dQ le] by auto ultimately show ?case by auto next case (Seq S1 S2) then have qpP: "is_quantum_predicate P" and qpQ: "is_quantum_predicate Q" and wc1: "well_com S1" and wc2: "well_com S2" and dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" and le: "P \\<^sub>L Q"using is_quantum_predicate_def by auto have qpP2: "is_quantum_predicate (wlp S2 P)" using Seq qpP wc2 by auto have qpQ2: "is_quantum_predicate (wlp S2 Q)" using Seq(2)[OF wc2 qpQ qpQ] lowner_le_refl dQ by blast have qpP1: "is_quantum_predicate (wlp S1 (wlp S2 P))" using Seq(1)[OF wc1 qpP2 qpP2] qpP2 is_quantum_predicate_def[of "wlp S2 P"] lowner_le_refl by auto have "wlp S2 P \\<^sub>L wlp S2 Q" using Seq(2) wc2 qpP qpQ le by auto then have "wlp S1 (wlp S2 P) \\<^sub>L wlp S1 (wlp S2 Q)" using Seq(1) wc1 qpP2 qpQ2 by auto then show ?case using qpP1 by auto next case (Measure n M S) then have wc: "well_com (Measure n M S)" and wck: "k < n \ well_com (S!k)" and l: "length S = n" and m: "measurement d n M" and le: "P \\<^sub>L Q" and qpP: "is_quantum_predicate P" and dP: "P \ carrier_mat d d" and qpQ: "is_quantum_predicate Q" and dQ: "Q \ carrier_mat d d" for k using measure_well_com is_quantum_predicate_def by auto have dMk: "k < n \ M k \ carrier_mat d d" for k using m measurement_def by auto have set: "k < n \ S!k \ set S" for k using l by auto have qpk: "k < n \ is_quantum_predicate (wlp (S!k) P)" for k using Measure(1)[OF set wck qpP qpP] lowner_le_refl[of P] dP by auto then have dWkP: "k < n \ wlp (S!k) P \ carrier_mat d d" for k using is_quantum_predicate_def by auto then have dMkP: "k < n \ adjoint (M k) * (wlp (S!k) P) * (M k) \ carrier_mat d d" for k using dMk by fastforce have "k < n \ is_quantum_predicate (wlp (S!k) Q)" for k using Measure(1)[OF set wck qpQ qpQ] lowner_le_refl[of Q] dQ by auto then have dWkQ: "k < n \ wlp (S!k) Q \ carrier_mat d d" for k using is_quantum_predicate_def by auto then have dMkQ: "k < n \ adjoint (M k) * (wlp (S!k) Q) * (M k) \ carrier_mat d d" for k using dMk by fastforce have "k < n \ wlp (S!k) P \\<^sub>L wlp (S!k) Q" for k using Measure(1)[OF set wck qpP qpQ le] by auto then have "k < n \ adjoint (M k) * (wlp (S!k) P) * (M k) \\<^sub>L adjoint (M k) * (wlp (S!k) Q) * (M k)" for k using lowner_le_keep_under_measurement[OF dMk dWkP dWkQ] by auto then have le': "wlp (Measure n M S) P \\<^sub>L wlp (Measure n M S) Q" unfolding wlp_measure_expand[OF wc] using lowner_le_matrix_sum dMkP dMkQ by auto have qp': "is_quantum_predicate (wlp (Measure n M S) P)" unfolding wlp_measure_expand[OF wc] using qps_after_measure_is_qp[OF m] qpk by auto show ?case using le' qp' by auto next case (While M S) then have m: "measurement d 2 M" and wcs: "well_com S" and qpP: "is_quantum_predicate P" by auto have closeWS: "is_quantum_predicate P \ is_quantum_predicate (wlp S P)" for P proof - assume asm: "is_quantum_predicate P" then have dP: "P \ carrier_mat d d" using is_quantum_predicate_def by auto then show "is_quantum_predicate (wlp S P)" using While(1)[OF wcs asm asm lowner_le_refl] dP by auto qed have monoWS: "is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ wlp S P \\<^sub>L wlp S Q" for P Q using While(1)[OF wcs] by auto have "is_quantum_predicate (wlp (While M S) P)" using wlp_while_exists[of "wlp S" M P] closeWS monoWS m qpP by auto moreover have "wlp (While M S) P \\<^sub>L wlp (While M S) Q" using wlp_while_mono[of "wlp S" M P Q] closeWS monoWS m While by auto ultimately show ?case by auto qed lemma wlp_close: assumes wc: "well_com S" and qp: "is_quantum_predicate P" shows "is_quantum_predicate (wlp S P)" using wlp_mono_and_close[OF wc qp qp] is_quantum_predicate_def[of P] qp lowner_le_refl by auto lemma wlp_soundness: "well_com S \ (\P. (is_quantum_predicate P \ (\\ \ density_states. trace (wlp S P * \) = trace (P * (denote S \)) + trace \ - trace (denote S \))))" proof (induct S) case SKIP then show ?case by auto next case (Utrans U) then have dU: "U \ carrier_mat d d" and u: "unitary U" and wc: "well_com (Utrans U)" and qp: "is_quantum_predicate P" and dP: "P \ carrier_mat d d" using is_quantum_predicate_def by auto have qp': "is_quantum_predicate (wlp (Utrans U) P)" using wlp_close[OF wc qp] by auto have eq1: "trace (adjoint U * P * U * \) = trace (P * (U * \ * adjoint U))" if dr: "\ \ carrier_mat d d" for \ using dr dP apply (mat_assoc d) using wc by auto have eq2: "trace (U * \ * adjoint U) = trace \" if dr: "\ \ carrier_mat d d" for \ using unitary_operator_keep_trace[OF adjoint_dim[OF dU] dr unitary_adjoint[OF dU u]] adjoint_adjoint[of U] by auto show ?case using qp' eq1 eq2 density_states_def by auto next case (Seq S1 S2) then have qp: "is_quantum_predicate P" and wc1: "well_com S1" and wc2: "well_com S2" by auto then have qp2: "is_quantum_predicate (wlp S2 P)" using wlp_close by auto then have qp1: "is_quantum_predicate (wlp S1 (wlp S2 P))" using wlp_close wc1 by auto have eq1: "trace (wlp S2 P * \) = trace (P * denote S2 \) + trace \ - trace (denote S2 \)" if ds: "\ \ density_states" for \ using Seq(2) wc2 qp ds by auto have eq2: "trace (wlp S1 (wlp S2 P) * \) = trace ((wlp S2 P) * denote S1 \) + trace \ - trace (denote S1 \)" if ds: "\ \ density_states" for \ using Seq(1) wc1 qp2 ds by auto have eq3: "trace (wlp S1 (wlp S2 P) * \) = trace (P * (denote S2 (denote S1 \))) + trace \ - trace (denote S2 (denote S1 \))" if ds: "\ \ density_states" for \ proof - have "denote S1 \ \ density_states" using ds denote_density_states wc1 by auto then have "trace ((wlp S2 P) * denote S1 \) = trace (P * denote S2 (denote S1 \)) + trace (denote S1 \) - trace (denote S2 (denote S1 \))" using eq1 by auto then show "trace (wlp S1 (wlp S2 P) * \) = trace (P * (denote S2 (denote S1 \))) + trace \ - trace (denote S2 (denote S1 \))" using eq2 ds by auto qed then show ?case using qp1 by auto next case (Measure n M S) then have wc: "well_com (Measure n M S)" and wck: "k < n \ well_com (S!k)" and qpP: "is_quantum_predicate P" and dP: "P \ carrier_mat d d" and qpWk: "k < n \ is_quantum_predicate (wlp (S!k) P)" and dWk: "k < n \ (wlp (S!k) P) \ carrier_mat d d" and c: "k < n \ \ \ density_states \ trace (wlp (S!k) P * \) = trace (P * denote (S!k) \) + trace \ - trace (denote (S!k) \)" and m: "measurement d n M" and aMMkleq: "k < n \ adjoint (M k) * M k \\<^sub>L 1\<^sub>m d" and dMk: "k < n \ M k \ carrier_mat d d" for k \ using is_quantum_predicate_def measurement_def measure_well_com measurement_le_one_mat wlp_close by auto { fix \ assume \: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto have dsr: "k < n \ (M k) * \ * adjoint (M k) \ density_states" for k unfolding density_states_def using dMk pdo_close_under_measurement[OF dMk dr pdor aMMkleq] dr by fastforce then have leqk: "k < n \ trace (wlp (S!k) P * ((M k) * \ * adjoint (M k))) = trace (P * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))" for k using c by auto have "k < n \ M k * \ * adjoint (M k) \ carrier_mat d d" for k using dMk dr by fastforce then have dsMrk: "k < n \ matrix_sum d (\k. (M k * \ * adjoint (M k))) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. (M k * \ * adjoint (M k))" d] by fastforce have "k < n \ adjoint (M k) * (wlp (S!k) P) * M k \ carrier_mat d d" for k using dMk by fastforce then have dsMW: "k < n \ matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. adjoint (M k) * (wlp (S!k) P) * M k" d] by fastforce have dSMrk: "k < n \ denote (S ! k) (M k * \ * adjoint (M k)) \ carrier_mat d d" for k using denote_dim[OF wck, of k "M k * \ * adjoint (M k)"] dsr density_states_def by fastforce have dsSMrk: "k < n \ matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. denote (S ! k) (M k * \ * adjoint (M k))" d, OF dSMrk] by fastforce have "k \ n \ trace (matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) k * \) = trace (P * (denote (Measure k M S) \)) + (trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) k) - trace (denote (Measure k M S) \))" for k unfolding denote_measure_expand[OF _ wc] proof (induct k) case 0 then show ?case unfolding matrix_sum.simps using dP dr by auto next case (Suc k) then have k: "k < n" by auto have eq1: "trace (matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) (Suc k) * \) = trace (adjoint (M k) * (wlp (S!k) P) * M k * \) + trace (matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) k * \)" unfolding matrix_sum.simps using dMk[OF k] dWk[OF k] dr dsMW[OF k] by (mat_assoc d) have "trace (adjoint (M k) * (wlp (S!k) P) * M k * \) = trace ((wlp (S!k) P) * (M k * \ * adjoint (M k)))" using dMk[OF k] dWk[OF k] dr by (mat_assoc d) also have "\ = trace (P * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))" using leqk k by auto finally have eq2: "trace (adjoint (M k) * (wlp (S!k) P) * M k * \) = trace (P * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))" . have eq3: "trace (P * matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) (Suc k)) = trace (P * (denote (S!k) (M k * \ * adjoint (M k)))) + trace (P * matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k)" unfolding matrix_sum.simps using dP dSMrk[OF k] dsSMrk[OF k] by (mat_assoc d) have eq4: "trace (denote (S ! k) (M k * \ * adjoint (M k)) + matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k) = trace (denote (S ! k) (M k * \ * adjoint (M k))) + trace (matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k)" using dSMrk[OF k] dsSMrk[OF k] by (mat_assoc d) show ?case apply (subst eq1) apply (subst eq3) apply (simp del: less_eq_complex_def) apply (subst trace_add_linear[of "M k * \ * adjoint (M k)" d "matrix_sum d (\k. M k * \ * adjoint (M k)) k"]) apply (simp add: dMk adjoint_dim[OF dMk] dr mult_carrier_mat[of _ d d _ d] k) apply (simp add: dsMrk k) apply (subst eq4) apply (insert eq2 Suc(1) k, fastforce) done qed then have leq: "trace (matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) n * \) = trace (P * denote (Measure n M S) \) + (trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) n) - trace (denote (Measure n M S) \))" by auto have "trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) n) = trace \" using trace_measurement m dr by auto with leq have "trace (matrix_sum d (\k. adjoint (M k) * (wlp (S!k) P) * M k) n * \) = trace (P * denote (Measure n M S) \) + (trace \ - trace (denote (Measure n M S) \))" unfolding denote_measure_def by auto } then show ?case unfolding wlp_measure_expand[OF wc] by auto next case (While M S) then have qpP: "is_quantum_predicate P" and dP: "P \ carrier_mat d d" and wcS: "well_com S" and m: "measurement d 2 M" and wc: "well_com (While M S)" using is_quantum_predicate_def by auto define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using m measurement_def M0_def M1_def by auto have leM1: "adjoint M1 * M1 \\<^sub>L 1\<^sub>m d" using measurement_le_one_mat m M1_def by auto define W where "W k = wlp_while_n M0 M1 (wlp S) k P" for k define DS where "DS = denote S" define D0 where "D0 = denote_while_n M0 M1 DS" define D1 where "D1 = denote_while_n_comp M0 M1 DS" define D where "D = denote_while_n_iter M0 M1 DS" have eqk: "\ \ density_states \ trace ((W k) * \) = (\k=0..))) + trace \ - (\k=0..))" for k \ proof (induct k arbitrary: \) case 0 then have dsr: "\ \ density_states" by auto show ?case unfolding W_def wlp_while_n.simps using dsr density_states_def by auto next case (Suc k) then have dsr: "\ \ density_states" and dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto then have dsM1r: "M1 * \ * adjoint M1 \ density_states" unfolding density_states_def using pdo_close_under_measurement[OF dM1 dr pdor leM1] dr dM1 by auto then have dsDSM1r: "(DS (M1 * \ * adjoint M1)) \ density_states" unfolding density_states_def DS_def using denote_dim[OF wcS] denote_partial_density_operator[OF wcS] dsM1r by auto have qpWk: "is_quantum_predicate (W k)" using wlp_while_n_close[OF _ m qpP, folded M0_def M1_def, of "wlp S", folded W_def] wlp_close[OF wcS] by auto then have "is_quantum_predicate (wlp S (W k))" using wlp_close[OF wcS] by auto then have dWWk: "wlp S (W k) \ carrier_mat d d" using is_quantum_predicate_def by auto have "trace (P * (M0 * \ * adjoint M0)) + (\k=0.. * adjoint M1))))) = trace (P * (D0 0 \)) + (\k=0..)))" unfolding D0_def by auto also have "\ = trace (P * (D0 0 \)) + (\k=1..<(Suc k). trace (P * (D0 k \)))" using sum.shift_bounds_Suc_ivl[symmetric, of "\k. trace (P * (D0 k \))"] by auto also have "\ = (\k=0..<(Suc k). trace (P * (D0 k \)))" using sum.atLeast_Suc_lessThan[of 0 "Suc k" "\k. trace (P * (D0 k \))"] by auto finally have eq1: "trace (P * (M0 * \ * adjoint M0)) + (\k=0.. * adjoint M1))))) = (\k=0..<(Suc k). trace (P * (D0 k \)))". have eq2: "trace (M1 * \ * adjoint M1) = trace \ - trace (M0 * \ * adjoint M0)" unfolding M0_def M1_def using m trace_measurement2[OF m dr] dr by (simp add: algebra_simps) have "trace (M0 * \ * adjoint M0) + (\k=0.. * adjoint M1)))) = trace (D0 0 \) + (\k=0..))" unfolding D0_def by auto also have "\ = trace (D0 0 \) + (\k=1..<(Suc k). trace (D0 k \))" using sum.shift_bounds_Suc_ivl[symmetric, of "\k. trace (D0 k \)"] by auto also have "\ = (\k=0..<(Suc k). trace (D0 k \))" using sum.atLeast_Suc_lessThan[of 0 "Suc k" "\k. trace (D0 k \)"] by auto finally have eq3: "trace (M0 * \ * adjoint M0) + (\k=0.. * adjoint M1)))) = (\k=0..<(Suc k). trace (D0 k \))". then have "trace (M1 * \ * adjoint M1) - (\k=0.. * adjoint M1)))) = trace \ - (trace (M0 * \ * adjoint M0) + (\k=0.. * adjoint M1)))))" by (simp add: algebra_simps eq2) then have eq4: "trace (M1 * \ * adjoint M1) - (\k=0.. * adjoint M1)))) = trace \ - (\k=0..<(Suc k). trace (D0 k \))" by (simp add: eq3) have "trace ((W (Suc k)) * \) = trace (P * (M0 * \ * adjoint M0)) + trace ((wlp S (W k)) * (M1 * \ * adjoint M1))" unfolding W_def wlp_while_n.simps apply (fold W_def) using dM0 dP dM1 dWWk dr by (mat_assoc d) also have "\ = trace (P * (M0 * \ * adjoint M0)) + trace ((W k) * (DS (M1 * \ * adjoint M1))) + trace (M1 * \ * adjoint M1) - trace (DS (M1 * \ * adjoint M1))" using While(1)[OF wcS, of "W k"] qpWk dsM1r DS_def by auto also have "\ = trace (P * (M0 * \ * adjoint M0)) + (\k=0.. * adjoint M1))))) + trace (DS (M1 * \ * adjoint M1)) - (\k=0.. * adjoint M1)))) + trace (M1 * \ * adjoint M1) - trace (DS (M1 * \ * adjoint M1))" using Suc(1)[OF dsDSM1r] by auto also have "\ = trace (P * (M0 * \ * adjoint M0)) + (\k=0.. * adjoint M1))))) + trace (M1 * \ * adjoint M1) - (\k=0.. * adjoint M1))))" by auto also have "\ = (\k=0..<(Suc k). trace (P * (D0 k \))) + trace \ - (\k=0..<(Suc k). trace (D0 k \))" by (simp add: eq1 eq4) finally show ?case. qed { fix \ assume dsr: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto have limDW: "limit_mat (\n. matrix_sum d (\k. D0 k \) (n)) (denote (While M S) \) d" using limit_mat_denote_while_n[OF wc dr pdor] unfolding D0_def M0_def M1_def DS_def by auto then have "limit_mat (\n. (P * (matrix_sum d (\k. D0 k \) (n)))) (P * (denote (While M S) \)) d" using mat_mult_limit[OF dP] unfolding mat_mult_seq_def by auto then have limtrPm: "(\n. trace (P * (matrix_sum d (\k. D0 k \) (n)))) \ trace (P * (denote (While M S) \))" using mat_trace_limit by auto with limDW have limtrDW:"(\n. trace (matrix_sum d (\k. D0 k \) (n))) \ trace (denote (While M S) \)" using mat_trace_limit by auto have limm: "(\n. trace (matrix_sum d (\k. D0 k \) (n))) \ trace (denote (While M S) \)" using mat_trace_limit limDW by auto have closeWS: "is_quantum_predicate P \ is_quantum_predicate (wlp S P)" for P proof - assume asm: "is_quantum_predicate P" then have dP: "P \ carrier_mat d d" using is_quantum_predicate_def by auto then show "is_quantum_predicate (wlp S P)" using wlp_mono_and_close[OF wcS asm asm] lowner_le_refl by auto qed have monoWS: "is_quantum_predicate P \ is_quantum_predicate Q \ P \\<^sub>L Q \ wlp S P \\<^sub>L wlp S Q" for P Q using wlp_mono_and_close[OF wcS] by auto have "is_quantum_predicate (wlp (While M S) P)" using wlp_while_exists[of "wlp S" M P] closeWS monoWS m qpP by auto have "limit_mat W (wlp_while M0 M1 (wlp S) P) d" unfolding W_def M0_def M1_def using wlp_while_exists[of "wlp S" M P] closeWS monoWS m qpP by auto then have "limit_mat (\k. (W k) * \) ((wlp_while M0 M1 (wlp S) P) * \) d" using mult_mat_limit dr by auto then have lim1: "(\k. trace ((W k) * \)) \ trace ((wlp_while M0 M1 (wlp S) P) * \)" using mat_trace_limit by auto have dD0kr: "D0 k \ \ carrier_mat d d" for k unfolding D0_def using denote_while_n_dim[OF dr dM0 dM1 pdor] denote_positive_trace_dim[OF wcS, folded DS_def] by auto then have "(P * (matrix_sum d (\k. D0 k \) (n))) = matrix_sum d (\k. P * (D0 k \)) n" for n using matrix_sum_distrib_left[OF dP] by auto moreover have "trace (matrix_sum d (\k. P * (D0 k \)) n) = (\k=0..)))" for n using trace_matrix_sum_linear dD0kr dP by auto ultimately have eqPsD0kr: "trace (P * (matrix_sum d (\k. D0 k \) (n))) = (\k=0..)))" for n by auto with limtrPm have lim2: "(\k. (\k=0..)))) \ trace (P * (denote (While M S) \))" by auto have "trace (matrix_sum d (\k. D0 k \) (n)) = (\k=0..))" for n using trace_matrix_sum_linear dD0kr by auto with limtrDW have lim3: "(\k. (\k=0..))) \ trace (denote (While M S) \)" by auto have "(\k. (\k=0..))) + trace \) \ trace (P * (denote (While M S) \)) + trace \" using tendsto_add[of "\k. (\k=0..)))"] lim2 by auto then have "(\k. (\k=0..))) + trace \ - (\k=0..))) \ trace (P * (denote (While M S) \)) + trace \ - trace (denote (While M S) \)" using tendsto_diff[of _ _ _ "\k. (\k=0..))"] lim3 by auto then have lim4: "(\k. trace ((W k) * \)) \ trace (P * (denote (While M S) \)) + trace \ - trace (denote (While M S) \)" using eqk[OF dsr] by auto then have "trace ((wlp_while M0 M1 (wlp S) P) * \) = trace (P * (denote (While M S) \)) + trace \ - trace (denote (While M S) \)" using eqk[OF dsr] tendsto_unique[OF _ lim1 lim4] by auto } then show ?case unfolding M0_def M1_def DS_def wlp.simps by auto qed lemma denote_while_split: assumes wc: "well_com (While M S)" and dsr: "\ \ density_states" shows "denote (While M S) \ = (M 0) * \ * adjoint (M 0) + denote (While M S) (denote S (M 1 * \ * adjoint (M 1)))" proof - have m: "measurement d 2 M" using wc by auto have wcs: "well_com S" using wc by auto define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using m measurement_def M0_def M1_def by auto have M1leq: "adjoint M1 * M1 \\<^sub>L 1\<^sub>m d" using measurement_le_one_mat m M1_def by auto define DS where "DS = denote S" define D0 where "D0 = denote_while_n M0 M1 DS" define D1 where "D1 = denote_while_n_comp M0 M1 DS" define D where "D = denote_while_n_iter M0 M1 DS" define DW where "DW \ = denote (While M S) \" for \ { fix \ assume dsr: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto have pdoDkr: "\k. partial_density_operator (D k \)" unfolding D_def using pdo_denote_while_n_iter[OF dr pdor dM1 M1leq] denote_partial_density_operator[OF wcs] denote_dim[OF wcs, folded DS_def] apply (fold DS_def) by auto then have pDkr: "\k. positive (D k \)" unfolding partial_density_operator_def by auto have dDkr: "\k. D k \ \ carrier_mat d d" using denote_while_n_iter_dim[OF dr pdor dM1 M1leq denote_dim_pdo[OF wcs, folded DS_def], of id M0, simplified, folded D_def] by auto then have dD0kr: "\k. D0 k \ \ carrier_mat d d" unfolding D0_def denote_while_n.simps apply (fold D_def) using dM0 by auto } note dD0k = this have "matrix_sum d (\k. D0 k \) k \ carrier_mat d d" if dsr: "\ \ density_states" for \ k using matrix_sum_dim[OF dD0k, of _ "\k. \" id, OF dsr] dsr by auto { fix k have "matrix_sum d (\k. D0 k \) (Suc k) = (D0 0 \) + matrix_sum d (\k. D0 (Suc k) \) k" using matrix_sum_shift_Suc[of _ "\k. D0 k \"] dD0k[OF dsr] by fastforce also have "\ = M0 * \ * adjoint M0 + matrix_sum d (\k. D0 k (DS (M1 * \ * adjoint M1))) k" unfolding D0_def by auto finally have "matrix_sum d (\k. D0 k \) (Suc k) = M0 * \ * adjoint M0 + matrix_sum d (\k. D0 k (DS (M1 * \ * adjoint M1))) k". } note eqk = this have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def dsr by auto then have "M1 * \ * adjoint M1 \ carrier_mat d d" and "partial_density_operator (M1 * \ * adjoint M1)" using dM1 dr pdo_close_under_measurement[OF dM1 dr pdor M1leq] by auto then have dSM1r: "(DS (M1 * \ * adjoint M1)) \ carrier_mat d d" and pdoSM1r: "partial_density_operator (DS (M1 * \ * adjoint M1))" unfolding DS_def using denote_dim_pdo[OF wcs] by auto have "limit_mat (matrix_sum d (\k. D0 k \)) (DW \) d" unfolding M0_def M1_def D0_def DS_def DW_def using limit_mat_denote_while_n[OF wc dr pdor] by auto then have liml: "limit_mat (\k. matrix_sum d (\k. D0 k \) (Suc k)) (DW \) d" using limit_mat_ignore_initial_segment[of "matrix_sum d (\k. D0 k \)" "DW \" d 1] by auto have dM0r: "M0 * \ * adjoint M0 \ carrier_mat d d" using dM0 dr by fastforce have "limit_mat (matrix_sum d (\k. D0 k (DS (M1 * \ * adjoint M1)))) (DW (DS (M1 * \ * adjoint M1))) d" using limit_mat_denote_while_n[OF wc dSM1r pdoSM1r] unfolding M0_def M1_def D0_def DS_def DW_def by auto then have limr: "limit_mat (mat_add_seq (M0 * \ * adjoint M0) (matrix_sum d (\k. D0 k (DS (M1 * \ * adjoint M1))))) (M0 * \ * adjoint M0 + (DW (DS (M1 * \ * adjoint M1)))) d" using mat_add_limit[OF dM0r] by auto moreover have "(\k. matrix_sum d (\k. D0 k \) (Suc k)) = (mat_add_seq (M0 * \ * adjoint M0) (matrix_sum d (\k. D0 k (DS (M1 * \ * adjoint M1)))))" using eqk mat_add_seq_def by auto ultimately have "limit_mat (\k. matrix_sum d (\k. D0 k \) (Suc k)) (M0 * \ * adjoint M0 + (DW (DS (M1 * \ * adjoint M1)))) d" by auto with liml limit_mat_unique have "DW \ = (M0 * \ * adjoint M0 + (DW (DS (M1 * \ * adjoint M1))))" by auto then show ?thesis unfolding DW_def M0_def M1_def DS_def by auto qed lemma wlp_while_split: assumes wc: "well_com (While M S)" and qpP: "is_quantum_predicate P" shows "wlp (While M S) P = adjoint (M 0) * P * (M 0) + adjoint (M 1) * (wlp S (wlp (While M S) P)) * (M 1)" proof - have m: "measurement d 2 M" using wc by auto have wcs: "well_com S" using wc by auto define M0 where "M0 = M 0" define M1 where "M1 = M 1" have dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" using m measurement_def M0_def M1_def by auto have M1leq: "adjoint M1 * M1 \\<^sub>L 1\<^sub>m d" using measurement_le_one_mat m M1_def by auto define DS where "DS = denote S" define D0 where "D0 = denote_while_n M0 M1 DS" define D1 where "D1 = denote_while_n_comp M0 M1 DS" define D where "D = denote_while_n_iter M0 M1 DS" define DW where "DW \ = denote (While M S) \" for \ have dP: "P \ carrier_mat d d" using qpP is_quantum_predicate_def by auto have qpWP: "is_quantum_predicate (wlp (While M S) P)" using qpP wc wlp_close[OF wc qpP] by auto then have "is_quantum_predicate (wlp S (wlp (While M S) P))" using wc wlp_close[OF wcs] by auto then have dWWP: "(wlp S (wlp (While M S) P)) \ carrier_mat d d" using is_quantum_predicate_def by auto have dWP: "(wlp (While M S) P) \ carrier_mat d d" using qpWP is_quantum_predicate_def by auto { fix \ assume dsr: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto have dsM1r: "M1 * \ * adjoint M1 \ density_states" unfolding density_states_def using pdo_close_under_measurement[OF dM1 dr pdor] M1leq dM1 dr by fastforce then have dsDSM1r: "DS (M1 * \ * adjoint M1) \ density_states" unfolding density_states_def DS_def using denote_dim_pdo[OF wcs] by auto have dM0r: "M0 * \ * adjoint M0 \ carrier_mat d d" using dM0 dr by fastforce have dDWDSM1r: "DW (DS (M1 * \ * adjoint M1)) \ carrier_mat d d" unfolding DW_def using denote_dim[OF wc] dsDSM1r density_states_def by auto have eq2: "trace ((wlp (While M S) P) * DS (M1 * \ * adjoint M1)) = trace (P * (DW (DS (M1 * \ * adjoint M1)))) + trace (DS (M1 * \ * adjoint M1)) - trace (DW (DS (M1 * \ * adjoint M1)))" unfolding DW_def using wlp_soundness[OF wc qpP] dsDSM1r by auto have eq3: "trace (M1 * \ * adjoint M1) = trace \ - trace (M0 * \ * adjoint M0)" unfolding M0_def M1_def using m trace_measurement2[OF m dr] dr by (simp add: algebra_simps) have "trace (adjoint M1 * (wlp S (wlp (While M S) P)) * M1 * \) = trace ((wlp S (wlp (While M S) P)) * (M1 * \ * adjoint M1))" using dWWP dM1 dr by (mat_assoc d) also have "\ = trace ((wlp (While M S) P) * DS (M1 * \ * adjoint M1)) + trace (M1 * \ * adjoint M1) - trace (DS (M1 * \ * adjoint M1))" unfolding DS_def using wlp_soundness[OF wcs qpWP] dsM1r by auto also have "\ = trace (P * (DW (DS (M1 * \ * adjoint M1)))) + trace (M1 * \ * adjoint M1) - trace (DW (DS (M1 * \ * adjoint M1)))" using eq2 by auto also have "\ = trace (P * (DW (DS (M1 * \ * adjoint M1)))) + trace \ - (trace (M0 * \ * adjoint M0) + trace (DW (DS (M1 * \ * adjoint M1))))" using eq3 by auto finally have eq4: "trace (adjoint M1 * (wlp S (wlp (While M S) P)) * M1 * \) = trace (P * (DW (DS (M1 * \ * adjoint M1)))) + trace \ - (trace (M0 * \ * adjoint M0) + trace (DW (DS (M1 * \ * adjoint M1))))". have "trace (adjoint M0 * P * M0 * \) + trace (P * (DW (DS (M1 * \ * adjoint M1)))) = trace (P * ((M0 * \ * adjoint M0) + (DW (DS (M1 * \ * adjoint M1)))))" using dP dr dM0 dDWDSM1r by (mat_assoc d) also have "\ = trace (P * (DW \))" unfolding DW_def M0_def M1_def DS_def using denote_while_split[OF wc dsr] by auto finally have eq5: "trace (adjoint M0 * P * M0 * \) + trace (P * (DW (DS (M1 * \ * adjoint M1)))) = trace (P * (DW \))". have "trace (M0 * \ * adjoint M0) + trace (DW (DS (M1 * \ * adjoint M1))) = trace (M0 * \ * adjoint M0 + (DW (DS (M1 * \ * adjoint M1))))" using dr dM0 dDWDSM1r by (mat_assoc d) also have "\ = trace (DW \)" unfolding DW_def DS_def M0_def M1_def denote_while_split[OF wc dsr] by auto finally have eq6: "trace (M0 * \ * adjoint M0) + trace (DW (DS (M1 * \ * adjoint M1))) = trace (DW \)". from eq5 eq4 eq6 have eq7: "trace (adjoint M0 * P * M0 * \) + trace (adjoint M1 * wlp S (wlp (While M S) P) * M1 * \) = trace (P * DW \) + trace \ - trace (DW \)" by auto have eq8: "trace (adjoint M0 * P * M0 * \) + trace (adjoint M1 * wlp S (wlp (While M S) P) * M1 * \) = trace ((adjoint M0 * P * M0 + adjoint M1 * wlp S (wlp (While M S) P) * M1) * \)" using dM0 dM1 dr dP dWWP by (mat_assoc d) from eq7 eq8 have eq9: "trace ((adjoint M0 * P * M0 + adjoint M1 * wlp S (wlp (While M S) P) * M1) * \) = trace (P * DW \) + trace \ - trace (DW \)" by auto have eq10: "trace ((wlp (While M S) P) * \) = trace (P * DW \) + trace \ - trace (DW \)" unfolding DW_def using wlp_soundness[OF wc qpP] dsr by auto with eq9 have "trace ((wlp (While M S) P) * \) = trace ((adjoint M0 * P * M0 + adjoint M1 * wlp S (wlp (While M S) P) * M1) * \)" by auto } then have "(wlp (While M S) P) = (adjoint M0 * P * M0 + adjoint M1 * wlp S (wlp (While M S) P) * M1)" using trace_pdo_eq_imp_eq[OF dWP, of "adjoint M0 * P * M0 + adjoint M1 * wlp S (wlp (While M S) P) * M1"] dM0 dP dM1 dWWP density_states_def by fastforce then show ?thesis using M0_def M1_def by auto qed lemma wlp_is_weakest_liberal_precondition: assumes "well_com S" and "is_quantum_predicate P" shows "is_weakest_liberal_precondition (wlp S P) S P" unfolding is_weakest_liberal_precondition_def proof (auto) show qpWP: "is_quantum_predicate (wlp S P)" using wlp_close assms by auto have eq: "trace (wlp S P * \) = trace (P * (denote S \)) + trace \ - trace (denote S \)" if dsr: "\ \ density_states" for \ using wlp_soundness assms dsr by auto then show "\\<^sub>p {wlp S P} S {P}" unfolding hoare_partial_correct_def by auto fix Q assume qpQ: "is_quantum_predicate Q" and p: "\\<^sub>p {Q} S {P}" { fix \ assume dsr: "\ \ density_states" then have "trace (Q * \) \ trace (P * (denote S \)) + trace \ - trace (denote S \)" using hoare_partial_correct_def p by auto then have "trace (Q * \) \ trace (wlp S P * \)" using eq[symmetric] dsr by auto } then show "Q \\<^sub>L wlp S P" using lowner_le_trace density_states_def qpQ qpWP is_quantum_predicate_def by auto qed subsection \Hoare triples for partial correctness\ inductive hoare_partial :: "complex mat \ com \ complex mat \ bool" ("\\<^sub>p ({(1_)}/ (_)/ {(1_)})" 50) where "is_quantum_predicate P \ \\<^sub>p {P} SKIP {P}" | "is_quantum_predicate P \ \\<^sub>p {adjoint U * P * U} Utrans U {P}" | "is_quantum_predicate P \ is_quantum_predicate Q \ is_quantum_predicate R \ \\<^sub>p {P} S1 {Q} \ \\<^sub>p {Q} S2 {R} \ \\<^sub>p {P} Seq S1 S2 {R}" | "(\k. k < n \ is_quantum_predicate (P k)) \ is_quantum_predicate Q \ (\k. k < n \ \\<^sub>p {P k} S ! k {Q}) \ \\<^sub>p {matrix_sum d (\k. adjoint (M k) * P k * M k) n} Measure n M S {Q}" | "is_quantum_predicate P \ is_quantum_predicate Q \ \\<^sub>p {Q} S {adjoint (M 0) * P * M 0 + adjoint (M 1) * Q * M 1} \ \\<^sub>p {adjoint (M 0) * P * M 0 + adjoint (M 1) * Q * M 1} While M S {P}" | "is_quantum_predicate P \ is_quantum_predicate Q \ is_quantum_predicate P' \ is_quantum_predicate Q' \ P \\<^sub>L P' \ \\<^sub>p {P'} S {Q'} \ Q' \\<^sub>L Q \ \\<^sub>p {P} S {Q}" theorem hoare_partial_sound: "\\<^sub>p {P} S {Q} \ well_com S \ \\<^sub>p {P} S {Q}" proof (induction rule: hoare_partial.induct) case (1 P) then show ?case unfolding hoare_partial_correct_def by auto next case (2 P U) (*utrans*) then have dU: "U \ carrier_mat d d" and "unitary U" and dP: "P \ carrier_mat d d" using is_quantum_predicate_def by auto then have uU: "adjoint U * U = 1\<^sub>m d" using unitary_def by auto show ?case unfolding hoare_partial_correct_def denote.simps(2) proof fix \ assume "\ \ density_states" then have dr: "\ \ carrier_mat d d" using density_states_def by auto have e1: "trace (U * \ * adjoint U) = trace ((adjoint U * U) * \)" using dr dU by (mat_assoc d) also have "\ = trace \" using uU dr by auto finally have e1: "trace (U * \ * adjoint U) = trace \" . have e2: "trace (P * (U * \ * adjoint U)) = trace (adjoint U * P * U * \)" using dU dP dr by (mat_assoc d) with e1 have "trace (P * (U * \ * adjoint U)) + (trace \ - trace (U * \ * adjoint U)) = trace (adjoint U * P * U * \)" using e1 by auto then show "trace (adjoint U * P * U * \) \ trace (P * (U * \ * adjoint U)) + (trace \ - trace (U * \ * adjoint U))" by auto qed next case (3 P Q R S1 S2) (*seq*) then have wc1: "\\<^sub>p {P} S1 {Q}" and wc2: "\\<^sub>p {Q} S2 {R}" by auto show ?case unfolding hoare_partial_correct_def denote.simps(3) proof clarify fix \ assume \: "\ \ density_states" have 1: "trace (P * \) \ trace (Q * denote S1 \) + (trace \ - trace (denote S1 \))" using wc1 hoare_partial_correct_def \ by auto have \': "denote S1 \ \ density_states" using 3(8) denote_density_states \ by auto have 2: "trace (Q * denote S1 \) \ trace (R * denote S2 (denote S1 \)) + (trace (denote S1 \) - trace (denote S2 (denote S1 \)))" using wc2 hoare_partial_correct_def \' by auto show "trace (P * \) \ trace (R * denote S2 (denote S1 \)) + (trace \ - trace (denote S2 (denote S1 \)))" using 1 2 by auto qed next case (4 n P Q S M) (*if*) then have wc: "k < n \ well_com (S!k)" and c: "k < n \ \\<^sub>p {P k} (S!k) {Q}" and m: "measurement d n M" and dMk: "k < n \ M k \ carrier_mat d d" and aMMkleq: "k < n \ adjoint (M k) * M k \\<^sub>L 1\<^sub>m d" and dPk: "k < n \ P k \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" for k using is_quantum_predicate_def measurement_def measure_well_com measurement_le_one_mat by auto { fix \ assume \: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pdor: "partial_density_operator \" using density_states_def by auto have dsr: "k < n \ (M k) * \ * adjoint (M k) \ density_states" for k unfolding density_states_def using dMk pdo_close_under_measurement[OF dMk dr pdor aMMkleq] dr by fastforce then have leqk: "k < n \ trace ((P k) * ((M k) * \ * adjoint (M k))) \ trace (Q * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))" for k using c unfolding hoare_partial_correct_def by auto have "k < n \ M k * \ * adjoint (M k) \ carrier_mat d d" for k using dMk dr by fastforce then have dsMrk: "k < n \ matrix_sum d (\k. (M k * \ * adjoint (M k))) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. (M k * \ * adjoint (M k))" d] by fastforce have "k < n \ adjoint (M k) * P k * M k \ carrier_mat d d" for k using dMk dPk by fastforce then have dsMP: "k < n \ matrix_sum d (\k. adjoint (M k) * P k * M k) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. adjoint (M k) * P k * M k" d] by fastforce have dSMrk: "k < n \ denote (S ! k) (M k * \ * adjoint (M k)) \ carrier_mat d d" for k using denote_dim[OF wc, of k "M k * \ * adjoint (M k)"] dsr density_states_def by fastforce have dsSMrk: "k < n \ matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k \ carrier_mat d d" for k using matrix_sum_dim[of k "\k. denote (S ! k) (M k * \ * adjoint (M k))" d, OF dSMrk] by fastforce have "k \ n \ trace (matrix_sum d (\k. adjoint (M k) * P k * M k) k * \) \ trace (Q * (denote (Measure k M S) \)) + (trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) k) - trace (denote (Measure k M S) \))" for k unfolding denote_measure_expand[OF _ 4(5)] proof (induct k) case 0 then show ?case using dQ dr pdor partial_density_operator_def positive_trace by auto next case (Suc k) then have k: "k < n" by auto have eq1: "trace (matrix_sum d (\k. adjoint (M k) * P k * M k) (Suc k) * \) = trace (adjoint (M k) * P k * M k * \) + trace (matrix_sum d (\k. adjoint (M k) * P k * M k) k * \)" unfolding matrix_sum.simps using dMk[OF k] dPk[OF k] dr dsMP[OF k] by (mat_assoc d) have "trace (adjoint (M k) * P k * M k * \) = trace (P k * (M k * \ * adjoint (M k)))" using dMk[OF k] dPk[OF k] dr by (mat_assoc d) also have "\ \ trace (Q * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))" using leqk k by auto finally have eq2: "trace (adjoint (M k) * P k * M k * \) \ trace (Q * (denote (S!k) ((M k) * \ * adjoint (M k)))) + (trace ((M k) * \ * adjoint (M k)) - trace (denote (S ! k) ((M k) * \ * adjoint (M k))))". have eq3: "trace (Q * matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) (Suc k)) = trace (Q * (denote (S!k) (M k * \ * adjoint (M k)))) + trace (Q * matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k)" unfolding matrix_sum.simps using dQ dSMrk[OF k] dsSMrk[OF k] by (mat_assoc d) have eq4: "trace (denote (S ! k) (M k * \ * adjoint (M k)) + matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k) = trace (denote (S ! k) (M k * \ * adjoint (M k))) + trace (matrix_sum d (\k. denote (S!k) (M k * \ * adjoint (M k))) k)" using dSMrk[OF k] dsSMrk[OF k] by (mat_assoc d) show ?case apply (subst eq1) apply (subst eq3) apply (simp del: less_eq_complex_def) apply (subst trace_add_linear[of "M k * \ * adjoint (M k)" d "matrix_sum d (\k. M k * \ * adjoint (M k)) k"]) apply (simp add: dMk adjoint_dim[OF dMk] dr mult_carrier_mat[of _ d d _ d] k) apply (simp add: dsMrk k) apply (subst eq4) apply (insert eq2 Suc(1) k, fastforce) done qed then have leq: "trace (matrix_sum d (\k. adjoint (M k) * P k * M k) n * \) \ trace (Q * denote (Measure n M S) \) + (trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) n) - trace (denote (Measure n M S) \))" by auto have "trace (matrix_sum d (\k. (M k) * \ * adjoint (M k)) n) = trace \" using trace_measurement m dr by auto with leq have "trace (matrix_sum d (\k. adjoint (M k) * P k * M k) n * \) \ trace (Q * denote (Measure n M S) \) + (trace \ - trace (denote (Measure n M S) \))" unfolding denote_measure_def by auto } then show ?case unfolding hoare_partial_correct_def by auto next case (5 P Q S M) (*while*) define M0 where "M0 = M 0" define M1 where "M1 = M 1" from 5 have wcs: "well_com S" and c: "\\<^sub>p {Q} S {adjoint M0 * P * M0 + adjoint M1 * Q * M1}" and m: "measurement d 2 M" and dM0: "M0 \ carrier_mat d d" and dM1: "M1 \ carrier_mat d d" and dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" and qpQ: "is_quantum_predicate Q" and wc: "well_com (While M S)" using measurement_def is_quantum_predicate_def M0_def M1_def by auto then have M0leq: "adjoint M0 * M0 \\<^sub>L 1\<^sub>m d" and M1leq: "adjoint M1 * M1 \\<^sub>L 1\<^sub>m d" using measurement_le_one_mat[OF m] M0_def M1_def by auto define DS where "DS = denote S" have "\\ \ density_states. trace (Q * \) \ trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * DS \) + trace \ - trace (DS \)" using hoare_partial_correct_def[of Q S "adjoint M0 * P * M0 + adjoint M1 * Q * M1"] c DS_def by auto define D0 where "D0 = denote_while_n M0 M1 DS" define D1 where "D1 = denote_while_n_comp M0 M1 DS" define D where "D = denote_while_n_iter M0 M1 DS" { fix \ assume dsr: "\ \ density_states" then have dr: "\ \ carrier_mat d d" and pr: "positive \" and pdor: "partial_density_operator \" using density_states_def partial_density_operator_def by auto have pdoDkr: "\k. partial_density_operator (D k \)" unfolding D_def using pdo_denote_while_n_iter[OF dr pdor dM1 M1leq] denote_partial_density_operator[OF wcs] denote_dim[OF wcs, folded DS_def] apply (fold DS_def) by auto then have pDkr: "\k. positive (D k \)" unfolding partial_density_operator_def by auto have dDkr: "\k. D k \ \ carrier_mat d d" using denote_while_n_iter_dim[OF dr pdor dM1 M1leq denote_dim_pdo[OF wcs, folded DS_def], of id M0, simplified, folded D_def] by auto then have dD0kr: "\k. D0 k \ \ carrier_mat d d" unfolding D0_def denote_while_n.simps apply (fold D_def) using dM0 by auto then have dPD0kr: "\k. P * (D0 k \) \ carrier_mat d d" using dP by auto have "\k. positive (D0 k \)" unfolding D0_def denote_while_n.simps by (fold D_def, rule positive_close_under_left_right_mult_adjoint[OF dM0 dDkr pDkr]) then have trge0: "\k. trace (D0 k \) \ 0" using positive_trace dD0kr by blast have DSr: "\ \ density_states \ DS \ \ density_states" for "\" unfolding DS_def density_states_def using denote_partial_density_operator denote_dim wcs by auto have dsD1nr: "D1 n \ \ density_states" for n unfolding D1_def denote_while_n_comp.simps apply (fold D_def) unfolding density_states_def apply (auto) apply (insert dDkr dM1 adjoint_dim[OF dM1], auto) apply (rule pdo_close_under_measurement[OF dM1 spec[OF allI[OF dDkr], of "\x. n"] spec[OF allI[OF pdoDkr], of "\x. n"] M1leq]) done have leQn: "trace (Q * D1 n \) \ trace (P * D0 (Suc n) \) + trace (Q * D1 (Suc n) \) + trace (D1 n \) - trace (D (Suc n) \)" for n proof - have "(\\\density_states. trace (Q * \) \ trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * denote S \) + (trace \ - trace (denote S \)))" using c hoare_partial_correct_def by auto then have leQn': "trace (Q * (D1 n \)) \ trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * (DS (D1 n \))) + (trace (D1 n \) - trace (DS (D1 n \)))" using dsD1nr[of n] DS_def by auto have "(DS (D1 n \)) \ carrier_mat d d" unfolding DS_def using denote_dim[OF wcs] dsD1nr density_states_def by auto then have "trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * (DS (D1 n \))) = trace (P * (M0 * (DS (D1 n \)) * adjoint M0)) + trace (Q * (M1 * (DS (D1 n \)) * adjoint M1))" using dP dQ dM0 dM1 by (mat_assoc d) moreover have "trace (P * (M0 * (DS (D1 n \)) * adjoint M0)) = trace (P * D0 (Suc n) \)" unfolding D0_def denote_while_n.simps apply (subst denote_while_n_iter_assoc) by (fold denote_while_n_comp.simps D1_def, auto) moreover have "trace (Q * (M1 * (DS (D1 n \)) * adjoint M1)) = trace (Q * D1 (Suc n) \)" apply (subst (2) D1_def) unfolding denote_while_n_comp.simps apply (subst denote_while_n_iter_assoc) by (fold denote_while_n_comp.simps D1_def, auto) ultimately have "trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * (DS (D1 n \))) = trace (P * D0 (Suc n) \) + trace (Q * D1 (Suc n) \)" by auto moreover have "trace (DS (D1 n \)) = trace (D (Suc n) \)" unfolding D_def apply (subst denote_while_n_iter_assoc) by (fold denote_while_n_comp.simps D1_def, auto) ultimately show ?thesis using leQn' by auto qed have 12: "trace (P * (M0 * \ * adjoint M0)) + trace (Q * (M1 * \ * adjoint M1)) \ (\k=0..<(n+2). trace (P * (D0 k \))) + trace (Q * (D1 (n+1) \)) + (\k=0..<(n+1). trace (D1 k \) - trace (D (k+1) \))" for n proof (induct n) case 0 show ?case apply (simp del: less_eq_complex_def) unfolding D0_def D1_def D_def denote_while_n_comp.simps denote_while_n.simps denote_while_n_iter.simps using leQn[of 0] unfolding D1_def D0_def D_def denote_while_n.simps denote_while_n_comp.simps denote_while_n_iter.simps by auto next case (Suc n) have "trace (Q * D1 (n + 1) \) \ trace (P * D0 (Suc (Suc n)) \) + trace (Q * D1 (Suc (Suc n)) \) + trace (D1 (Suc n) \) - trace (D (Suc (Suc n)) \)" using leQn[of "n + 1"] by auto with Suc show ?case apply (simp del: less_eq_complex_def) by auto qed have tr_measurement: "\ \ carrier_mat d d \ trace (M0 * \ * adjoint M0) + trace (M1 * \ * adjoint M1) = trace \" for \ using trace_measurement2[OF m, folded M0_def M1_def] by auto have 14: "(\k=0..<(n+1). trace (D1 k \) - trace (D (k+1) \)) = trace \ - trace (D (n+1) \) - (\k=0..<(n+1). trace (D0 k \))" for n proof (induct n) case 0 show ?case apply (simp) unfolding D1_def D0_def denote_while_n_comp.simps denote_while_n.simps denote_while_n_iter.simps using tr_measurement[OF dr] by (auto simp add: algebra_simps) next case (Suc n) have "trace (D0 (Suc n) \) + trace (D1 (Suc n) \) = trace (D (Suc n) \)" unfolding D0_def D1_def denote_while_n.simps denote_while_n_comp.simps apply (fold D_def) using tr_measurement dDkr by auto then have "trace (D1 (Suc n) \) = trace (D (Suc n) \) - trace (D0 (Suc n) \)" by (auto simp add: algebra_simps) then show ?case using Suc by simp qed have 15: "trace (Q * (D1 n \)) \ trace (D n \) - trace (D0 n \)" for n proof - have QleId: "Q \\<^sub>L 1\<^sub>m d" using is_quantum_predicate_def qpQ by auto then have "trace (Q * (D1 n \)) \ trace (1\<^sub>m d * (D1 n \))" using dsD1nr[of n] unfolding density_states_def lowner_le_trace[OF dQ one_carrier_mat] by auto also have "\ = trace (D1 n \)" using dsD1nr[of n] unfolding density_states_def by auto also have "\ = trace (M1 * (D n \) * adjoint M1)" unfolding D1_def denote_while_n_comp.simps D_def by auto also have "\ = trace (D n \) - trace (M0 * (D n \) * adjoint M0)" using tr_measurement[OF dDkr[of n]] by (simp add: algebra_simps) also have "\ = trace (D n \) - trace (D0 n \)" unfolding D0_def denote_while_n.simps by (fold D_def, auto) finally show ?thesis. qed have tmp: "\a b c. 0 \ a \ b \ c - a \ b \ (c::complex)" by simp then have 151: "\n. trace (Q * (D1 n \)) \ trace (D n \)" by (auto simp add: tmp[OF trge0 15] simp del: less_eq_complex_def) have main_leq: "\n. trace (P * (M0 * \ * adjoint M0)) + trace (Q * (M1 * \ * adjoint M1)) \ trace (P * (matrix_sum d (\k. D0 k \) (n+2))) + trace \ - trace (matrix_sum d (\k. D0 k \) (n+2))" proof - fix n have "(\k=0..<(n+2). trace (P * (D0 k \))) + trace (Q * (D1 (n+1) \)) + (\k=0..<(n+1). trace (D1 k \) - trace (D (k+1) \)) \ (\k=0..<(n+2). trace (P * (D0 k \))) + trace (Q * (D1 (n+1) \)) + trace \ - trace (D (n+1) \) - (\k=0..<(n+1). trace (D0 k \))" by (subst 14, auto) also have "\ \ (\k=0..<(n+2). trace (P * (D0 k \))) + trace (D (n+1) \) - trace (D0 (n+1) \) + trace \ - trace (D (n+1) \) - (\k=0..<(n+1). trace (D0 k \))" using 15[of "n+1"] by auto also have "\ = (\k=0..<(n+2). trace (P * (D0 k \))) + trace \ - (\k=0..<(n+2). trace (D0 k \))" by auto also have "\ = trace (matrix_sum d (\k. (P * (D0 k \))) (n+2)) + trace \ - (\k=0..<(n+2). trace (D0 k \))" using trace_matrix_sum_linear[of "n+2" "\k. (P * (D0 k \))" d, symmetric] dPD0kr by auto also have "\ = trace (matrix_sum d (\k. (P * (D0 k \))) (n+2)) + trace \ - trace (matrix_sum d (\k. D0 k \) (n+2))" using trace_matrix_sum_linear[of "n+2" "\k. D0 k \" d, symmetric] dD0kr by auto also have "\ = trace (P * (matrix_sum d (\k. D0 k \) (n+2))) + trace \ - trace (matrix_sum d (\k. D0 k \) (n+2))" using matrix_sum_distrib_left[OF dP dD0kr, of id "n+2"] by auto finally have "(\k=0..<(n+2). trace (P * (D0 k \))) + trace (Q * (D1 (n+1) \)) + (\k=0..<(n+1). trace (D1 k \) - trace (D (k+1) \)) \ trace (P * (matrix_sum d (\k. D0 k \) (n+2))) + trace \ - trace (matrix_sum d (\k. D0 k \) (n+2))" . then show "trace (P * (M0 * \ * adjoint M0)) + trace (Q * (M1 * \ * adjoint M1)) \ trace (P * (matrix_sum d (\k. D0 k \) (n+2))) + trace \ - trace (matrix_sum d (\k. D0 k \) (n+2))" using 12[of n] by auto qed have "limit_mat (\n. matrix_sum d (\k. D0 k \) (n)) (denote (While M S) \) d" using limit_mat_denote_while_n[OF wc dr pdor] unfolding D0_def M0_def M1_def DS_def by auto then have limp2: "limit_mat (\n. matrix_sum d (\k. D0 k \) (n + 2)) (denote (While M S) \) d" using limit_mat_ignore_initial_segment[of "\n. matrix_sum d (\k. D0 k \) (n)" "(denote (While M S) \)" d 2] by auto then have "limit_mat (\n. (P * (matrix_sum d (\k. D0 k \) (n+2)))) (P * (denote (While M S) \)) d" using mat_mult_limit[OF dP] unfolding mat_mult_seq_def by auto then have limPm: "(\n. trace (P * (matrix_sum d (\k. D0 k \) (n+2)))) \ trace (P * (denote (While M S) \))" using mat_trace_limit by auto have limm: "(\n. trace (matrix_sum d (\k. D0 k \) (n+2))) \ trace (denote (While M S) \)" using mat_trace_limit limp2 by auto have leq_lim: "trace (P * (M0 * \ * adjoint M0)) + trace (Q * (M1 * \ * adjoint M1)) \ trace (P * (denote (While M S) \)) + trace \ - trace (denote (While M S) \)" (is "?lhs \ ?rhs") using main_leq proof - define seq where "seq n = trace (P * matrix_sum d (\k. D0 k \) (n + 2)) - trace (matrix_sum d (\k. D0 k \) (n + 2)) " for n define seqlim where "seqlim = trace (P * (denote (While M S) \)) - trace (denote (While M S) \)" have main_leq': "?lhs \ trace \ + seq n" for n unfolding seq_def using main_leq by (simp add: algebra_simps) have limseq: "seq \ seqlim" - unfolding seq_def seqlim_def using tendsto_intros(15)[OF limPm limm] by auto + unfolding seq_def seqlim_def using tendsto_diff[OF limPm limm] by auto have limrs: "(\n. trace \ + seq n) \ (trace \ + seqlim)" using tendsto_add[OF _ limseq] by auto have limrsRe: "(\n. Re (trace \ + seq n)) \ Re (trace \ + seqlim)" using tendsto_Re[OF limrs] by auto have main_leq_Re: "Re ?lhs \ Re (trace \ + seq n)" for n using main_leq' by auto have Re: "Re ?lhs \ Re (trace \ + seqlim)" using Lim_bounded2[OF limrsRe ] main_leq_Re by auto have limrsIm: "(\n. Im (trace \ + seq n)) \ Im (trace \ + seqlim)" using tendsto_Im[OF limrs] by auto have main_leq_Im: "Im ?lhs = Im (trace \ + seq n)" for n using main_leq' unfolding less_eq_complex_def by auto then have limIm: "(\n. Im (trace \ + seq n)) \ Im ?lhs" using tendsto_intros(1) by auto have Im: "Im ?lhs = Im (trace \ + seqlim)" using tendsto_unique[OF _ limIm limrsIm] by auto have "?lhs \ trace \ + seqlim" using Re Im by auto then show "?lhs \ ?rhs" unfolding seqlim_def by auto qed have "trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * \) = trace (P * (M0 * \ * adjoint M0)) + trace (Q * (M1 * \ * adjoint M1))" using dr dM0 dM1 dP dQ by (mat_assoc d) then have "trace ((adjoint M0 * P * M0 + adjoint M1 * Q * M1) * \) \ trace (P * (denote (While M S) \)) + (trace \ - trace (denote (While M S) \))" using leq_lim by auto } then show ?case unfolding hoare_partial_correct_def denote.simps(5) apply (fold M0_def M1_def DS_def D0_def D1_def) by auto next case (6 P Q P' Q' S) then have wcs: "well_com S" and c: "\\<^sub>p {P'} S {Q'}" and dP: "P \ carrier_mat d d" and dQ: "Q \ carrier_mat d d" and dP': "P' \ carrier_mat d d" and dQ': "Q' \ carrier_mat d d" using is_quantum_predicate_def by auto show ?case unfolding hoare_partial_correct_def proof fix \ assume pds: "\ \ density_states" then have pdor: "partial_density_operator \" and dr: "\ \ carrier_mat d d" using density_states_def by auto have pdoSr: "partial_density_operator (denote S \)" using denote_partial_density_operator pdor dr wcs by auto have dSr: "denote S \ \ carrier_mat d d" using denote_dim pdor dr wcs by auto have "trace (P * \) \ trace (P' * \)" using lowner_le_trace[OF dP dP'] 6 dr pdor by auto also have "\ \ trace (Q' * denote S \) + (trace \ - trace (denote S \))" using c unfolding hoare_partial_correct_def using pds by auto also have "\ \ trace (Q * denote S \) + (trace \ - trace (denote S \))" using lowner_le_trace[OF dQ' dQ] 6 dSr pdoSr by auto finally show "trace (P * \) \ trace (Q * denote S \) + (trace \ - trace (denote S \)) ". qed qed lemma wlp_complete: "well_com S \ is_quantum_predicate P \ \\<^sub>p {wlp S P} S {P}" proof (induct S arbitrary: P) case SKIP then show ?case unfolding wlp.simps using hoare_partial.intros by auto next case (Utrans U) then show ?case unfolding wlp.simps using hoare_partial.intros by auto next case (Seq S1 S2) then have wc1: "well_com S1" and wc2: "well_com S2" and qpP: "is_quantum_predicate P" and p2: "\\<^sub>p {wlp S2 P} S2 {P}" by auto have qpW2P: "is_quantum_predicate (wlp S2 P)" using wlp_close[OF wc2 qpP] by auto then have p1: "\\<^sub>p {wlp S1 (wlp S2 P)} S1 {wlp S2 P}" using Seq by auto have qpW1W2P: "is_quantum_predicate (wlp S1 (wlp S2 P))" using wlp_close[OF wc1 qpW2P] by auto then show ?case unfolding wlp.simps using hoare_partial.intros qpW1W2P qpW2P qpP p1 p2 by auto next case (Measure n M S) then have wc: "well_com (Measure n M S)" and qpP: "is_quantum_predicate P" by auto have set: "k < n \ (S!k) \ set S" for k using wc by auto have wck: "k < n \ well_com (S!k)" for k using wc measure_well_com by auto then have qpWkP: "k < n \ is_quantum_predicate (wlp (S!k) P)" for k using wlp_close qpP by auto have pk: "k < n \ \\<^sub>p {(wlp (S!k) P)} (S!k) {P}" for k using Measure(1) set wck qpP by auto show ?case unfolding wlp_measure_expand[OF wc] using hoare_partial.intros qpWkP qpP pk by auto next case (While M S) then have wc: "well_com (While M S)" and wcS: "well_com S" and qpP: "is_quantum_predicate P " by auto have qpWP: "is_quantum_predicate (wlp (While M S) P)" using wlp_close[OF wc qpP] by auto then have qpWWP: "is_quantum_predicate (wlp S (wlp (While M S) P))" using wlp_close wcS by auto have "\\<^sub>p {wlp S (wlp (While M S) P)} S {wlp (While M S) P}" using While(1) wcS qpWP by auto moreover have eq: "wlp (While M S) P = adjoint (M 0) * P * M 0 + adjoint (M 1) * wlp S (wlp (While M S) P) * M 1" using wlp_while_split wc qpP by auto ultimately have p: "\\<^sub>p {wlp S (wlp (While M S) P)} S {adjoint (M 0) * P * M 0 + adjoint (M 1) * wlp S (wlp (While M S) P) * M 1}" by auto then show ?case using hoare_partial.intros(5)[OF qpP qpWWP p] eq by auto qed theorem hoare_partial_complete: "\\<^sub>p {P} S {Q} \ well_com S \ is_quantum_predicate P \ is_quantum_predicate Q \ \\<^sub>p {P} S {Q}" proof - assume p: "\\<^sub>p {P} S {Q}" and wc: "well_com S" and qpP: "is_quantum_predicate P" and qpQ: "is_quantum_predicate Q" then have dQ: "Q \ carrier_mat d d" using is_quantum_predicate_def by auto have qpWP: "is_quantum_predicate (wlp S Q)" using wlp_close wc qpQ by auto then have dWP: "wlp S Q \ carrier_mat d d" using is_quantum_predicate_def by auto have eq: "trace (wlp S Q * \) = trace (Q * (denote S \)) + trace \ - trace (denote S \)" if dsr: "\ \ density_states" for \ using wlp_soundness wc qpQ dsr by auto then have "\\<^sub>p {wlp S Q} S {Q}" unfolding hoare_partial_correct_def by auto { fix \ assume dsr: "\ \ density_states" then have "trace (P * \) \ trace (Q * (denote S \)) + trace \ - trace (denote S \)" using hoare_partial_correct_def p by auto then have "trace (P * \) \ trace (wlp S Q * \)" using eq[symmetric] dsr by auto } then have le: "P \\<^sub>L wlp S Q" using lowner_le_trace density_states_def qpP qpWP is_quantum_predicate_def by auto moreover have wlp: "\\<^sub>p {wlp S Q} S {Q}" using wlp_complete wc qpQ by auto ultimately show "\\<^sub>p {P} S {Q}" using hoare_partial.intros(6)[OF qpP qpQ qpWP qpQ] lowner_le_refl[OF dQ] by auto qed subsection \Consequences of completeness\ lemma hoare_patial_seq_assoc_sem: shows "\\<^sub>p {A} (S1 ;; S2) ;; S3 {B} \ \\<^sub>p {A} S1 ;; (S2 ;; S3) {B}" unfolding hoare_partial_correct_def denote.simps by auto lemma hoare_patial_seq_assoc: assumes "well_com S1" and "well_com S2" and "well_com S3" and "is_quantum_predicate A" and "is_quantum_predicate B" shows "\\<^sub>p {A} (S1 ;; S2) ;; S3 {B} \ \\<^sub>p {A} S1 ;; (S2 ;; S3) {B}" proof assume "\\<^sub>p {A} S1;; S2;; S3 {B}" then have "\\<^sub>p {A} (S1 ;; S2) ;; S3 {B}" using hoare_partial_sound assms by auto then have "\\<^sub>p {A} S1 ;; (S2 ;; S3) {B}" using hoare_patial_seq_assoc_sem by auto then show "\\<^sub>p {A} S1 ;; (S2 ;; S3) {B}" using hoare_partial_complete assms by auto next assume "\\<^sub>p {A} S1;; (S2;; S3) {B}" then have "\\<^sub>p {A} S1;; (S2;; S3) {B}" using hoare_partial_sound assms by auto then have "\\<^sub>p {A} S1;; S2;; S3 {B}" using hoare_patial_seq_assoc_sem by auto then show "\\<^sub>p {A} S1;; S2;; S3 {B}" using hoare_partial_complete assms by auto qed end end diff --git a/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy b/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy --- a/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy +++ b/thys/Transcendence_Series_Hancl_Rucki/Transcendence_Series.thy @@ -1,1042 +1,1042 @@ (* Title: Transcendence_Series.thy Authors: Angeliki Koutsoukou-Argyraki and Wenda Li, University of Cambridge *) section \The transcendence of certain infinite series\ theory "Transcendence_Series" imports "HOL-Analysis.Multivariate_Analysis" "HOL-Computational_Algebra.Polynomial" Prime_Number_Theorem.Prime_Number_Theorem_Library begin text \ We formalise the proofs of two transcendence criteria by J. Han\v{c}l and P. Rucki that assert the transcendence of the sums of certain infinite series built up by sequences that fulfil certain properties (Theorems 2.1 and 2.2 in \cite{hancl2005}, HanclRucki1 and HanclRucki2 here respectively). Both proofs make use of Roth's celebrated theorem on diophantine approximations to algebraic numbers from 1955 \cite{roth_1955} which we assume and implement within the locale RothsTheorem. A small mistake was detected in the original proof of Theorem 2.1, and the authors suggested to us a fix for the problem (in communication by email). Our formalised proof incorporates this correction (see the Remark in the proof of HanclRucki1). \ subsection \Misc\ lemma powr_less_inverse_iff: fixes x y z::real assumes "x>0""y>0""z>0" shows "x powr y < z \ x < z powr (inverse y)" proof assume "x powr y < z" from powr_less_mono2[OF _ _ this,of "inverse y"] show "x < z powr inverse y" using assms by (auto simp:powr_powr) next assume *:"x < z powr inverse y" from powr_less_mono2[OF _ _ *,of y] show "x powr y < z" using assms by (auto simp:powr_powr) qed lemma powr_less_inverse_iff': fixes x y z::real assumes "x>0""y>0""z>0" shows "z< x powr y \ z powr (inverse y) < x" using powr_less_inverse_iff[symmetric,of _ "inverse y"] assms by auto lemma powr_less_eq_inverse_iff: fixes x y z::real assumes "x>0""y>0""z>0" shows "x powr y \ z \ x \ z powr (inverse y)" by (meson assms(1) assms(2) assms(3) not_less powr_less_inverse_iff') lemma powr_less_eq_inverse_iff': fixes x y z::real assumes "x>0""y>0""z>0" shows "z\ x powr y \ z powr (inverse y) \ x" by (simp add: assms(1) assms(2) assms(3) powr_less_eq_inverse_iff) lemma tendsto_PInfty_mono: assumes "(ereal o f) \ \" "\\<^sub>F x in sequentially. f x \ g x" shows "(ereal o g) \ \" using assms unfolding comp_def tendsto_PInfty_eq_at_top by (elim filterlim_at_top_mono, simp) lemma limsup_infinity_imp_Inf_many: assumes "limsup f = \" shows "(\ m. (\\<^sub>\i. f i > ereal m))" unfolding INFM_nat proof (clarify,rule ccontr) fix m k assume "\ (\n>k. ereal m < f n)" then have "\n>k. f n \ ereal m" by auto then have "\\<^sub>F n in sequentially. f n \ ereal m" using eventually_at_top_dense by blast then have "limsup f \ ereal m" using Limsup_bounded by auto then show False using assms by simp qed lemma snd_quotient_plus_leq: defines "de\(snd o quotient_of)" shows "de (x+y) \ de x * de y " proof - obtain x1 x2 y1 y2 where xy: "quotient_of x = (x1,x2)" "quotient_of y=(y1,y2)" by (meson surj_pair) have "x2>0" "y2>0" using xy quotient_of_denom_pos by blast+ then show ?thesis unfolding de_def comp_def rat_plus_code xy apply (auto split:prod.split simp:Rat.normalize_def Let_def) by (smt div_by_1 gcd_pos_int int_div_less_self mult_eq_0_iff mult_sign_intros(1)) qed lemma quotient_of_inj: "inj quotient_of" unfolding inj_def by (simp add: quotient_of_inject) lemma infinite_inj_imageE: assumes "infinite A" "inj_on f A" "f ` A \ B" shows "infinite B" using assms inj_on_finite by blast lemma incseq_tendsto_limsup: fixes f::"nat \ 'a::{complete_linorder,linorder_topology}" assumes "incseq f" shows "f \ limsup f" using LIMSEQ_SUP assms convergent_def convergent_ereal tendsto_Limsup trivial_limit_sequentially by blast subsection \Main proofs\ text \Since the proof of Roths theorem has not been formalized yet, we implement it into a locale and used it as an assumption.\ locale RothsTheorem = assumes RothsTheorem:"\\ \. algebraic \ \ \ \ \ \ infinite {(p,q). q>0 \ coprime p q \ \\ - of_int p / of_int q\ < 1 / q powr \} \ \ \ 2" theorem (in RothsTheorem) HanclRucki1: fixes a b ::"nat\int" and \ ::real defines "aa\(\n. real_of_int (a n))" and "bb\(\n. real_of_int (b n))" assumes a_pos:"\ k. a k >0" and b_pos:"\ k. b k >0" and "\ >0" and limsup_infy:"limsup (\ k. aa (k+1)/(\i = 0..k. aa i)powr(2+\)*(1/bb (k+1))) = \" and liminf_1:"liminf (\k. aa (k+1) / aa k * bb k / bb (k+1)) > 1" shows "\ algebraic(suminf (\ k. bb k / aa k))" proof - have summable:"summable (\ k. bb k / aa k)" proof (rule ratio_test_convergence) have [simp]:"aa k>0" "bb k>0" for k unfolding aa_def bb_def using a_pos b_pos by auto show "\\<^sub>F n in sequentially. 0 < bb n / aa n" apply (rule eventuallyI) by auto show "1 < liminf (\n. ereal (bb n / aa n / (bb (Suc n) / aa (Suc n))))" using liminf_1 by (auto simp:algebra_simps) qed have [simp]: "aa k>0" "bb k>0" for k unfolding aa_def bb_def by (auto simp add: a_pos b_pos) have ab_1:"aa k\1" "bb k\1" for k unfolding aa_def bb_def using a_pos b_pos by (auto simp add: int_one_le_iff_zero_less) define B where "B=liminf (\x. ereal (aa (x + 1) / aa x * bb x / bb (x + 1)))" define M where "M= (case B of ereal m \ (m+1)/2 | _ \ 2)" have "M > 1" "M < B" using liminf_1 unfolding M_def apply (fold B_def) by (cases B,auto)+ text \ Remark: In the original proof of Theorem 2.1 in \cite{hancl2005} it was claimed in p.534 that from assumption (3) (i.e. @{thm liminf_1}), we obtain that: $\forall A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $, however note the counterexample where $a_{k+1}= k (a_1 a_2 ... a_k)^{\lceil 2+ \delta \rceil}$ if k is odd, and $a_{k+1} = 2 a_k$ otherwise, with $b_k =1$ for all $k$. In commmunication by email the authors suggested to replace the claim $\forall A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $ with $\exists A>1~\exists k_0~ \forall k > k_0~ \frac{1}{A} \frac{b_k}{ a_k} > \frac{ b_{k+1}}{ a_{k+1}} $ which solves the problem and the proof proceeds as in the paper. The witness for $\exists A>1 $ is denoted by $M$ here.\ have bb_aa_event:"\\<^sub>F k in sequentially. (1/M)*(bb k / aa k)> bb(k+1)/ aa (k+1)" using less_LiminfD[OF \M < B\[unfolded B_def],simplified] apply eventually_elim using \M > 1\ by (auto simp:divide_simps algebra_simps) have bb_aa_p:"\\<^sub>F k in sequentially. \p. bb(k+p)/ aa (k+p) \ (1/M^p)*(bb k / aa k)" proof - obtain k0 where k0_ineq: "\n\k0. bb (n + 1) / aa (n + 1) < 1 / M * (bb n / aa n)" using bb_aa_event unfolding eventually_sequentially by auto have "bb(k+p)/ aa (k+p) \ (1/M^p)*(bb k / aa k)" when "k\k0" for p k proof (induct p) case 0 then show ?case by auto next case (Suc p) have " bb (k + Suc p) / aa (k + Suc p) < 1 / M * (bb (k+p) / aa (k+p))" using k0_ineq[rule_format,of "k+p"] that by auto also have "... \ 1 / M ^(Suc p) * (bb k / aa k)" using Suc \M>1\ by (auto simp add:field_simps) finally show ?case by auto qed then show ?thesis unfolding eventually_sequentially by auto qed define \ where "\ = suminf (\ k. bb k / aa k)" have \_Inf_many:"\\<^sub>\ k. \\ - (\k = 0..k. bb k / aa k)\ < 1 / prod aa {0..k} powr (2 + \)" proof - have "\\ - (\i = 0..k. bb i / aa i)\ = \\i. bb (i+(k+1)) / aa (i+(k+1))\" for k unfolding \_def apply (subst suminf_minus_initial_segment[of _ "k+1",OF summable]) using atLeast0AtMost lessThan_Suc_atMost by auto moreover have "\\<^sub>\ k. \\i. bb(i+(k+1))/ aa (i+(k+1))\ < 1 / prod aa {0..k} powr (2 + \)" proof - define P where "P = (\ i. \p. bb (i + 1 + p) / aa (i + 1 + p) \ 1 / M ^ p * (bb (i + 1) / aa (i + 1)))" define Q where "Q= (\ i. ereal (M / (M - 1)) < ereal (aa (i + 1) / prod aa {0..i} powr (2 + \) * (1 / bb (i + 1))))" have "\\<^sub>\ i. P i" using bb_aa_p[THEN sequentially_offset, of 1] cofinite_eq_sequentially unfolding P_def by auto moreover have "\\<^sub>\i. Q i" using limsup_infy[THEN limsup_infinity_imp_Inf_many,rule_format,of "(M / (M -1))"] unfolding Q_def . moreover have "\\i. bb(i+(k+1))/ aa (i+(k+1))\ < 1 / prod aa {0..k} powr (2 + \)" when "P k" "Q k" for k proof - have summable_M:"summable (\i. 1 / M ^ i)" apply (rule summable_ratio_test[of "1/M"]) using \M>1\ by auto have "(\i. bb (i + (k + 1)) / aa (i + (k + 1))) \ 0" apply (rule suminf_nonneg) subgoal using summable_ignore_initial_segment[OF summable,of "k+1"] by auto subgoal by (simp add: less_imp_le) done then have "\\i. bb (i + (k + 1)) / aa (i + (k + 1))\ = (\i. bb (i + (k + 1)) / aa (i + (k + 1)))" by auto also have "... \ (\i. 1 / M ^ i * (bb (k + 1) / aa (k + 1)))" apply (rule suminf_le) subgoal using that(1) unfolding P_def by (auto simp add:algebra_simps) subgoal using summable_ignore_initial_segment[OF summable,of "k+1"] by auto subgoal using summable_mult2[OF summable_M,of " bb (k + 1) / aa (k + 1)"] by auto done also have "... = (bb (k + 1) / aa (k + 1)) * (\i. 1 / M ^ i)" using suminf_mult2[OF summable_M,of " bb (k + 1) / aa (k + 1)"] by (auto simp:algebra_simps) also have "... = (bb (k + 1) / aa (k + 1)) * (\i. (1 / M) ^ i)" using \M>1\ by (auto simp: field_simps) also have "... = (bb (k + 1) / aa (k + 1)) * (M / (M -1))" apply (subst suminf_geometric) using \M>1\ by (auto simp: field_simps) also have "... < (bb (k + 1) / aa (k + 1)) * (aa (k + 1) / prod aa {0..k} powr (2 + \) * (1 / bb (k + 1)))" apply (subst mult_less_cancel_left_pos) using that(2) unfolding Q_def by auto also have "... = 1/ prod aa {0..k} powr (2 + \)" using ab_1[of "Suc k"] by auto finally show ?thesis . qed ultimately show ?thesis by (smt INFM_conjI INFM_mono) qed ultimately show ?thesis by auto qed define pq where "pq = (\k. quotient_of (\i=0..k. of_int (b i) / of_int (a i)))" define p q where "p = fst \ pq" and "q = snd \ pq" have coprime_pq:"coprime (p k) (q k)" and q_pos:"q k > 0" and pq_sum:"p k / q k = (\i=0..k. b i / a i)" for k proof - have eq: "quotient_of (\i=0..k. of_int (b i) / of_int (a i)) = (p k, q k)" by (simp add: p_def q_def pq_def) from quotient_of_coprime[OF eq] show "coprime (p k) (q k)" . from quotient_of_denom_pos[OF eq] show "q k > 0" . have "(\i=0..k. b i / a i) = of_rat (\i=0..k. of_int (b i) / of_int (a i))" by (simp add: of_rat_sum of_rat_divide) also have "(\i=0..k. rat_of_int (b i) / rat_of_int (a i)) = rat_of_int (p k) / rat_of_int (q k)" using quotient_of_div[OF eq] by simp finally show "p k / q k = (\i=0..k. b i / a i)" by (simp add:of_rat_divide) qed have \_Inf_many2:"\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \)" using \_Inf_many proof (elim INFM_mono) fix k assume asm:"\\ - (\k = 0..k. bb k / aa k)\ < 1 / prod aa {0..k} powr (2 + \)" have "\\ - real_of_int (p k) / real_of_int (q k)\ = \\ - (\k = 0..k. bb k / aa k)\" using pq_sum unfolding aa_def bb_def by auto also have "... < 1 / prod aa {0..k} powr (2 + \)" using asm by auto also have "... \ 1 / q k powr (2 + \)" proof - have "q k \ prod aa {0..k}" proof (induct k) case 0 then show ?case unfolding q_def pq_def aa_def apply (simp add:rat_divide_code of_int_rat quotient_of_Fract) using ab_1[of 0,unfolded aa_def bb_def] unfolding Let_def normalize_def apply auto by (metis div_by_1 gcd_pos_int less_imp_le less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2) next case (Suc k) define de where "de=snd \ quotient_of" have "real_of_int (q (Suc k)) = de (\i=0..Suc k. of_int (b i) / of_int (a i))" unfolding q_def pq_def de_def by simp also have "... = de ((\i=0..k. of_int (b i) / of_int (a i)) + of_int (b (Suc k)) / of_int (a (Suc k)))" by simp also have "... \ de (\i=0..k. of_int (b i) / of_int (a i)) * de (of_int (b (Suc k)) / of_int (a (Suc k)))" using snd_quotient_plus_leq[folded de_def] by presburger also have "... = q k * de (of_int (b (Suc k)) / of_int (a (Suc k)))" unfolding q_def pq_def de_def by auto also have "... = q k * snd (Rat.normalize (b (Suc k), a (Suc k)))" by (simp add:rat_divide_code of_int_rat quotient_of_Fract de_def) also have "... \ q k * aa (Suc k)" using ab_1[of "Suc k"] q_pos[of "k"] unfolding normalize_def aa_def bb_def Let_def apply auto by (metis div_by_1 int_one_le_iff_zero_less less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2 zero_less_one) also have "... \ prod aa {0..k} * aa (Suc k)" using Suc ab_1[of "Suc k"] by auto also have "... = prod aa {0..Suc k}" by (simp add: prod.atLeast0_atMost_Suc) finally show ?case . qed then show ?thesis by (smt \0 < \\ frac_le of_int_0 of_int_le_iff powr_gt_zero powr_mono2 q_pos) qed finally show "\\ - real_of_int (p k) / real_of_int (q k)\ < 1 / real_of_int (q k) powr (2 + \)" . qed define pqs where "pqs={(p, q). q>0 \ coprime p q \ \\ - real_of_int p / real_of_int q\ < 1 / q powr (2 + \)}" have \_infinite:"infinite pqs" proof - define A where "A={k. \\ - (p k) / (q k)\ < 1 / (q k) powr (2 + \)}" have "\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \)" using \_Inf_many2 . then have "infinite A" unfolding Inf_many_def A_def by auto moreover have "inj_on (\k. (p k, q k)) A" proof - define g where "g=(\i. rat_of_int (b i) / rat_of_int (a i))" define f where "f=(\k. \i = 0..k. g i)" have g_pos:"g i>0" for i unfolding g_def by (simp add: a_pos b_pos) have "strict_mono f" unfolding strict_mono_def f_def proof safe fix x y::nat assume "x < y" then have "sum g {0..y} - sum g {0..x} = sum g {x<..y}" apply (subst Groups_Big.sum_diff[symmetric]) by (auto intro:arg_cong2[where f=sum]) also have "... > 0" apply (rule ordered_comm_monoid_add_class.sum_pos) using \x g_pos by auto finally have "sum g {0..y} - sum g {0..x} >0" . then show "sum g {0..x} < sum g {0..y}" by auto qed then have "inj f" using strict_mono_imp_inj_on by auto then have "inj (quotient_of o f)" by (simp add: inj_compose quotient_of_inj) then have "inj (\k. (p k, q k))" unfolding f_def p_def q_def pq_def comp_def apply (fold g_def) by auto then show ?thesis by (auto elim:subset_inj_on) qed moreover have "(\k. (p k, q k)) ` A \ pqs" unfolding A_def pqs_def using coprime_pq q_pos by auto ultimately show ?thesis apply (elim infinite_inj_imageE) by auto qed moreover have "finite pqs" if "\ \ \" proof - obtain m n where \_mn:"\ = (of_int m / of_int n)" and "coprime m n" "n>0" proof - obtain m n where mn:"\\\ = (of_nat m / of_nat n)" "coprime m n" "n\0" using Rats_abs_nat_div_natE[OF \\ \ \\ Rats_abs_nat_div_natE] by metis define m' and n'::int where "m'=(if \ > 0 then nat m else -nat m)" and "n'=nat n" then have "\ = (of_int m' / of_int n')" "coprime m' n'" "n'>0" using mn by auto then show ?thesis using that by auto qed have "pqs \ {(m,n)} \ {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof (rule subsetI) fix x assume "x \ pqs" define p q where "p=fst x" and "q=snd x" have "q>0" "coprime p q" and pq_less:"\\ - p / q\ < 1 / q powr (2 + \)" using \x\pqs\ unfolding p_def q_def pqs_def by auto have q_lt_n:"qp \ n\q" proof - have "m * q \ n * p" using that \coprime m n\ \coprime p q\ \q>0\ \n>0\ by (metis eq_rat(1) fst_conv int_one_le_iff_zero_less mult.commute normalize_stable not_one_le_zero quotient_of_Fract snd_conv) then have "1/(n*q) \ \m/n - p/q\" using \q>0\ \n>0\ apply (auto simp:field_simps) by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero less_irrefl not_le of_int_diff of_int_lessD of_int_mult) also have "... < 1 / q powr (2 + \)" using pq_less unfolding \_mn by auto also have "... \ 1 / q ^2" proof - have "real_of_int (q\<^sup>2) = q powr 2" apply (subst powr_numeral) unfolding power2_eq_square using \q>0\ by auto also have "... \ q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "real_of_int (q\<^sup>2) \ real_of_int q powr (2 + \)" . moreover have "real_of_int q powr (2 + \) > 0" using \0 < q\ by auto ultimately show ?thesis by (auto simp:field_simps) qed finally have " 1 / (n * q) < 1 / q\<^sup>2" . then show ?thesis using \q>0\ \n>0\ unfolding power2_eq_square by (auto simp:field_simps) qed moreover have "- \m\ - 1 \ p \ p \ \m\ + 1" when "m\p \ n\q" proof - define qn where "qn=q/n" have "0m\p \ n\q\] \q>0\ by auto have "\m/n - p / q\ < 1 / q powr (2 + \)" using pq_less unfolding \_mn by simp then have "\p / q - m/n\ < 1 / q powr (2 + \)" by simp then have " m/n- 1 / q powr (2 + \) < p/q \ p/q < m/n + 1 / q powr (2 + \)" unfolding abs_diff_less_iff by auto then have "qn*m- q / q powr (2 + \) < p \ p < qn*m + q / q powr (2 + \)" unfolding qn_def using \q>0\ by (auto simp:field_simps) moreover have "- \m\ - 1 \ qn*m- q / q powr (2 + \)" proof - have "- \m\ \ qn*m" using \0 \qn<1\ apply (cases "m\0") subgoal apply simp by (meson less_eq_real_def mult_nonneg_nonneg neg_le_0_iff_le of_int_0_le_iff order_trans) subgoal by simp done moreover have "- 1 \ - q / q powr (2 + \)" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "q \ q powr (2 + \)" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed moreover have "qn*m + q / q powr (2 + \) \ \m\ + 1" proof - have "qn*m \ \m\" using \0 \qn<1\ apply (cases "m\0") subgoal by (simp add: mult_left_le_one_le) subgoal by (smt of_int_0_le_iff zero_le_mult_iff) done moreover have "q / q powr (2 + \) \ 1" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \)" apply (rule powr_mono) using \q>0\ \\>0\ by auto finally have "q \ q powr (2 + \)" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show "x \ {(m, n)} \ {x \ pqs. - \m\ - 1 \ fst x \ fst x \ \m\ + 1 \ 0 < snd x \ snd x < n}" using \x \ pqs\ \q>0\ unfolding p_def q_def by force qed moreover have "finite {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof - have "finite ({- \m\ - 1..\m\ +1} \ {0<..pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n } \ ({- \m\ - 1..\m\ +1} \ {0<.._def) using RothsTheorem[rule_format,of \ "2 + \",folded pqs_def] \\ >0\ by auto qed theorem (in RothsTheorem) HanclRucki2: fixes a b ::"nat\int" and \ \ ::real defines "aa\(\n. real_of_int (a n))" and "bb\(\n. real_of_int (b n))" assumes a_pos:"\ k. a k >0" and b_pos:"\ k. b k >0" and "\ >0" and "\ >0" and limsup_infi:"limsup (\ k.(aa (k+1)/(\i = 0..k. aa i)powr(2+(2/\) + \)) * (1/(bb (k+1)))) = \" and ratio_large:"\ k. ( k \ t \ (( aa(k+1)/bb( k+1)) ) powr(1/(1+\)) \ (( aa k/bb k) powr(1/(1+\)))+1)" shows "\ algebraic(suminf (\ k. bb k /aa k)) " proof- have aa_bb_pos[simp]:"aa k>0" "bb k>0" for k unfolding aa_def bb_def using a_pos b_pos by auto have summable:"summable (\ k. bb k / aa k)" proof - define c0 where "c0\(aa t/bb t) powr(1/(1+\)) - t" have ab_k:"(aa k / bb k) powr(1/(1+\)) \ k + c0" when "k\t" for k using that proof (induct k) case 0 then show ?case unfolding c0_def by simp next case (Suc k) have ?case when "\ t\k" proof - have "t = Suc k" using that Suc.prems by linarith then show ?thesis unfolding c0_def by auto qed moreover have ?case when "t \ k" proof - have "(aa(k+1)/bb(k+1)) powr(1/(1+\)) \ ( aa k/bb k) powr(1/(1+\))+1" using ratio_large[rule_format,OF that] by blast then show ?thesis using Suc(1)[OF that] by simp qed ultimately show ?case by auto qed have "summable (\k. 1 / (k + c0) powr (1+\))" proof - have "c0 + t > 0" unfolding c0_def using aa_bb_pos[of t] by (simp,linarith) then have "summable (\k. 1 / (k + (c0+t)) powr (1+\))" using summable_hurwitz_zeta_real[of "1+\" "c0+t"] apply (subst (asm) powr_minus_divide) using \\>0\ by auto then show ?thesis apply (rule_tac summable_offset[of _ t]) by (auto simp:algebra_simps) qed moreover have "bb k / aa k \ 1 / (k + c0) powr (1+\)" when "k\t" for k proof - have "(aa t / bb t) powr (1 / (1 + \)) > 0" apply simp by (metis \\k. 0 < aa k\ \\k. 0 < bb k\ less_numeral_extra(3)) then have "k + c0 >0" unfolding c0_def using that by linarith then have "aa k / bb k \ (k + c0) powr (1+\)" using ab_k[OF that] apply (subst (asm) powr_less_eq_inverse_iff') using \\ >0\ by auto then have "inverse (aa k / bb k) \ inverse ((k + c0) powr (1+\))" apply (elim le_imp_inverse_le) using \k + c0 >0\ by auto then show ?thesis by (simp add: inverse_eq_divide) qed ultimately show ?thesis apply (elim summable_comparison_test'[where N=t]) using aa_bb_pos by (simp add: less_eq_real_def) qed have 7:"\\<^sub>\ k. 1 / (M powr (\/(1+\)) * (\i = 0..k. aa i) powr(2+ \ * \ / (1+\))) > (bb (k+1) / aa(k+1)) powr (\ / (1+\))" when "M > 0" for M proof - define tt where "tt= (\i. prod aa {0..i} powr (2 + 2 / \ + \))" have tt_pos:"tt i > 0" for i unfolding tt_def apply(subst powr_gt_zero,induct i) subgoal by (metis aa_bb_pos(1) order_less_irrefl prod_pos) subgoal by (metis \\k. 0 < aa k\ order_less_irrefl prod_pos) done have "\\<^sub>\i. M < (aa (i + 1) / tt i * (1 / bb (i + 1)))" using limsup_infinity_imp_Inf_many[OF limsup_infi,rule_format,of M] unfolding tt_def by auto then have "\\<^sub>\i. 1 / (M * tt i) > (bb (i+1) / aa (i+1))" apply (elim INFM_mono) using \M>0\ tt_pos by (auto simp:divide_simps algebra_simps) then have "\\<^sub>\i. (1 / (M * tt i)) powr (\/(1+\)) > (bb (i+1) / aa (i+1)) powr (\/(1+\))" apply (elim INFM_mono powr_less_mono2[rotated 2]) by (simp_all add: assms(6) pos_add_strict less_eq_real_def) moreover have "tt i powr (\/(1+\)) = prod aa {0..i} powr (2 + \ * \ / (1+\))" for i unfolding tt_def apply (auto simp:powr_powr) using \\>0\ by (simp add:divide_simps,simp add:algebra_simps) ultimately show ?thesis apply (elim INFM_mono) by (smt nonzero_mult_div_cancel_left powr_divide powr_mult powr_one_eq_one that tt_pos zero_less_divide_iff) qed have 8:"\\<^sub>F k in sequentially. \s. (( aa(k+s)/bb( k+s))) \ ((( aa k/bb k) powr(1/(1+\))) +s)powr(1+\)" using eventually_ge_at_top[of t] proof eventually_elim case (elim k) define ff where "ff=(\k. (aa k / bb k) powr (1 / (1 + \)))" have 11:"ff k+s \ ff (k+s)" for s proof (induct s) case 0 then show ?case unfolding ff_def by auto next case (Suc s) then have "ff k + Suc s \ ff (k+Suc s)" using ratio_large[rule_format,of "k+s"] \ t \ k\ unfolding ff_def by auto then show ?case by simp qed then have "(ff k+s) powr (1+\) \ ff (k+s) powr (1+\)" for s apply (rule powr_mono2[of "1+\",rotated 2]) unfolding ff_def using \\>0\ by auto then show ?case unfolding ff_def using \\>0\ apply (auto simp add:powr_powr) by (simp add: a_pos aa_def b_pos bb_def le_less) qed have 9: "(\r. 1/((z+real r)powr(1+\))) \ 1/(\ *(z-1) powr \)" "summable (\i. 1/((z+ real i)powr(1+\)))" when "z>1" for z proof - define f where "f= (\r. 1/((z+ r)powr(1+\)))" have "((\x. f (x - 1)) has_integral ((z-1) powr - \) / \) {0..}" proof - have "((\x. (z-1 + x) powr (- 1 - \)) has_integral ((z-1) powr - \) / \) {0<..}" using powr_has_integral_at_top[of 0 "z-1" "- 1 - \",simplified] \z>1\ \\>0\ by auto then have "((\x. (z-1 + x) powr (- 1 - \)) has_integral ((z-1) powr - \) / \) {0..}" apply (subst (asm) has_integral_closure[symmetric]) by (auto simp add: negligible_convex_frontier) then show ?thesis apply (rule has_integral_cong[THEN iffD1, rotated 1]) unfolding f_def by (smt powr_minus_divide) qed moreover have "\x. 0 \ x \ 0 \ f (x - 1)" unfolding f_def by simp moreover have "\x y. 0 \ x \ x \ y \ f (y - 1) \ f (x - 1)" unfolding f_def by (smt assms(6) frac_le powr_mono2 powr_nonneg_iff that) ultimately have "summable (\i. f (real i))" "(\i. f (real i)) \ (z - 1) powr - \ / \" using decreasing_sum_le_integral[of "\x. f (x-1)" "((z-1) powr - \) / \",simplified] by auto moreover have "(z - 1) powr - \ / \ = 1/(\ *(z-1) powr \)" by (simp add: powr_minus_divide) ultimately show "(\i. f (real i)) \ 1/(\ *(z-1) powr \)" by auto show "summable (\i. f (real i))" using \summable (\i. f (real i))\ . qed have u:"(\k.( aa k / bb k)) \ \" proof - define ff where "ff\(\x. ereal (aa x / bb x))" have "limsup ff = \" proof - define tt where "tt= (\i. prod aa {0..i} powr (2 + 2 / \ + \))" have "tt i \ 1" for i unfolding tt_def apply (rule ge_one_powr_ge_zero) subgoal apply (rule linordered_nonzero_semiring_class.prod_ge_1) by (simp add: a_pos aa_def int_one_le_iff_zero_less) subgoal by (simp add: \\>0\ \\>0\ less_imp_le) done then have "limsup (\x. (aa (x + 1) / tt x * (1 / bb (x + 1)))) \ limsup (\x. aa (x + 1) / bb (x + 1))" apply (intro Limsup_mono eventuallyI) apply (auto simp add:field_simps order.order_iff_strict) by (metis aa_bb_pos(1) div_by_1 frac_less2 less_irrefl less_numeral_extra(1) not_le) also have "... = limsup (\x. aa x / bb x)" by (subst limsup_shift,simp) finally have "limsup (\x. ereal (aa (x + 1) / tt x * (1 / bb (x + 1)))) \ limsup (\x. ereal (aa x / bb x))" . moreover have "limsup (\x. ereal (aa (x + 1) / tt x * (1 / bb (x + 1)))) = \" using limsup_infi unfolding tt_def by auto ultimately show ?thesis unfolding ff_def using ereal_infty_less_eq2(1) by blast qed then have "limsup (\k. ff (k+t)) = \" by (simp add: limsup_shift_k) moreover have "incseq (\k. ff (k+t))" proof (rule incseq_SucI) fix k::nat define gg where "gg\(\x. (aa x / bb x))" have "(gg (k+t)) powr (1 / (1 + \)) + 1 \ (gg (Suc (k+t))) powr (1 / (1 + \))" using ratio_large[rule_format, of "k+t",simplified] unfolding gg_def by auto then have "(gg (k+t)) powr (1 / (1 + \)) \ (gg (Suc (k+t))) powr (1 / (1 + \))" by auto then have "gg (k+t)\ gg (Suc (k+t))" by (smt aa_bb_pos(1) aa_bb_pos(2) assms(6) divide_pos_pos gg_def powr_less_mono2) then show "ff (k + t) \ ff (Suc k + t)" unfolding gg_def ff_def by auto qed ultimately have "(\k. ff (k+t)) \ \" using incseq_tendsto_limsup by fastforce then show ?thesis unfolding ff_def unfolding tendsto_def apply (subst eventually_sequentially_seg[symmetric,of _ t]) by simp qed define \ where "\ = suminf (\ k. bb k / aa k)" have 10:"\\<^sub>F k in sequentially. \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k+1) / aa (k+1)) powr (\ / (1+\))" using 8[THEN sequentially_offset,of 1] eventually_ge_at_top[of t] u[unfolded tendsto_PInfty,rule_format,THEN sequentially_offset ,of "(2 powr (1/\) / (2 powr (1/\) - 1)) powr (1+\)" 1] proof eventually_elim case (elim k) define tt where "tt=(aa (k + 1) / bb (k + 1)) powr (1 / (1 + \))" have "tt>1" proof - have "(aa k / bb k) powr (1 / (1 + \)) > 0" by (metis a_pos aa_def b_pos bb_def divide_eq_0_iff less_irrefl of_int_eq_0_iff powr_gt_zero) then show ?thesis using ratio_large[rule_format,OF \k\t\] unfolding tt_def by argo qed have "\\ - (\i = 0..k. bb i / aa i)\ = \\i. bb (i+(k+1)) / aa (i+(k+1))\" unfolding \_def apply (subst suminf_minus_initial_segment[of _ "k+1",OF summable]) using atLeast0AtMost lessThan_Suc_atMost by auto also have "... = (\i. bb (i+(k+1)) / aa (i+(k+1)))" proof - have "(\i. bb (i+(k+1)) / aa (i+(k+1))) > 0" apply (rule suminf_pos) subgoal using summable[THEN summable_ignore_initial_segment,of "k+1"] . subgoal by (simp add: a_pos aa_def b_pos bb_def) done then show ?thesis by auto qed also have "... \ (\i. 1 / (tt + i) powr (1 + \))" proof (rule suminf_le) define ff where "ff=(\k n. (aa (k+1) / bb (k+1)) powr (1 / (1 + \)) + real n)" have "bb (n + (k + 1)) / aa (n + (k + 1)) \ 1 / (ff k n) powr (1 + \)" for n proof - have "ff k n powr (1 + \) \ aa (k + n +1 ) / bb (k + n+1 )" using elim(1)[rule_format,of n] unfolding ff_def by auto moreover have "ff k n powr (1 + \) > 0" unfolding ff_def by (smt elim(2) of_nat_0_le_iff powr_gt_zero ratio_large) ultimately have "1 / ff k n powr (1 + \) \ bb (k + (n + 1)) / aa (k + (n + 1))" apply (drule_tac le_imp_inverse_le) by (auto simp add: inverse_eq_divide) then show ?thesis by (auto simp:algebra_simps) qed then show " \n. bb (n + (k + 1)) / aa (n + (k + 1)) \ 1 / (tt + real n) powr (1 + \)" unfolding ff_def tt_def by auto show "summable (\i. bb (i + (k + 1)) / aa (i + (k + 1)))" using summable[THEN summable_ignore_initial_segment,of "k+1"] . show "summable (\x. 1 / (tt + real x) powr (1 + \))" using 9(2)[OF \tt>1\] . qed also have "... \ 1/(\ *(tt-1) powr \)" using 9[OF \tt>1\] by simp also have "... < 2 / (\ *tt powr \)" proof - have "((2 powr (1/\) / (2 powr (1/\) - 1)) powr (1 + \)) < (aa (k+1) / bb (k+1))" using elim(3) by auto then have "2 powr (1/\) / (2 powr (1/\) - 1) < tt" unfolding tt_def apply (drule_tac powr_less_mono2[rotated 2,where a="1/ (1 + \)"]) using \\>0\ apply (auto simp:powr_powr ) by (subst (asm) powr_one,auto simp add:field_simps) then have " tt < (tt-1) * (2 powr (1/\))" using \\>0\ by (auto simp:divide_simps algebra_simps) then have "tt powr \ < 2 * (tt - 1) powr \" apply (drule_tac powr_less_mono2[rotated 2,where a="\"]) using \\>0\ \tt>1\ by (auto simp:powr_powr powr_mult) then show ?thesis using \\>0\ \tt>1\ by (auto simp:divide_simps) qed also have "... = 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" unfolding tt_def using \\>0\ - by (auto simp:powr_powr divide_simps algebra_simps powr_divide) + by (auto simp:powr_powr divide_simps algebra_simps powr_divide less_imp_le) finally show ?case . qed define pq where "pq = (\k. quotient_of (\i=0..k. of_int (b i) / of_int (a i)))" define p q where "p = fst \ pq" and "q = snd \ pq" have coprime_pq:"coprime (p k) (q k)" and q_pos:"q k > 0" and pq_sum:"p k / q k = (\i=0..k. b i / a i)" for k proof - have eq: "quotient_of (\i=0..k. of_int (b i) / of_int (a i)) = (p k, q k)" by (simp add: p_def q_def pq_def) from quotient_of_coprime[OF eq] show "coprime (p k) (q k)" . from quotient_of_denom_pos[OF eq] show "q k > 0" . have "(\i=0..k. b i / a i) = of_rat (\i=0..k. of_int (b i) / of_int (a i))" by (simp add: of_rat_sum of_rat_divide) also have "(\i=0..k. rat_of_int (b i) / rat_of_int (a i)) = rat_of_int (p k) / rat_of_int (q k)" using quotient_of_div[OF eq] by simp finally show "p k / q k = (\i=0..k. b i / a i)" by (simp add:of_rat_divide) qed have \_Inf_many:"\\<^sub>\ k. \\ - p k / q k\ < 1 / q k powr (2 + \ * \ / (1 + \))" proof - have *:"\\<^sub>\k. (bb (Suc k) / aa (Suc k)) powr (\ / (1 + \)) < \ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \)))" using 7[of "(2 / \) powr ((1+\)/ \)"] \\>0\ by (auto simp:powr_powr) have **:"\\<^sub>\k. \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" using 10[unfolded cofinite_eq_sequentially[symmetric]] . from INFM_conjI[OF * **] show ?thesis proof (elim INFM_mono) fix k assume asm:"(bb (Suc k) / aa (Suc k)) powr (\ / (1 + \)) < \ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \))) \ \\ - (\k = 0..k. bb k / aa k)\ < 2 / \ * (bb (k + 1) / aa (k + 1)) powr (\ / (1 + \))" have "\\ - real_of_int (p k) / real_of_int (q k)\ = \\ - (\k = 0..k. bb k / aa k)\" using pq_sum unfolding aa_def bb_def by auto also have "... < (2 / \) * (bb (k+1) / aa (k+1)) powr (\ / (1+\))" using asm by auto also have "... < (2 / \) * (\ / (2 * prod aa {0..k} powr (2 + \ * \ / (1 + \))))" apply (rule mult_strict_left_mono) using asm \\>0\ by auto also have "... = 1/ prod aa {0..k} powr (2 + \ * \ / (1 + \))" using \\>0\ by auto also have "... \ 1 / q k powr (2 + \ * \ / (1 + \))" proof - have "q k \ prod aa {0..k}" proof (induct k) case 0 then show ?case unfolding q_def pq_def aa_def apply (simp add:rat_divide_code of_int_rat quotient_of_Fract) using aa_bb_pos[of 0,unfolded aa_def bb_def] unfolding Let_def normalize_def apply auto by (metis div_by_1 less_imp_le less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2) next case (Suc k) define de where "de=snd \ quotient_of" have "real_of_int (q (Suc k)) = de (\i=0..Suc k. of_int (b i) / of_int (a i))" unfolding q_def pq_def de_def by simp also have "... = de ((\i=0..k. of_int (b i) / of_int (a i)) + of_int (b (Suc k)) / of_int (a (Suc k)))" by simp also have "... \ de (\i=0..k. of_int (b i) / of_int (a i)) * de (of_int (b (Suc k)) / of_int (a (Suc k)))" using snd_quotient_plus_leq[folded de_def] by presburger also have "... = q k * de (of_int (b (Suc k)) / of_int (a (Suc k)))" unfolding q_def pq_def de_def by auto also have "... = q k * snd (Rat.normalize (b (Suc k), a (Suc k)))" by (simp add:rat_divide_code of_int_rat quotient_of_Fract de_def) also have "... \ q k * aa (Suc k)" using aa_bb_pos[of "Suc k"] q_pos[of "k"] unfolding normalize_def aa_def bb_def Let_def apply auto by (metis div_by_1 int_one_le_iff_zero_less less_trans nonneg1_imp_zdiv_pos_iff not_less zdiv_mono2 zero_less_one) also have "... \ prod aa {0..k} * aa (Suc k)" using Suc aa_bb_pos[of "Suc k"] by auto also have "... = prod aa {0..Suc k}" by (simp add: prod.atLeast0_atMost_Suc) finally show ?case . qed then show ?thesis apply (rule_tac divide_left_mono) apply (rule powr_mono2) using \\>0\ \\>0\ q_pos[of k] apply (auto simp:powr_mult[symmetric]) by (metis aa_bb_pos(1) less_irrefl) qed finally show "\\ - p k / q k\ < 1 / q k powr (2 + \ * \ / (1 + \))" . qed qed define pqs where "pqs={(p, q). q>0 \ coprime p q \ \\ - real_of_int p / real_of_int q\ < 1 / q powr (2 + \ * \ / (1 + \))}" have \_infinite:"infinite pqs" proof - define A where "A={k. \\ - (p k) / (q k)\ < 1 / (q k) powr (2 + \ * \ / (1 + \))}" note \_Inf_many then have "infinite A" unfolding Inf_many_def A_def by auto moreover have "inj_on (\k. (p k, q k)) A" proof - define g where "g=(\i. rat_of_int (b i) / rat_of_int (a i))" define f where "f=(\k. \i = 0..k. g i)" have g_pos:"g i>0" for i unfolding g_def by (simp add: a_pos b_pos) have "strict_mono f" unfolding strict_mono_def f_def proof safe fix x y::nat assume "x < y" then have "sum g {0..y} - sum g {0..x} = sum g {x<..y}" apply (subst Groups_Big.sum_diff[symmetric]) by (auto intro:arg_cong2[where f=sum]) also have "... > 0" apply (rule ordered_comm_monoid_add_class.sum_pos) using \x g_pos by auto finally have "sum g {0..y} - sum g {0..x} >0" . then show "sum g {0..x} < sum g {0..y}" by auto qed then have "inj f" using strict_mono_imp_inj_on by auto then have "inj (quotient_of o f)" by (simp add: inj_compose quotient_of_inj) then have "inj (\k. (p k, q k))" unfolding f_def p_def q_def pq_def comp_def apply (fold g_def) by auto then show ?thesis by (auto elim:subset_inj_on) qed moreover have "(\k. (p k, q k)) ` A \ pqs" unfolding A_def pqs_def using coprime_pq q_pos by auto ultimately show ?thesis apply (elim infinite_inj_imageE) by auto qed moreover have "finite pqs" if "\ \ \" proof - obtain m n where \_mn:"\ = (of_int m / of_int n)" and "coprime m n" "n>0" proof - obtain m n where mn:"\\\ = (of_nat m / of_nat n)" "coprime m n" "n\0" using Rats_abs_nat_div_natE[OF \\ \ \\ Rats_abs_nat_div_natE] by metis define m' and n'::int where "m'=(if \ > 0 then nat m else -nat m)" and "n'=nat n" then have "\ = (of_int m' / of_int n')" "coprime m' n'" "n'>0" using mn by auto then show ?thesis using that by auto qed have "pqs \ {(m,n)} \ {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof (rule subsetI) fix x assume "x \ pqs" define p q where "p=fst x" and "q=snd x" have "q>0" "coprime p q" and pq_less:"\\ - p / q\ < 1 / q powr (2 + \ * \ / (1 + \))" using \x\pqs\ unfolding p_def q_def pqs_def by auto have q_lt_n:"qp \ n\q" proof - have "m * q \ n * p" using that \coprime m n\ \coprime p q\ \q>0\ \n>0\ by (metis eq_rat(1) fst_conv int_one_le_iff_zero_less mult.commute normalize_stable not_one_le_zero quotient_of_Fract snd_conv) then have "1/(n*q) \ \m/n - p/q\" using \q>0\ \n>0\ apply (auto simp:field_simps) by (metis add_diff_cancel_left' diff_diff_eq2 diff_zero less_irrefl not_le of_int_diff of_int_lessD of_int_mult) also have "... < 1 / q powr (2 + \ * \ / (1 + \))" using pq_less unfolding \_mn by auto also have "... \ 1 / q ^2" proof - have "real_of_int (q\<^sup>2) = q powr 2" apply (subst powr_numeral) unfolding power2_eq_square using \q>0\ by auto also have "... \ q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "real_of_int (q\<^sup>2) \ real_of_int q powr (2 + \ * \ / (1 + \))" . moreover have "real_of_int q powr (2 + \ * \ / (1 + \)) > 0" using \0 < q\ by auto ultimately show ?thesis by (auto simp:field_simps) qed finally have " 1 / (n * q) < 1 / q\<^sup>2" . then show ?thesis using \q>0\ \n>0\ unfolding power2_eq_square by (auto simp:field_simps) qed moreover have "- \m\ - 1 \ p \ p \ \m\ + 1" when "m\p \ n\q" proof - define qn where "qn=q/n" have "0m\p \ n\q\] \q>0\ by auto have "\m/n - p / q\ < 1 / q powr (2 + \ * \ / (1 + \))" using pq_less unfolding \_mn by simp then have "\p / q - m/n\ < 1 / q powr (2 + \ * \ / (1 + \))" by simp then have " m/n- 1 / q powr (2 + \ * \ / (1 + \)) < p/q \ p/q < m/n + 1 / q powr (2 + \ * \ / (1 + \))" unfolding abs_diff_less_iff by auto then have "qn*m- q / q powr (2 + \ * \ / (1 + \)) < p \ p < qn*m + q / q powr (2 + \ * \ / (1 + \))" unfolding qn_def using \q>0\ by (auto simp:field_simps) moreover have "- \m\ - 1 \ qn*m- q / q powr (2 + \ * \ / (1 + \))" proof - have "- \m\ \ qn*m" using \0 \qn<1\ apply (cases "m\0") subgoal apply simp by (meson less_eq_real_def mult_nonneg_nonneg neg_le_0_iff_le of_int_0_le_iff order_trans) subgoal by simp done moreover have "- 1 \ - q / q powr (2 + \ * \ / (1 + \))" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "q \ q powr (2 + \ * \ / (1 + \))" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed moreover have "qn*m + q / q powr (2 + \ * \ / (1 + \)) \ \m\ + 1" proof - have "qn*m \ \m\" using \0 \qn<1\ apply (cases "m\0") subgoal by (simp add: mult_left_le_one_le) subgoal by (smt of_int_0_le_iff zero_le_mult_iff) done moreover have "q / q powr (2 + \ * \ / (1 + \)) \ 1" proof - have "q = q powr 1" using \0 < q\ by auto also have "... \q powr (2 + \ * \ / (1 + \))" apply (rule powr_mono) using \q>0\ \\>0\ \\>0\ by auto finally have "q \ q powr (2 + \ * \ / (1 + \))" . then show ?thesis using \0 < q\ by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis by auto qed ultimately show "x \ {(m, n)} \ {x \ pqs. - \m\ - 1 \ fst x \ fst x \ \m\ + 1 \ 0 < snd x \ snd x < n}" using \x \ pqs\ \q>0\ unfolding p_def q_def by force qed moreover have "finite {x. x \pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n }" proof - have "finite ({- \m\ - 1..\m\ +1} \ {0<..pqs \ - \m\ - 1 \fst x \ fst x \ \m\ + 1 \ 0 snd x < n } \ ({- \m\ - 1..\m\ +1} \ {0<.._def) using RothsTheorem[rule_format,of \ "2 + \ * \ / (1 + \)",folded pqs_def] \\ >0\ \\>0\ apply (auto simp:divide_simps ) by (meson mult_le_0_iff not_less) qed subsection\ Acknowledgements\ text\A.K.-A. and W.L. were supported by the ERC Advanced Grant ALEXANDRIA (Project 742178) funded by the European Research Council and led by Professor Lawrence Paulson at the University of Cambridge, UK. Thanks to Iosif Pinelis for his clarification on MathOverflow regarding the summmability of the series in @{thm [source] RothsTheorem.HanclRucki2} @{url "https://mathoverflow.net/questions/323069/why-is-this-series-summable"} and to Manuel Eberl for his helpful comments.\ end \ No newline at end of file