diff --git a/thys/Deep_Learning/Lebesgue_Functional.thy b/thys/Deep_Learning/Lebesgue_Functional.thy --- a/thys/Deep_Learning/Lebesgue_Functional.thy +++ b/thys/Deep_Learning/Lebesgue_Functional.thy @@ -1,95 +1,95 @@ (* Author: Alexander Bentkamp, Universität des Saarlandes *) section \Alternative Lebesgue Measure Definition\ theory Lebesgue_Functional imports "HOL-Analysis.Lebesgue_Measure" begin text \Lebesgue\_Measure.lborel is defined on the typeclass euclidean\_space, which does not allow the space dimension to be dependent on a variable. As the Lebesgue measure of higher dimensions is the product measure of the one dimensional Lebesgue measure, we can easily define a more flexible version of the Lebesgue measure as follows. This version of the Lebesgue measure measures sets of functions from nat to real whose values are undefined for arguments higher than n. These "Extensional Function Spaces" are defined in HOL/FuncSet. \ definition lborel_f :: "nat \ (nat \ real) measure" where "lborel_f n = (\\<^sub>M b\{..b. interval_measure (\x. x))" unfolding product_sigma_finite_def using sigma_finite_interval_measure by auto lemma l_borel_f_1: "distr (lborel_f 1) lborel (\x. x 0) = lborel" unfolding lborel_f_def using product_sigma_finite.distr_singleton[OF product_sigma_finite_interval, of 0] lborel_eq_real lessThan_Suc by auto lemma space_lborel_f: "space (lborel_f n) = Pi\<^sub>E {.._. UNIV)" unfolding lborel_f_def unfolding space_PiM space_lborel space_borel by metis lemma space_lborel_f_subset: "space (lborel_f n) \ space (lborel_f (Suc n))" unfolding space_lborel_f by (rule subsetI, rule PiE_I, blast, metis PiE_E Suc_n_not_le_n le_cases lessThan_subset_iff subsetCE) lemma space_lborel_add_dim: assumes "f \ space (lborel_f n)" shows "f(n:=x) \ space (lborel_f (Suc n))" unfolding space_lborel_f using assms lessThan_Suc space_lborel_f by auto lemma integral_lborel_f: assumes "f \ borel_measurable (lborel_f (Suc n))" shows "integral\<^sup>N (lborel_f (Suc n)) f = \\<^sup>+ y. \\<^sup>+ x. f (x(n := y)) \lborel_f n \lborel" unfolding lborel_f_def using product_sigma_finite.product_nn_integral_insert_rev[OF product_sigma_finite_interval, of "{.. sets (lborel_f (Suc n))" assumes "\y. {x\space (lborel_f n). x(n := y) \ A} \ sets (lborel_f n)" shows "emeasure (lborel_f (Suc n)) A = \\<^sup>+ y. emeasure (lborel_f n) {x\space (lborel_f n). x(n := y) \ A} \lborel" proof - { fix x y assume "x\space (lborel_f n)" - then have "(indicator A) (x(n := y)) = (indicator {x\space (lborel_f n). x(n := y) \ A}) x" using indicator_def - by (metis (no_types, lifting) mem_Collect_eq) + then have "(indicator A) (x(n := y)) = (indicator {x\space (lborel_f n). x(n := y) \ A}) x" + by (simp add: indicator_def) } then show ?thesis unfolding nn_integral_indicator[OF assms(1), symmetric] nn_integral_indicator[OF assms(2), symmetric] integral_lborel_f[OF borel_measurable_indicator, OF assms(1)] using nn_integral_cong by (metis (no_types, lifting)) qed lemma lborel_f_measurable_add_dim: "(\f. f(n := x)) \ measurable (lborel_f n) (lborel_f (Suc n))" proof - have "x \ space lborel" by simp have 0:"(\(f, y). f(n := y)) \ (\xa. (xa, x)) = (\f. f(n := x))" unfolding comp_def using case_prod_conv by fast show ?thesis unfolding lborel_f_def using measurable_comp[OF measurable_Pair2'[of x lborel "Pi\<^sub>M {..b. lborel)", OF \x \ space lborel\] measurable_add_dim[of n "{..b. lborel"], unfolded 0] lessThan_Suc by auto qed lemma sets_lborel_f_sub_dim: assumes "A \ sets (lborel_f (Suc n))" shows "{x. x(n := y) \ A} \ space (lborel_f n) \ sets (lborel_f n)" proof - have "(\f. f(n := y)) -` A \ space (lborel_f n) \ sets (lborel_f n)" using measurable_sets[OF lborel_f_measurable_add_dim assms] by auto moreover have "(\f. f(n := y)) -` A = {x. x(n := y) \ A}" by auto finally show ?thesis by metis qed lemma lborel_f_measurable_restrict: assumes "m \ n" shows "(\f. restrict f {.. measurable (lborel_f n) (lborel_f m)" using measurable_restrict_subset lborel_f_def assms by auto lemma lborel_measurable_sub_dim: "(\f. restrict f {.. measurable (lborel_f (Suc n)) (lborel_f n)" using lborel_f_measurable_restrict[of "n" "Suc n"] by linarith lemma measurable_lborel_component [measurable]: assumes "kx. x k) \ borel_measurable (lborel_f n)" unfolding lborel_f_def using assms measurable_PiM_component_rev by simp_all end diff --git a/thys/Ergodic_Theory/Recurrence.thy b/thys/Ergodic_Theory/Recurrence.thy --- a/thys/Ergodic_Theory/Recurrence.thy +++ b/thys/Ergodic_Theory/Recurrence.thy @@ -1,2293 +1,2293 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Conservativity, recurrence\ theory Recurrence imports Measure_Preserving_Transformations begin text \A dynamical system is conservative if almost every point comes back close to its starting point. This is always the case if the measure is finite, not when it is infinite (think of the translation on $\mathbb{Z}$). In conservative systems, an important construction is the induced map: the first return map to a set of finite measure. It is measure-preserving and conservative if the original system is. This makes it possible to reduce statements about general conservative systems in infinite measure to statements about systems in finite measure, and as such is extremely useful.\ subsection \Definition of conservativity\ locale conservative = qmpt + assumes conservative: "\A. A \ sets M \ emeasure M A > 0 \ \n>0. emeasure M ((T^^n)-`A \ A) >0" lemma conservativeI: assumes "qmpt M T" "\A. A \ sets M \ emeasure M A > 0 \ \n>0. emeasure M ((T^^n)-`A \ A) >0" shows "conservative M T" unfolding conservative_def conservative_axioms_def using assms by auto text \To prove conservativity, it is in fact sufficient to show that the preimages of a set of positive measure intersect it, without any measure control. Indeed, in a non-conservative system, one can construct a set which does not satisfy this property.\ lemma conservativeI2: assumes "qmpt M T" "\A. A \ sets M \ emeasure M A > 0 \ \n>0. (T^^n)-`A \ A \ {}" shows "conservative M T" unfolding conservative_def conservative_axioms_def proof (auto simp add: assms) interpret qmpt M T using assms by auto fix A assume A_meas [measurable]: "A \ sets M" and "emeasure M A > 0" show "\n>0. 0 < emeasure M ((T ^^ n) -` A \ A)" proof (rule ccontr) assume "\ (\n>0. 0 < emeasure M ((T ^^ n) -` A \ A))" then have meas_0: "emeasure M ((T ^^ n) -` A \ A) = 0" if "n>0" for n by (metis zero_less_iff_neq_zero that) define C where "C = (\n. (T^^(Suc n))-`A \ A)" have C_meas [measurable]: "C \ sets M" unfolding C_def by measurable have "emeasure M C = 0" unfolding C_def by (intro emeasure_UN_eq_0[of M, of "\n. (T^^(Suc n))-`A \ A", OF meas_0], auto) define A2 where "A2 = A-C" then have A2_meas [measurable]: "A2 \ sets M" by simp have "\(\n>0. (T^^n)-`A2 \ A2 \ {})" proof (rule ccontr, simp) assume "\n>0. (T^^n)-`A2 \ A2 \ {}" then obtain n where n: "n > 0" "(T^^n)-`A2 \ A2 \ {}" by auto define m where "m = n-1" have "(T^^(m+1))-`A2 \ A2 \ {}" unfolding m_def using n by auto then show False using C_def A2_def by auto qed then have "emeasure M A2 = 0" using assms(2)[OF A2_meas] by (meson zero_less_iff_neq_zero) then have "emeasure M (C \ A2) = 0" using \emeasure M C = 0\ by (simp add: emeasure_Un_null_set null_setsI) moreover have "A \ C \ A2" unfolding A2_def by auto ultimately have "emeasure M A = 0" by (meson A2_meas C_meas emeasure_eq_0 sets.Un) then show False using \emeasure M A > 0\ by auto qed qed text \There is also a dual formulation, saying that conservativity follows from the fact that a set disjoint from all its preimages has to be null.\ lemma conservativeI3: assumes "qmpt M T" "\A. A \ sets M \ (\n>0. (T^^n)-`A \ A = {}) \ A \ null_sets M" shows "conservative M T" proof (rule conservativeI2[OF assms(1)]) fix A assume "A \ sets M" "0 < emeasure M A" then have "\(A \ null_sets M)" unfolding null_sets_def by auto then show "\n>0. (T ^^ n) -` A \ A \ {}" using assms(2)[OF \A \ sets M\] by auto qed text \The inverse of a conservative map is still conservative\ lemma (in conservative) conservative_Tinv: assumes "invertible_qmpt" shows "conservative M Tinv" proof (rule conservativeI2) show "qmpt M Tinv" using Tinv_qmpt[OF assms]. have "bij T" using assms unfolding invertible_qmpt_def by auto fix A assume [measurable]: "A \ sets M" and "emeasure M A > 0" then obtain n where *: "n>0" "emeasure M ((T^^n)-`A \ A) > 0" using conservative[OF \A \ sets M\ \emeasure M A > 0\] by blast have "bij (T^^n)" using bij_fn[OF \bij T\] by auto then have "bij(inv (T^^n))" using bij_imp_bij_inv by auto then have "bij (Tinv^^n)" unfolding Tinv_def using inv_fn[OF \bij T\, of n] by auto have "(T^^n)-`A \ A \ {}" using * by auto then have "(Tinv^^n)-`((T^^n)-`A \ A) \ {}" using surj_vimage_empty[OF bij_is_surj[OF \bij (Tinv^^n)\]] by meson then have **: "(Tinv^^n)-`((T^^n)-`A) \ (Tinv^^n)-` A \ {}" by auto have "(Tinv^^n)-`((T^^n)-`A) = ((T^^n) o (Tinv^^n))-`A" by auto moreover have "(T^^n) o (Tinv^^n) = (\x. x)" unfolding Tinv_def using \bij T\ fn_o_inv_fn_is_id by blast ultimately have "(Tinv^^n)-`((T^^n)-`A) = A" by auto then have "(Tinv^^n)-` A \ A \ {}" using ** by auto then show "\n>0. (Tinv ^^ n) -` A \ A \ {}" using \n>0\ by auto qed text \We introduce the locale of a conservative measure preserving map.\ locale conservative_mpt = mpt + conservative lemma conservative_mptI: assumes "mpt M T" "\A. A \ sets M \ emeasure M A > 0 \ \n>0. (T^^n)-`A \ A \ {}" shows "conservative_mpt M T" unfolding conservative_mpt_def apply (auto simp add: assms(1), rule conservativeI2) using assms(1) by (auto simp add: mpt_def assms(2)) text \The fact that finite measure preserving transformations are conservative, albeit easy, is extremely important. This result is known as Poincaré recurrence theorem.\ sublocale fmpt \ conservative_mpt proof (rule conservative_mptI) show "mpt M T" by (simp add: mpt_axioms) fix A assume A_meas [measurable]: "A \ sets M" and "emeasure M A > 0" show "\n>0. (T^^n)-`A \ A \ {}" proof (rule ccontr) assume "\(\n>0. (T^^n)-`A \ A \ {})" then have disj: "(T^^(Suc n))--`A \ A = {}" for n unfolding vimage_restr_def using zero_less_one by blast define B where "B = (\ n. (T^^n)--`A)" then have B_meas [measurable]: "B n \ sets M" for n by simp have same: "measure M (B n) = measure M A" for n by (simp add: B_def A_meas T_vrestr_same_measure(2)) have "B n \ B m = {}" if "n > m" for m n proof - have "B n \ B m = (T^^m)--` (B (n-m) \ A)" using B_def \m < n\ A_meas vrestr_intersec T_vrestr_composed(1) by auto moreover have "B (n-m) \ A = {}" unfolding B_def by (metis disj \m < n\ Suc_diff_Suc) ultimately show ?thesis by simp qed then have "disjoint_family B" by (metis disjoint_family_on_def inf_sup_aci(1) less_linear) have "measure M A < e" if "e>0" for e::real proof - obtain N::nat where "N>0" "(measure M (space M))/e0 < e\ by (metis divide_less_0_iff reals_Archimedean2 less_eq_real_def measure_nonneg not_gr0 not_le of_nat_0) then have "(measure M (space M))/N < e" using \0 < e\ \N>0\ by (metis bounded_measure div_0 le_less_trans measure_empty mult.commute pos_divide_less_eq) have *: "disjoint_family_on B {..disjoint_family B\ disjoint_family_on_mono subsetI) then have "(\i\{.. measure M (space M)" by (metis bounded_measure \\n. B n \ sets M\ image_subset_iff finite_lessThan finite_measure_finite_Union) also have "(\i\{..i\{.. measure M (space M)" by simp then have "measure M A \ (measure M (space M))/N" using \N>0\ by (simp add: mult.commute mult_imp_le_div_pos) then show "measure M A < e" using \(measure M (space M))/N by simp qed then have "measure M A \ 0" using not_less by blast then have "measure M A = 0" by (simp add: measure_le_0_iff) then have "emeasure M A = 0" using emeasure_eq_measure by simp then show False using \emeasure M A > 0\ by simp qed qed text \The following fact that powers of conservative maps are also conservative is true, but nontrivial. It is proved as follows: consider a set $A$ with positive measure, take a time $n_1$ such that $A_1 = T^{-n_1} A \cap A$ has positive measure, then a time $n_2$ such that $A_2 = T^{-n_2} A_1 \cap A$ has positive measure, and so on. It follows that $T^{-(n_i+n_{i+1}+\dots+n_j)}A \cap A$ has positive measure for all $i proposition (in conservative) conservative_power: "conservative M (T^^n)" proof (unfold_locales) show "T ^^ n \ quasi_measure_preserving M M" by (auto simp add: Tn_quasi_measure_preserving) fix A assume [measurable]: "A \ sets M" "0 < emeasure M A" define good_time where "good_time = (\K. Inf{(i::nat). i > 0 \ emeasure M ((T^^i)-`K \ A) > 0})" define next_good_set where "next_good_set = (\K. (T^^(good_time K))-`K \ A)" have good_rec: "((good_time K > 0) \ (next_good_set K \ A) \ (next_good_set K \ sets M) \ (emeasure M (next_good_set K) > 0))" if [measurable]: "K \ sets M" and "K \ A" "emeasure M K > 0" for K proof - have a: "next_good_set K \ sets M" "next_good_set K \ A" using next_good_set_def by simp_all obtain k where "k > 0" and posK: "emeasure M ((T^^k)-`K \ K) > 0" using conservative[OF \K \ sets M\, OF \emeasure M K > 0\] by auto have *:"(T^^k)-`K \ K \ (T^^k)-`K \ A" using \K \ A\ by auto have posKA: "emeasure M ((T^^k)-`K \ A) > 0" using emeasure_mono[OF *, of M] posK by simp let ?S = "{(i::nat). i>0 \ emeasure M ((T^^i)-`K \ A) > 0}" have "k \ ?S" using \k>0\ posKA by simp then have "?S \ {}" by auto then have "Inf ?S \ ?S" using Inf_nat_def1[of ?S] by simp then have "good_time K \ ?S" using good_time_def by simp then show "(good_time K > 0) \ (next_good_set K \ A) \ (next_good_set K \ sets M) \ (emeasure M (next_good_set K) > 0)" using a next_good_set_def by auto qed define B where "B = (\i. (next_good_set^^i) A)" define t where "t = (\i. good_time (B i))" have good_B: "(B i \ A) \ (B i \ sets M) \ (emeasure M (B i) > 0)" for i proof (induction i) case 0 have "B 0 = A" using B_def by simp then show ?case using \B 0 = A\ \A \ sets M\ \emeasure M A > 0\ by simp next case (Suc i) moreover have "B (i+1) = next_good_set (B i)" using B_def by simp ultimately show ?case using good_rec[of "B i"] by auto qed have t_pos: "\i. t i > 0" using t_def by (simp add: good_B good_rec) define s where "s = (\i k. (\n \ {i.. (T^^(s i k))-`A \ A" for i k proof (induction k) case 0 show ?case using s_def good_B[of i] by simp next case (Suc k) have "B(i+k+1) = (T^^(t (i+k)))-`(B (i+k)) \ A" using t_def B_def next_good_set_def by simp moreover have "B(i+k) \ (T^^(s i k))-`A" using Suc.IH by simp ultimately have "B(i+k+1) \ (T^^(t (i+k)))-` (T^^(s i k))-`A \ A" by auto then have "B(i+k+1) \ (T^^(t(i+k) + s i k))-`A \ A" by (simp add: add.commute funpow_add vimage_comp) moreover have "t(i+k) + s i k = s i (k+1)" using s_def by simp ultimately show ?case by simp qed moreover have "(T^^j)-`A \ A \ sets M" for j by simp ultimately have *: "emeasure M ((T^^(s i k))-`A \ A) > 0" for i k by (metis inf.orderE inf.strict_boundedE good_B emeasure_mono) show "\k>0. 0 < emeasure M (((T ^^ n) ^^ k) -` A \ A)" proof (cases) assume "n = 0" then have "((T ^^ n) ^^ 1) -` A = A" by simp then show ?thesis using \emeasure M A > 0\ by auto next assume "\(n = 0)" then have "n > 0" by simp define u where "u = (\i. s 0 i mod n)" have "range u \ {..0 < n\ image_subset_iff u_def) then have "finite (range u)" using finite_nat_iff_bounded by auto then have "\i j. (i (u i = u j)" by (metis finite_imageD infinite_UNIV_nat injI less_linear) then obtain i k where "k>0" "u i = u (i+k)" using less_imp_add_positive by blast moreover have "s 0 (i+k) = s 0 i + s i k" unfolding s_def by (simp add: sum.atLeastLessThan_concat) ultimately have "(s i k) mod n = 0" using u_def nat_mod_cong by metis then obtain r where "s i k = n * r" by auto moreover have "s i k > 0" unfolding s_def using \k > 0\ t_pos sum_strict_mono[of "{i..x. 0", of "\x. t x"] by simp ultimately have "r > 0" by simp moreover have "emeasure M ((T^^(n * r))-`A \ A) > 0" using * \s i k = n * r\ by metis ultimately show ?thesis by (metis funpow_mult) qed qed proposition (in conservative_mpt) conservative_mpt_power: "conservative_mpt M (T^^n)" using conservative_power mpt_power unfolding conservative_mpt_def by auto text \The standard way to use conservativity is as follows: if a set is almost disjoint from all its preimages, then it is null:\ lemma (in conservative) ae_disjoint_then_null: assumes "A \ sets M" "\n. n > 0 \ A \ (T^^n)-`A \ null_sets M" shows "A \ null_sets M" by (metis Int_commute assms(1) assms(2) conservative zero_less_iff_neq_zero null_setsD1 null_setsI) lemma (in conservative) disjoint_then_null: assumes "A \ sets M" "\n. n > 0 \ A \ (T^^n)-`A = {}" shows "A \ null_sets M" by (rule ae_disjoint_then_null, auto simp add: assms) text \Conservativity is preserved by replacing the measure by an equivalent one.\ lemma (in conservative) conservative_density: assumes [measurable]: "h \ borel_measurable M" and "AE x in M. h x \ 0" "AE x in M. h x \ \" shows "conservative (density M h) T" proof - interpret A: qmpt "density M h" T by (rule qmpt_density[OF assms]) show ?thesis apply (rule conservativeI3) apply (simp add: A.qmpt_axioms) unfolding sets_density null_sets_density[OF assms(1) assms(2)] by (metis conservative emeasure_empty not_gr_zero null_setsI) qed context qmpt begin text \We introduce the recurrent subset of $A$, i.e., the set of points of $A$ that return to $A$, and the infinitely recurrent subset, i.e., the set of points of $A$ that return infinitely often to $A$. In conservative systems, both coincide with $A$ almost everywhere.\ definition recurrent_subset::"'a set \ 'a set" where "recurrent_subset A = (\n \ {1..}. A \ (T^^n)-`A)" definition recurrent_subset_infty::"'a set \ 'a set" where "recurrent_subset_infty A = A - (\n. (T^^n)-` (A - recurrent_subset A))" lemma recurrent_subset_infty_inf_returns: "x \ recurrent_subset_infty A \ (x \ A \ infinite {n. (T^^n) x \ A})" proof assume *: "x \ recurrent_subset_infty A" have "infinite {n. (T^^n) x \ A}" proof (rule ccontr) assume "\(infinite {n. (T^^n) x \ A})" then have F: "finite {n. (T^^n) x \ A}" by auto have "0 \ {n. (T^^n) x \ A}" using * recurrent_subset_infty_def by auto then have NE: "{n. (T^^n) x \ A} \ {}" by blast define N where "N = Max {n. (T^^n) x \ A}" have "N \ {n. (T^^n) x \ A}" unfolding N_def using F NE using Max_in by auto then have "(T^^N) x \ A" by auto moreover have "x \ (T^^N)-` (A - recurrent_subset A)" using * unfolding recurrent_subset_infty_def by auto ultimately have "(T^^N) x \ recurrent_subset A" by auto then have "(T ^^ N) x \ A \ (\n. n \ {1..} \ (T ^^ n) ((T ^^ N) x) \ A)" unfolding recurrent_subset_def by blast then obtain n where "n>0" "(T^^n) ((T^^N) x) \ A" by (metis atLeast_iff gr0I not_one_le_zero) then have "n+N \ {n. (T^^n) x \ A}" by (simp add: funpow_add) then show False unfolding N_def using \n>0\ F NE by (metis Max_ge Nat.add_0_right add.commute nat_add_left_cancel_less not_le) qed then show "x \ A \ infinite {n. (T^^n) x \ A}" using * recurrent_subset_infty_def by auto next assume *: "(x \ A \ infinite {n. (T ^^ n) x \ A})" { fix n obtain N where "N>n" "(T^^N) x \ A" using * using infinite_nat_iff_unbounded by force define k where "k = N-n" then have "k>0" "N = n+k" using \N>n\ by auto then have "(T^^k) ((T^^n) x) \ A" by (metis \(T ^^ N) x \ A\ \N = n + k\ add.commute comp_def funpow_add) then have "(T^^ n) x \ A - recurrent_subset A" unfolding recurrent_subset_def using \k>0\ by auto } then show "x \ recurrent_subset_infty A" unfolding recurrent_subset_infty_def using * by auto qed lemma recurrent_subset_infty_series_infinite: assumes "x \ recurrent_subset_infty A" shows "(\n. indicator A ((T^^n) x)) = (\::ennreal)" proof (rule ennreal_ge_nat_imp_PInf) have *: "\ finite {n. (T^^n) x \ A}" using recurrent_subset_infty_inf_returns assms by auto fix N::nat obtain F where F: "finite F" "F \ {n. (T^^n) x \ A}" "card F = N" using infinite_arbitrarily_large[OF *] by blast have "N = (\n \ F. 1::ennreal)" using F(3) by auto also have "... = (\n \ F. (indicator A ((T^^n) x))::ennreal)" apply (rule sum.cong) using F(2) indicator_def by auto also have "... \ (\n. indicator A ((T^^n) x))" by (rule sum_le_suminf, auto simp add: F) finally show "N \ (\n. (indicator A ((T^^n) x))::ennreal)" by auto qed lemma recurrent_subset_infty_def': "recurrent_subset_infty A = (\m. (\n\{m..}. A \ (T^^n)-`A))" proof (auto) fix x assume x: "x \ recurrent_subset_infty A" then show "x \ A" unfolding recurrent_subset_infty_def by auto fix N::nat show "\n\{N..}. (T^^n) x \ A" using recurrent_subset_infty_inf_returns x using infinite_nat_iff_unbounded_le by auto next fix x assume "x \ A" "\N. \n\{N..}. (T^^n) x \ A" then show "x \ recurrent_subset_infty A" unfolding recurrent_subset_infty_inf_returns using infinite_nat_iff_unbounded_le by auto qed lemma recurrent_subset_incl: "recurrent_subset A \ A" "recurrent_subset_infty A \ A" "recurrent_subset_infty A \ recurrent_subset A" unfolding recurrent_subset_def recurrent_subset_infty_def' by (simp, simp, fast) lemma recurrent_subset_meas [measurable]: assumes [measurable]: "A \ sets M" shows "recurrent_subset A \ sets M" "recurrent_subset_infty A \ sets M" unfolding recurrent_subset_def recurrent_subset_infty_def' by measurable lemma recurrent_subset_rel_incl: assumes "A \ B" shows "recurrent_subset A \ recurrent_subset B" "recurrent_subset_infty A \ recurrent_subset_infty B" proof - show "recurrent_subset A \ recurrent_subset B" unfolding recurrent_subset_def using assms by auto show "recurrent_subset_infty A \ recurrent_subset_infty B" apply (auto, subst recurrent_subset_infty_inf_returns) using assms recurrent_subset_incl(2) infinite_nat_iff_unbounded_le recurrent_subset_infty_inf_returns by fastforce qed text \If a point belongs to the infinitely recurrent subset of $A$, then when they return to $A$ its iterates also belong to the infinitely recurrent subset.\ lemma recurrent_subset_infty_returns: assumes "x \ recurrent_subset_infty A" "(T^^n) x \ A" shows "(T^^n) x \ recurrent_subset_infty A" proof (subst recurrent_subset_infty_inf_returns, rule ccontr) assume "\ ((T ^^ n) x \ A \ infinite {k. (T ^^ k) ((T ^^ n) x) \ A})" then have 1: "finite {k. (T^^k) ((T^^n) x) \ A}" using assms(2) by auto have "0 \ {k. (T^^k) ((T^^n) x) \ A}" using assms(2) by auto then have 2: "{k. (T^^k) ((T^^n) x) \ A} \ {}" by blast define M where "M = Max {k. (T^^k) ((T^^n) x) \ A}" have M_prop: "\k. k > M \ (T^^k) ((T^^n) x) \ A" unfolding M_def using 1 2 by auto { fix N assume *: "(T^^N) x \ A" have "N \ n+M" proof (cases) assume "N \ n" then show ?thesis by auto next assume "\(N \ n)" then have "N > n" by simp define k where "k = N-n" have "N = n + k" unfolding k_def using \N > n\ by auto then have "(T^^k) ((T^^n)x) \ A" using * by (simp add: add.commute funpow_add) then have "k \ M" using M_prop using not_le by blast then show ?thesis unfolding k_def by auto qed } then have "finite {N. (T^^N) x \ A}" by (metis (no_types, lifting) infinite_nat_iff_unbounded mem_Collect_eq not_less) moreover have "infinite {N. (T^^N) x \ A}" using recurrent_subset_infty_inf_returns assms(1) by auto ultimately show False by auto qed lemma recurrent_subset_of_recurrent_subset: "recurrent_subset_infty(recurrent_subset_infty A) = recurrent_subset_infty A" proof show "recurrent_subset_infty (recurrent_subset_infty A) \ recurrent_subset_infty A" using recurrent_subset_incl(2)[of A] recurrent_subset_rel_incl(2) by auto show "recurrent_subset_infty A \ recurrent_subset_infty (recurrent_subset_infty A)" using recurrent_subset_infty_returns recurrent_subset_infty_inf_returns by (metis (no_types, lifting) Collect_cong subsetI) qed text \The Poincare recurrence theorem states that almost every point of $A$ returns (infinitely often) to $A$, i.e., the recurrent and infinitely recurrent subsets of $A$ coincide almost everywhere with $A$. This is essentially trivial in conservative systems, as it is a reformulation of the definition of conservativity. (What is not trivial, and has been proved above, is that it is true in finite measure preserving systems, i.e., finite measure preserving systems are automatically conservative.)\ theorem (in conservative) Poincare_recurrence_thm: assumes [measurable]: "A \ sets M" shows "A - recurrent_subset A \ null_sets M" "A - recurrent_subset_infty A \ null_sets M" "A \ recurrent_subset A \ null_sets M" "A \ recurrent_subset_infty A \ null_sets M" "emeasure M (recurrent_subset A) = emeasure M A" "emeasure M (recurrent_subset_infty A) = emeasure M A" "AE x \ A in M. x \ recurrent_subset_infty A" proof - define B where "B = {x \ A. \ n\{1..}. (T^^n) x \ (space M - A)}" have rs: "recurrent_subset A = A - B" by (auto simp add: B_def recurrent_subset_def) (meson Tn_meas assms measurable_space sets.sets_into_space subsetCE) then have *: "A - recurrent_subset A = B" using B_def by blast have "B \ null_sets M" by (rule disjoint_then_null, auto simp add: B_def) then show "A - recurrent_subset A \ null_sets M" using * by simp then have *: "(\n. (T^^n)--`(A-recurrent_subset A)) \ null_sets M" using T_quasi_preserves_null2(2) by blast have "recurrent_subset_infty A = recurrent_subset_infty A \ space M" using sets.sets_into_space by auto also have "... = A \ space M - (\n. (T^^n)-`(A-recurrent_subset A) \ space M)" unfolding recurrent_subset_infty_def by blast also have "... = A - (\n. (T^^n)--`(A-recurrent_subset A))" unfolding vimage_restr_def using sets.sets_into_space by auto finally have **: "recurrent_subset_infty A = A - (\n. (T ^^ n) --` (A - recurrent_subset A))" . then have "A - recurrent_subset_infty A \ (\n. (T^^n)--`(A-recurrent_subset A))" by auto with * ** show "A - recurrent_subset_infty A \ null_sets M" by (simp add: Diff_Diff_Int null_set_Int1) have "A \ recurrent_subset A = A - recurrent_subset A" using recurrent_subset_incl(1)[of A] by blast then show "A \ recurrent_subset A \ null_sets M" using \A - recurrent_subset A \ null_sets M\ by auto then show "emeasure M (recurrent_subset A) = emeasure M A" by (rule Delta_null_same_emeasure[symmetric], auto) have "A \ recurrent_subset_infty A = A - recurrent_subset_infty A" using recurrent_subset_incl(2)[of A] by blast then show "A \ recurrent_subset_infty A \ null_sets M" using \A - recurrent_subset_infty A \ null_sets M\ by auto then show "emeasure M (recurrent_subset_infty A) = emeasure M A" by (rule Delta_null_same_emeasure[symmetric], auto) show "AE x\A in M. x \ recurrent_subset_infty A" unfolding eventually_ae_filter by (metis (no_types, lifting) DiffI \A - recurrent_subset_infty A \ null_sets M\ mem_Collect_eq subsetI) qed text \A convenient way to use conservativity is given in the following theorem: if $T$ is conservative, then the series $\sum_n f(T^n x)$ is infinite for almost every $x$ with $f x > 0$. When $f$ is an indicator function, this is the fact that, starting from $B$, one returns infinitely many times to $B$ almost surely. The general case follows by approximating $f$ from below by constants time indicators.\ theorem (in conservative) recurrence_series_infinite: fixes f::"'a \ ennreal" assumes [measurable]: "f \ borel_measurable M" shows "AE x in M. f x > 0 \ (\n. f ((T^^n) x)) = \" proof - have *: "AE x in M. f x > epsilon \ (\n. f ((T^^n) x)) = \" if "epsilon > 0" for epsilon proof - define B where "B = {x \ space M. f x > epsilon}" have [measurable]: "B \ sets M" unfolding B_def by auto have "(\n. f ((T^^n) x)) = \" if "x \ recurrent_subset_infty B" for x proof - have "\ = epsilon * \" using \epsilon > 0\ ennreal_mult_top by auto also have "... = epsilon * (\n. indicator B ((T^^n) x))" using recurrent_subset_infty_series_infinite[OF that] by simp also have "... = (\n. epsilon * indicator B ((T^^n) x))" by auto also have "... \ (\n. f ((T^^n) x))" apply (rule suminf_le) unfolding indicator_def B_def by auto finally show ?thesis by (simp add: dual_order.antisym) qed moreover have "AE x in M. f x > epsilon \ x \ recurrent_subset_infty B" using Poincare_recurrence_thm(7)[OF \B \ sets M\] unfolding B_def by auto ultimately show ?thesis by auto qed have "\u::(nat \ ennreal). (\n. u n > 0) \ u \ 0" by (meson approx_from_above_dense_linorder ex_gt_or_lt gr_implies_not_zero) then obtain u::"nat \ ennreal" where u: "\n. u n > 0" "u \ 0" by auto have "AE x in M. (\n::nat. (f x > u n \ (\n. f ((T^^n) x)) = \))" unfolding AE_all_countable using u by (auto intro!: *) moreover have "f x > 0 \ (\n. f ((T^^n) x)) = \" if "(\n::nat. (f x > u n \ (\n. f ((T^^n) x)) = \))" for x proof (auto) assume "f x > 0" obtain n where "u n < f x" using order_tendstoD(2)[OF u(2) \f x > 0\] eventually_False_sequentially eventually_mono by blast then show "(\n. f ((T^^n) x)) = \" using that by auto qed ultimately show ?thesis by auto qed subsection \The first return time\ text \The first return time to a set $A$ under the dynamics $T$ is the smallest integer $n$ such that $T^n(x) \in A$. The first return time is only well defined on the recurrent subset of $A$, elsewhere we set it to $0$ for definiteness. We can partition $A$ according to the value of the return time on it, thus defining the return partition of $A$.\ definition return_time_function::"'a set \ ('a \ nat)" where "return_time_function A x = ( if (x \ recurrent_subset A) then (Inf {n::nat\{1..}. (T^^n) x \ A}) else 0)" definition return_partition::"'a set \ nat \ 'a set" where "return_partition A k = A \ (T^^k)--`A - (\i\{0<..Basic properties of the return partition.\ lemma return_partition_basics: assumes A_meas [measurable]: "A \ sets M" shows [measurable]: "return_partition A n \ sets M" and "disjoint_family (\n. return_partition A (n+1))" "(\n. return_partition A (n+1)) = recurrent_subset A" proof - show "return_partition A n \ sets M" for n unfolding return_partition_def by auto define B where "B = (\n. A \ (T^^(n+1))--`A)" have "return_partition A (n+1) = B n -(\i\{0..n. return_partition A (n+1) = disjointed B n" using disjointed_def[of B] by simp then show "disjoint_family (\n. return_partition A (n+1))" using disjoint_family_disjointed by simp have "A \ (T^^n)-`A = A \ (T^^n)--`A" for n using sets.sets_into_space[OF A_meas] by auto then have "recurrent_subset A = (\n\ {1..}. A \ (T^^n)--`A)" unfolding recurrent_subset_def by simp also have "... = (\n. B n)" by (simp add: B_def atLeast_Suc_greaterThan greaterThan_0) also have "... = (\n. return_partition A (n+1))" using * UN_disjointed_eq[of B] by simp finally show "(\n. return_partition A (n+1)) = recurrent_subset A" by simp qed text \Basic properties of the return time, relationship with the return partition.\ lemma return_time0: "(return_time_function A)-`{0} = UNIV - recurrent_subset A" proof (auto) fix x assume *: "x \ recurrent_subset A" "return_time_function A x = 0" define K where "K = {n::nat\{1..}. (T^^n) x \ A}" have **: "return_time_function A x = Inf K" using K_def return_time_function_def * by simp have "K \ {}" using K_def recurrent_subset_def * by auto moreover have "0 \ K" using K_def by auto ultimately have "Inf K >0" by (metis (no_types, lifting) K_def One_nat_def atLeast_iff cInf_lessD mem_Collect_eq neq0_conv not_le zero_less_Suc) then have "return_time_function A x > 0" using ** by simp then show "False" using * by simp qed (auto simp add: return_time_function_def) lemma return_time_n: assumes [measurable]: "A \ sets M" shows "(return_time_function A)-`{Suc n} = return_partition A (Suc n)" proof (auto) fix x assume *: "return_time_function A x = Suc n" then have rx: "x \ recurrent_subset A" using return_time_function_def by (auto, meson Zero_not_Suc) define K where "K = {i\{1..}. (T^^i) x \ A}" have "return_time_function A x = Inf K" using return_time_function_def rx K_def by auto then have "Inf K = Suc n" using * by simp moreover have "K \ {}" using rx recurrent_subset_def K_def by auto ultimately have "Suc n \ K" using Inf_nat_def1[of K] by simp then have "(T^^(Suc n))x \ A" using K_def by auto then have a: "x \ A \ (T^^(Suc n))--`A" using rx recurrent_subset_incl[of A] sets.sets_into_space[OF assms] by auto have "\i. i\{1.. i \ K" using cInf_lower \Inf K = Suc n\ by force then have "\i. i\{1.. x \ (T^^i)--`A" using K_def by auto then have "x \ (\i\{1.. return_partition A (Suc n)" using a return_partition_def by simp next fix x assume *: "x \ return_partition A (Suc n)" then have a: "x \ space M" unfolding return_partition_def using vimage_restr_def by blast define K where "K = {i::nat\{1..}. (T^^i) x \ A}" have "Inf K = Suc n" apply (rule cInf_eq_minimum) using * by (auto simp add: a assms K_def return_partition_def) have "x \ recurrent_subset A" using * return_partition_basics(3)[OF assms] by auto then show "return_time_function A x = Suc n" using return_time_function_def K_def \Inf K = Suc n\ by auto qed text \The return time is measurable.\ lemma return_time_function_meas [measurable]: assumes [measurable]: "A \ sets M" shows "return_time_function A \ measurable M (count_space UNIV)" "return_time_function A \ borel_measurable M" proof - have "(return_time_function A)-`{n} \ space M \ sets M" for n proof (cases "n = 0") case True then show ?thesis using return_time0 recurrent_subset_meas[OF assms] by auto next case False show ?thesis using return_time_n return_partition_basics(1)[OF assms] not0_implies_Suc[OF False] by auto qed then show "return_time_function A \ measurable M (count_space UNIV)" by (simp add: measurable_count_space_eq2_countable assms) then show "return_time_function A \ borel_measurable M" using measurable_cong_sets sets_borel_eq_count_space by blast qed text \A close cousin of the return time and the return partition is the first entrance set: we partition the space according to the first positive time where a point enters $A$.\ definition first_entrance_set::"'a set \ nat \ 'a set" where "first_entrance_set A n = (T^^n) --` A - (\ i sets M" shows "first_entrance_set A n \ sets M" unfolding first_entrance_set_def by measurable lemma first_entrance_disjoint: "disjoint_family (first_entrance_set A)" proof - have "first_entrance_set A = disjointed (\i. (T^^i)--`A)" by (auto simp add: disjointed_def first_entrance_set_def) then show ?thesis by (simp add: disjoint_family_disjointed) qed text \There is an important dynamical phenomenon: if a point has first entrance time equal to $n$, then their preimages either have first entrance time equal to $n+1$ (these are the preimages not in $A$) or they belong to $A$ and have first return time equal to $n+1$. When $T$ preserves the measure, this gives an inductive control on the measure of the first entrance set, that will be used again and again in the proof of Kac's Formula. We formulate these (simple but extremely useful) facts now.\ lemma first_entrance_rec: assumes [measurable]: "A \ sets M" shows "first_entrance_set A (Suc n) = T--`(first_entrance_set A n) - A" proof - have A0: "A = (T^^0)--`A" by auto have "first_entrance_set A n = (T^^n) --` A - (\ i iA \ sets M\ by simp then have *: "T--`(first_entrance_set A n) - A = (T^^(n+1))--`A - (A \ (\ i i j\{1.. (\ i j\{0.. sets M" shows "(return_time_function A)-`{Suc n} = T--`(first_entrance_set A n) \ A" proof - have "return_partition A (Suc n) = T--`(first_entrance_set A n) \ A" unfolding return_partition_def first_entrance_set_def by (auto simp add: T_vrestr_composed[OF assms]) (auto simp add: less_Suc_eq_0_disj) then show ?thesis using return_time_n[OF assms] by simp qed subsection \Local time controls\ text \The local time is the time that an orbit spends in a given set. Local time controls are basic to all the forthcoming developments.\ definition local_time::"'a set \ nat \ 'a \ nat" where "local_time A n x = card {i\{.. A}" lemma local_time_birkhoff: "local_time A n x = birkhoff_sum (indicator A) n x" proof (induction n) case 0 then show ?case unfolding local_time_def birkhoff_sum_def by simp next case (Suc n) have "local_time A (n+1) x = local_time A n x + indicator A ((T^^n) x)" proof (cases) assume *: "(T^^n) x \ A" then have "{i\{.. A} = {i\{.. A} \ {n}" by auto then have "card {i\{.. A} = card {i\{.. A} + card {n}" using card_Un_disjoint by auto then have "local_time A (n+1) x = local_time A n x + 1" using local_time_def by simp moreover have "indicator A ((T^^n)x) = (1::nat)" using * indicator_def by auto ultimately show ?thesis by simp next assume *: "\((T^^n) x \ A)" then have "{i\{.. A} = {i\{.. A}" using less_Suc_eq by force then have "card {i\{.. A} = card {i\{.. A}" by auto then have "local_time A (n+1) x = local_time A n x" using local_time_def by simp moreover have "indicator A ((T^^n)x) = (0::nat)" using * indicator_def by auto ultimately show ?thesis by simp qed then have "local_time A (n+1) x = birkhoff_sum (indicator A) n x + indicator A ((T^^n) x)" using Suc.IH by auto moreover have "birkhoff_sum (indicator A) (n+1) x = birkhoff_sum (indicator A) n x + indicator A ((T^^n) x)" by (metis birkhoff_sum_cocycle[where ?n = "n" and ?m = "1"] birkhoff_sum_1(2)) ultimately have "local_time A (n+1) x = birkhoff_sum (indicator A) (n+1) x" by metis then show ?case by (metis Suc_eq_plus1) qed lemma local_time_meas [measurable]: assumes [measurable]: "A \ sets M" shows "local_time A n \ borel_measurable M" unfolding local_time_birkhoff by auto lemma local_time_cocycle: "local_time A n x + local_time A m ((T^^n)x) = local_time A (n+m) x" by (metis local_time_birkhoff birkhoff_sum_cocycle) lemma local_time_incseq: "incseq (\n. local_time A n x)" using local_time_cocycle incseq_def by (metis le_iff_add) lemma local_time_Suc: "local_time A (n+1) x = local_time A n x + indicator A ((T^^n)x)" by (metis local_time_birkhoff birkhoff_sum_cocycle birkhoff_sum_1(2)) text \The local time is bounded by $n$: at most, one returns to $A$ all the time!\ lemma local_time_bound: "local_time A n x \ n" proof - have "card {i\{.. A} \ card {..The fact that local times are unbounded will be the main technical tool in the proof of recurrence results or Kac formula below. In this direction, we prove more and more general results in the lemmas below. We show that, in $T^{-n}(A)$, the number of visits to $A$ tends to infinity in measure, when $A$ has finite measure. In other words, the points in $T^{-n}(A)$ with local time $ lemma (in conservative_mpt) local_time_unbounded1: assumes A_meas [measurable]: "A \ sets M" and fin: "emeasure M A < \" shows "(\n. emeasure M {x \ (T^^n)--`A. local_time A n x < k}) \ 0" proof (induction k) case 0 have "{x \ (T^^n)--`A. local_time A n x < 0} = {}" for n by simp then show ?case by simp next case (Suc k) define K where "K = (\p n. {x \ (T^^n)--`A. local_time A n x < p})" have K_meas [measurable]: "K p n \ sets M" for n p unfolding K_def by measurable show ?case proof (rule tendsto_zero_ennreal) fix e :: real assume "0 < e" define e2 where "e2 = e/3" have "e2 > 0" using e2_def \e>0\ by simp have "(\n. emeasure M (return_partition A (n+1))) = emeasure M ((\n. return_partition A (n + 1)))" apply (rule suminf_emeasure) using return_partition_basics[OF A_meas] by auto also have "... = emeasure M (recurrent_subset A)" using return_partition_basics(3)[OF A_meas] by simp also have "... = emeasure M A" by (metis A_meas double_diff emeasure_Diff_null_set order_refl Poincare_recurrence_thm(1)[OF A_meas] recurrent_subset_incl(1)) finally have "(\n. emeasure M (return_partition A (n+1))) = emeasure M A" by simp moreover have "summable (\n. emeasure M (return_partition A (n+1)))" by simp ultimately have "(\N. (\n emeasure M A" unfolding sums_def[symmetric] sums_iff by simp then have "(\N. (\n emeasure M A + e2" by (intro tendsto_add) auto moreover have "emeasure M A < emeasure M A + e2" using \emeasure M A < \\ \0 < e2\ by auto ultimately have "eventually (\N. (\n emeasure M A) sequentially" by (simp add: order_tendsto_iff) then obtain N where "N>0" and largeM: "(\n emeasure M A" by (metis (no_types, lifting) add.commute add_Suc_right eventually_at_top_linorder le_add2 zero_less_Suc) have upper: "emeasure M (K (Suc k) n) \ e2 + (\iN" for n proof - define B where "B = (\i. (T^^(n-i-1))--`(return_partition A (i+1)))" have B_meas [measurable]: "B i \ sets M" for i unfolding B_def by measurable have disj_B: "disjoint_family_on B {.. B j = {}" if "i\{..{.. i" "n>j" using \n>N\ that by auto let ?k = "j-i" have "x \ B i" if "x \ B j" for x proof - have "(T^^(n-j-1)) x \ return_partition A (j+1)" using B_def that by auto moreover have "?k>0" using \i < j\ by simp moreover have "?k < j+1" by simp ultimately have "(T^^(n-j-1)) x \ (T^^?k)--`A" using return_partition_def by auto then have "x \ (T^^(n-j-1))--` (T^^?k)--`A" by auto then have "x \ (T^^(n-j-1 + ?k))--`A" using T_vrestr_composed[OF A_meas] by simp then have "x \ (T^^(n-i-1))--`A" using \i \n>j\ by auto then have "x \ (T^^(n-i-1))--` (return_partition A (i+1))" using return_partition_def by auto then show "x \ B i" using B_def by auto qed then show "B i \ B j = {}" by auto qed then have "\i j. i\{.. j\{.. i \ j \ B i \ B j = {}" by (metis Int_commute linorder_neqE_nat) then show ?thesis unfolding disjoint_family_on_def by auto qed have incl_B: "B i \ (T^^n)--`A" if "i \ {.. i" using \n>N\ that by auto have "B i \ (T^^(n-i-1))--` (T^^(i+1))--` A" using B_def return_partition_def by auto then show "B i \ (T^^n)--`A" using T_vrestr_composed(1)[OF A_meas, of "n-i-1", of "i+1"] \n>i\ by auto qed define R where "R = (T^^n)--`A - (\i \ {.. sets M" unfolding R_def by measurable have dec_n: "(T^^n)--`A = R \ (\i \ {.. (\i \ {..i \ {..i \ {..i \ {..i \ {..i \ {..i. emeasure M (B i) = emeasure M (return_partition A (i+1))" using T_vrestr_same_emeasure(2) B_def return_partition_basics(1)[OF A_meas] by simp ultimately have a: "emeasure M A = emeasure M R + (\i \ {..i \ {.. \" using fin by (simp add: a less_top) ultimately show ?thesis using largeM fin b by simp qed have "K (Suc k) n \ R \ (\i K (Suc k) n" show "x \ R \ (\i(x \ R)" have "x \ (T^^n)--`A" using a K_def by simp then have "x\ (\i \ {..\(x \ R)\ by simp then obtain i where "i\{.. B i" by auto then have "n>i" using \n>N\ by auto then have "(T^^(n-i-1)) x \ return_partition A (i+1)" using B_def \x \ B i\ by auto then have i: "(T^^(n-i-1)) x \ A" using return_partition_def by auto then have "indicator A ((T^^(n-i-1)) x) = (1::nat)" by auto then have "local_time A (n-i) x = local_time A (n-i-1) x + 1" by (metis Suc_diff_Suc Suc_eq_plus1 diff_diff_add local_time_Suc[of A, of "n-i-1"] \n>i\) then have "local_time A (n-i) x > local_time A (n-i-1) x" by simp moreover have "local_time A n x \ local_time A (n-i) x" using local_time_incseq by (metis \i < n\ le_add_diff_inverse2 less_or_eq_imp_le local_time_cocycle le_iff_add) ultimately have "local_time A n x > local_time A (n-i-1) x" by simp moreover have "local_time A n x < Suc k" using a K_def by simp ultimately have *: "local_time A (n-i-1) x < k" by simp have "x \ space M" using \x \ (T^^n)--`A\ by auto then have "x \ (T^^(n-i-1))--`A" using i A_meas vimage_restr_def by (metis IntI sets.Int_space_eq2 vimageI) then have "x \ K k (n-i-1)" using * K_def by blast then show ?thesis using \i\{.. by auto qed (simp) qed then have "emeasure M (K (Suc k) n) \ emeasure M (R \ (\i emeasure M R + emeasure M (\i emeasure M R + (\i e2 + (\i e2 + (\in. (\i\{.. (\i\{..n. (\i\{..e2 > 0\ by (simp add: order_tendsto_iff) then obtain N2 where N2bound: "\n. n > N2 \ (\i\{.. N3" for n proof - have "n>N2" "n > N" using N3_def that by auto then have "emeasure M (K (Suc k) n) \ ennreal e2 + (\i\{.. ennreal e2 + ennreal e2" using N2bound[OF \n > N2\] less_imp_le by auto also have "... < e" using e2_def \e > 0\ by (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus intro!: ennreal_lessI) ultimately show "emeasure M (K (Suc k) n) < e" using le_less_trans by blast qed then show "\\<^sub>F x in sequentially. emeasure M {xa \ (T ^^ x) --` A. local_time A x xa < Suc k} < ennreal e" unfolding K_def by (auto simp: eventually_at_top_dense intro!: exI[of _ N3]) qed qed text \We deduce that local times to a set $B$ also tend to infinity on $T^{-n}A$ if $B$ is related to $A$, i.e., if points in $A$ have some iterate in $B$. This is clearly a necessary condition for the lemmas to hold: otherwise, points of $A$ that never visit $B$ have a local time equal to $B$ equal to $0$, and so do all their preimages. The lemmas are readily reduced to the previous one on the local time to $A$, since if one visits $A$ then one visits $B$ in finite time by assumption (uniformly bounded in the first lemma, uniformly bounded on a set of large measure in the second lemma).\ lemma (in conservative_mpt) local_time_unbounded2: assumes A_meas [measurable]: "A \ sets M" and fin: "emeasure M A < \" and incl: "A \ (T^^i)--`B" shows "(\n. emeasure M {x \ (T^^n)--`A. local_time B n x < k}) \ 0" proof - have "emeasure M {x \ (T^^n)--`A. local_time B n x < k} \ emeasure M {x \ (T^^n)--`A. local_time A n x < k + i}" if "n > i" for n proof - have "local_time A n x \ local_time B n x + i" for x proof - have "local_time B n x \ local_time A (n-i) x" proof - define KA where "KA = {t \ {0.. A}" define KB where "KB = {t \ {0.. B}" then have "KB \ {0..t. t \ KA \ ?g t \ KB" proof - fix t assume "t \ KA" then have "(T^^t) x \ A" using KA_def by simp then have "(T^^i) ((T^^t) x) \ B" using incl by auto then have "(T^^(t+i)) x \ B" by (simp add: funpow_add add.commute) moreover have "t+i < n" using \t \ KA\ KA_def \n > i\ by auto ultimately show "?g t \ KB" unfolding KB_def by simp qed then have "?g`KA \ KB" by auto moreover have "inj_on ?g KA" by simp ultimately have "card KB \ card KA" using card_inj_on_le[where ?f = "?g" and ?A = KA and ?B = KB] \finite KB\ by simp then show ?thesis using KA_def KB_def local_time_def by simp qed moreover have "i \ local_time A i ((T^^(n-i))x)" using local_time_bound by auto ultimately show "local_time B n x + i \ local_time A n x" using local_time_cocycle[where ?n = "n-i" and ?m = i and ?x = x and ?A = A] \n>i\ by auto qed then have "local_time B n x < k \ local_time A n x < k + i" for x by (meson add_le_cancel_right le_trans not_less) then show ?thesis by (intro emeasure_mono, auto) qed then have "eventually (\n. emeasure M {x \ (T^^n)--`A. local_time B n x < k} \ emeasure M {x \ (T^^n)--`A. local_time A n x < k + i}) sequentially" using eventually_at_top_dense by blast from tendsto_sandwich[OF _ this tendsto_const local_time_unbounded1[OF A_meas fin, of "k+i"]] show ?thesis by auto qed lemma (in conservative_mpt) local_time_unbounded3: assumes A_meas[measurable]: "A \ sets M" and B_meas[measurable]: "B \ sets M" and fin: "emeasure M A < \" and incl: "A - (\i. (T^^i)--`B) \ null_sets M" shows "(\n. emeasure M {x \ (T^^n)--`A. local_time B n x < k}) \ 0" proof - define R where "R = A - (\i. (T^^i)--`B)" have R_meas[measurable]: "R \ sets M" by (simp add: A_meas B_meas T_vrestr_meas(2)[OF B_meas] R_def countable_Un_Int(1) sets.Diff) have "emeasure M R = 0" using incl R_def by auto define A2 where "A2 = A - R" have A2_meas [measurable]: "A2 \ sets M" unfolding A2_def by auto have meq: "emeasure M A2 = emeasure M A" using \emeasure M R = 0\ unfolding A2_def by (subst emeasure_Diff) (auto simp: R_def) then have A2_fin: "emeasure M A2 < \" using fin by auto define K where "K = (\N. A2 \ (\i sets M" for N unfolding K_def by auto have K_incl: "\N. K N \ A" using K_def A2_def by blast have "(\N. K N) = A2" using A2_def R_def K_def by blast moreover have "incseq K" unfolding K_def incseq_def by fastforce ultimately have "(\N. emeasure M (K N)) \ emeasure M A2" by (auto intro: Lim_emeasure_incseq) then have conv: "(\N. emeasure M (K N)) \ emeasure M A" using meq by simp define Bad where "Bad = (\U n. {x \ (T^^n)--`U. local_time B n x < k})" define Bad0 where "Bad0 = (\n. {x \ space M. local_time B n x < k})" have Bad0_meas [measurable]: "Bad0 n \ sets M" for n unfolding Bad0_def by auto have Bad_inter: "\U n. Bad U n = (T^^n)--`U \ Bad0 n" unfolding Bad_def Bad0_def by auto have Bad_meas [measurable]: "\U n. U \ sets M \ Bad U n \ sets M" unfolding Bad_def by auto show ?thesis proof (rule tendsto_zero_ennreal) fix e::real assume "e > 0" define e2 where "e2 = e/3" then have "e2 > 0" using \e>0\ by simp then have "ennreal e2 > 0" by simp have "(\N. emeasure M (K N) + e2) \ emeasure M A + e2" using conv by (intro tendsto_add) auto moreover have "emeasure M A < emeasure M A + e2" using fin \e2 > 0\ by simp ultimately have "eventually (\N. emeasure M (K N) + e2 > emeasure M A) sequentially" by (simp add: order_tendsto_iff) then obtain N where "N>0" and largeK: "emeasure M (K N) + e2 > emeasure M A" by (metis (no_types, lifting) add.commute add_Suc_right eventually_at_top_linorder le_add2 zero_less_Suc) define S where "S = A - (K N)" have S_meas [measurable]: "S \ sets M" using A_meas K_meas S_def by simp have "emeasure M A = emeasure M (K N) + emeasure M S" by (metis Diff_disjoint Diff_partition plus_emeasure[OF K_meas[of N], OF S_meas] S_def K_incl[of N]) then have S_small: "emeasure M S < e2" using largeK fin by simp have A_incl: "A \ S \ (\i (T^^i)--`B)" using S_def K_def by auto define L where "L = (\i. A2 \ (T^^i)--`B)" have L_meas [measurable]: "L i \ sets M" for i unfolding L_def by auto have "\i. L i \ A2" using L_def by simp then have L_fin: "emeasure M (L i) < \" for i using emeasure_mono[of "L i" A2 M] A2_meas A2_fin by simp have "\i. L i \ (T^^i)--`B" using L_def by auto then have a: "\i. (\n. emeasure M (Bad (L i) n)) \ 0" unfolding Bad_def using local_time_unbounded2[OF L_meas, OF L_fin] by blast have "(\n. (\i 0" using tendsto_sum[OF a] by auto then have "eventually (\n. (\iennreal e2 > 0\ order_tendsto_iff by metis then obtain N2 where *: "\n. n > N2 \ (\i N2" for n proof - have "emeasure M (Bad S n) \ emeasure M ((T^^n)--`S)" apply (rule emeasure_mono) unfolding Bad_def by auto also have "... = emeasure M S" using T_vrestr_same_emeasure(2) by simp also have "... \ e2" using S_small by simp finally have SBad_small: "emeasure M (Bad S n) \ e2" by simp have "(T^^n)--`A \ (T^^n)--`S \ (\i Bad S n \ (\i emeasure M (Bad S n \ (\i emeasure M (Bad S n) + emeasure M (\i emeasure M (Bad S n) + (\i ennreal e2 + ennreal e2" using SBad_small less_imp_le[OF *[OF \n > N2\]] by (rule add_mono) also have "... < e" using e2_def \e>0\ by (simp del: ennreal_plus add: ennreal_plus[symmetric] ennreal_lessI) finally show "emeasure M (Bad A n) < e" by simp qed then show "\\<^sub>F x in sequentially. emeasure M {xa \ (T ^^ x) --` A. local_time B x xa < k} < e" unfolding eventually_at_top_dense Bad_def by auto qed qed subsection \The induced map\ text \The map induced by $T$ on a set $A$ is obtained by iterating $T$ until one lands again in $A$. (Outside of $A$, we take the identity for definiteness.) It has very nice properties: if $T$ is conservative, then the induced map $T_A$ also is. If $T$ is measure preserving, then so is $T_A$. (In particular, even if $T$ preserves an infinite measure, $T_A$ is a probability preserving map if $A$ has measure $1$ -- this makes it possible to prove some statements in infinite measure by using results in finite measure systems). If $T$ is invertible, then so is $T_A$. We prove all these properties in this paragraph.\ definition induced_map::"'a set \ ('a \ 'a)" where "induced_map A = (\ x. (T^^(return_time_function A x)) x)" text \The set $A$ is stabilized by the induced map.\ lemma induced_map_stabilizes_A: "x \ A \ induced_map A x \ A" proof (rule equiv_neg) fix x assume "x \ A" show "induced_map A x \ A" proof (cases) assume "x \ recurrent_subset A" then have "induced_map A x = x" using induced_map_def return_time_function_def by simp then show ?thesis using \x \ A\ by simp next assume H: "\(x \ recurrent_subset A)" define K where "K = {n\{1..}. (T^^n) x \ A}" have "K \ {}" using H recurrent_subset_def K_def by blast moreover have "return_time_function A x = Inf K" using return_time_function_def K_def H by simp ultimately have "return_time_function A x \ K" using Inf_nat_def1 by simp then show ?thesis unfolding induced_map_def K_def by blast qed next fix x assume "x \ A" then have "x \ recurrent_subset A" using recurrent_subset_def by simp then have "induced_map A x = x" using induced_map_def return_time_function_def by simp then show "induced_map A x \ A" using \x \ A\ by simp qed lemma induced_map_iterates_stabilize_A: assumes "x \ A" shows "((induced_map A)^^n) x \ A" proof (induction n) case 0 show ?case using \x \ A\ by auto next case (Suc n) have "((induced_map A)^^(Suc n)) x = (induced_map A) (((induced_map A)^^n) x)" by auto then show ?case using Suc.IH induced_map_stabilizes_A by auto qed lemma induced_map_meas [measurable]: assumes [measurable]: "A \ sets M" shows "induced_map A \ measurable M M" unfolding induced_map_def by auto text \The iterates of the induced map are given by a power of the original map, where the power is the Birkhoff sum (for the induced map) of the first return time. This is obvious, but useful.\ lemma induced_map_iterates: "((induced_map A)^^n) x = (T^^(\i < n. return_time_function A ((induced_map A ^^i) x))) x" proof (induction n) case 0 show ?case by auto next case (Suc n) have "((induced_map A)^^(n+1)) x = induced_map A (((induced_map A)^^n) x)" by (simp add: funpow_add) also have "... = (T^^(return_time_function A (((induced_map A)^^n) x))) (((induced_map A)^^n) x)" using induced_map_def by auto also have "... = (T^^(return_time_function A (((induced_map A)^^n) x))) ((T^^(\i < n. return_time_function A ((induced_map A ^^i) x))) x)" using Suc.IH by auto also have "... = (T^^(return_time_function A (((induced_map A)^^n) x) + (\i < n. return_time_function A ((induced_map A ^^i) x)))) x" by (simp add: funpow_add) also have "... = (T^^(\i < Suc n. return_time_function A ((induced_map A ^^i) x))) x" by (simp add: add.commute) finally show ?case by simp qed lemma induced_map_stabilizes_recurrent_infty: assumes "x \ recurrent_subset_infty A" shows "((induced_map A)^^n) x \ recurrent_subset_infty A" proof - have "x \ A" using assms(1) recurrent_subset_incl(2) by auto define R where "R = (\i < n. return_time_function A ((induced_map A ^^i) x))" have *: "((induced_map A)^^n) x = (T^^R) x" unfolding R_def by (rule induced_map_iterates) moreover have "((induced_map A)^^n) x \ A" by (rule induced_map_iterates_stabilize_A, simp add: \x \ A\) ultimately have "(T^^R) x \ A" by simp then show ?thesis using recurrent_subset_infty_returns[OF assms] * by auto qed text \If $x \in A$, then its successive returns to $A$ are exactly given by the iterations of the induced map.\ lemma induced_map_returns: assumes "x \ A" shows "((T^^n) x \ A) \ (\N\n. n = (\i A" have "\y. y \ A \ (T^^n)y \ A \ \N\n. n = (\iN\n. n = (\i(n = 0)" then have "n > 0" by simp then have y_rec: "y \ recurrent_subset A" using \y \ A\ \(T^^n) y \ A\ recurrent_subset_def by auto then have *: "return_time_function A y > 0" by (metis DiffE insert_iff neq0_conv vimage_eq return_time0) define m where "m = return_time_function A y" have "m > 0" using * m_def by simp define K where "K = {t \ {1..}. (T ^^ t) y \ A}" have "n \ K" unfolding K_def using \n > 0\ \(T^^n)y \ A\ by simp then have "n \ Inf K" by (simp add: cInf_lower) moreover have "m = Inf K" unfolding m_def K_def return_time_function_def using y_rec by simp ultimately have "n \ m" by simp define z where "z = induced_map A y" have "z \ A" using \y \ A\ induced_map_stabilizes_A z_def by simp have "z = (T^^m) y" using induced_map_def y_rec z_def m_def by auto then have "(T^^(n-m)) z = (T^^n) y" using \n \ m\ funpow_add[of "n-m" m T, symmetric] by (metis comp_apply le_add_diff_inverse2) then have "(T^^(n-m)) z \ A" using \(T^^n) y \ A\ by simp moreover have "n-m < n" using \m > 0\ \n > 0\ by simp ultimately obtain N0 where "N0 \ n-m" "n-m = (\iz \ A\ "1.IH" by blast then have "n-m = (\ii. ((induced_map A)^^i) (induced_map A y) = ((induced_map A)^^(i+1)) y" by (metis Suc_eq_plus1 comp_apply funpow_Suc_right) ultimately have "n-m = (\ii \ {1..i. return_time_function A (((induced_map A)^^i) y)", of 0, of 1, of N0, symmetric] atLeast0LessThan by auto moreover have "m = (\i\{0..<1}. return_time_function A (((induced_map A)^^i) y))" using m_def by simp ultimately have "n = (\i\{0..<1}. return_time_function A (((induced_map A)^^i) y)) + (\i \ {1..n \ m\ by simp then have "n = (\i\{0.. n" using \N0 \ n-m\ \n - m < n\ by linarith ultimately show ?thesis by (metis atLeast0LessThan) qed qed then show "\N\n. n = (\ix \ A\ \(T^^n) x \ A\ by simp next assume "\N\n. n = (\ii A" using \x \ A\ induced_map_iterates_stabilize_A by auto qed text \If a map is conservative, then the induced map is still conservative. Note that this statement is not true if one replaces the word "conservative" with "qmpt": inducion only works well in conservative settings. For instance, the right translation on $\mathbb{Z}$ is qmpt, but the induced map on $\mathbb{N}$ (again the right translation) is not, since the measure of $\{0\}$ is nonzero, while its preimage, the empty set, has zero measure. To prove conservativity, given a subset $B$ of $A$, there exists some time $n$ such that $T^{-n} B \cap B$ has positive measure. But this time $n$ corresponds to some returns to $A$ for the induced map, so $T^{-n} B \cap B$ is included in $\bigcup_m T_A^{-m} B \cap B$, hence one of these sets must have positive measure. The fact that the map is qmpt is then deduced from the conservativity. \ proposition (in conservative) induced_map_conservative: assumes A_meas: "A \ sets M" shows "conservative (restrict_space M A) (induced_map A)" proof have "sigma_finite_measure M" by unfold_locales then have "sigma_finite_measure (restrict_space M A)" using sigma_finite_measure_restrict_space assms by auto then show "\Aa. countable Aa \ Aa \ sets (restrict_space M A) \ \Aa = space (restrict_space M A) \ (\a\Aa. emeasure (restrict_space M A) a \ \)" using sigma_finite_measure_def by auto have imp: "\B. (B \ sets M \ B \ A \ emeasure M B > 0) \ (\N>0. emeasure M (((induced_map A)^^N)-`B \ B) > 0)" proof - fix B assume assm: "B \ sets M \ B \ A \ emeasure M B > 0" then have "B \ A" by simp have inc: "(\n\{1..}. (T^^n)-`B \ B) \ (\N\{1..}. ((induced_map A)^^N)-` B \ B)" proof fix x assume "x \ (\n\{1..}. (T^^n)-`B \ B)" then obtain n where "n\{1..}" and *: "x \ (T^^n)-`B \ B" by auto then have "n > 0" by auto have "x \ A" "(T^^n) x \ A" using * \B \ A\ by auto then obtain N where **: "n = (\i B" using * by simp then have "x \ ((induced_map A)^^N)-` B \ B" using * by simp moreover have "N > 0" using ** \n > 0\ by (metis leD lessThan_iff less_nat_zero_code neq0_conv sum.neutral_const sum_mono) ultimately show "x \ (\N\{1..}. ((induced_map A)^^N)-` B \ B)" by auto qed have B_meas [measurable]: "B \ sets M" and B_pos: "emeasure M B > 0" using assm by auto obtain n where "n > 0" and pos: "emeasure M ((T^^n)-`B \ B) > 0" using conservative[OF B_meas, OF B_pos] by auto then have "n \ {1..}" by auto have itB_meas: "\i. ((induced_map A)^^i)-` B \ B \ sets M" using B_meas measurable_compose_n[OF induced_map_meas[OF A_meas]] by (metis Int_assoc measurable_sets sets.Int sets.Int_space_eq1) then have "(\i\{1..}. ((induced_map A)^^i)-` B \ B) \ sets M" by measurable moreover have "(T^^n)-`B \ B \ (\i\{1..}. ((induced_map A)^^i)-` B \ B)" using inc \n \ {1..}\ by force ultimately have "emeasure M (\i\{1..}. ((induced_map A)^^i)-` B \ B) > 0" by (metis (no_types, lifting) emeasure_eq_0 zero_less_iff_neq_zero pos) then have "emeasure M (\i\{1..}. ((induced_map A)^^i)-` B \ B) \ 0" by simp have "\i\{1..}. emeasure M (((induced_map A)^^i)-` B \ B) \ 0" proof (rule ccontr) assume "\(\i\{1..}. emeasure M (((induced_map A)^^i)-` B \ B) \ 0)" then have a: "\i. i \ {1..} \ ((induced_map A)^^i)-` B \ B \ null_sets M" using itB_meas by auto have "(\i\{1..}. ((induced_map A)^^i)-` B \ B) \ null_sets M" by (rule null_sets_UN', simp_all add: a) then show "False" using \emeasure M (\i\{1..}. ((induced_map A)^^i)-` B \ B) > 0\ by auto qed then show "\N>0. emeasure M (((induced_map A)^^N)-` B \ B) > 0" by (simp add: Bex_def less_eq_Suc_le zero_less_iff_neq_zero) qed define K where "K = {B. B \ sets M \ B \ A}" have K_stable: "(induced_map A)-`B \ K" if "B \ K" for B proof - have B_meas: "B \ sets M" and "B \ A" using that unfolding K_def by auto then have a: "(induced_map A)-`B \ A" using induced_map_stabilizes_A by auto then have "(induced_map A)-`B = (induced_map A)-`B \ space M" using assms sets.sets_into_space by auto then have "(induced_map A)-`B \ sets M" using induced_map_meas[OF assms] B_meas by (metis vrestr_meas vrestr_of_set) then show "(induced_map A)-`B \ K" unfolding K_def using a by auto qed define K0 where "K0 = K \ (null_sets M)" have K0_stable: "(induced_map A)-`B \ K0" if "B \ K0" for B proof - have "B \ K" using that unfolding K0_def by simp then have a: "(induced_map A)-`B \ A" and b: "(induced_map A)-`B \ sets M" using K_stable unfolding K_def by auto have B_meas [measurable]: "B \ sets M" using \B \ K\ unfolding K_def by simp have B0: "B \ null_sets M" using \B \ K0\ unfolding K0_def by simp have "(induced_map A)-`B \ (\n. (T^^n)-`B)" unfolding induced_map_def by auto then have "(induced_map A)-`B \ (\n. (T^^n)-`B \ space M)" using b sets.sets_into_space by simp blast then have inc: "(induced_map A)-`B \ (\n. (T^^n)--`B)" unfolding vimage_restr_def using sets.sets_into_space [OF B_meas] by simp have "(T^^n)--`B \ null_sets M" for n using B0 T_quasi_preserves_null(2)[OF B_meas] by simp then have "(\n. (T^^n)--`B) \ null_sets M" using null_sets_UN by auto then have "(induced_map A)-`B \ null_sets M" using null_sets_subset[OF _ b inc] by auto then show "(induced_map A)-`B \ K0" unfolding K0_def K_def by (simp add: a b) qed have *: "D \ null_sets M \ D \ null_sets (restrict_space M A)" if "D\K" for D using that unfolding K_def apply auto apply (metis assms emeasure_restrict_space null_setsD1 null_setsI sets.Int_space_eq2 sets_restrict_space_iff) by (metis assms emeasure_restrict_space null_setsD1 null_setsI sets.Int_space_eq2) show "induced_map A \ quasi_measure_preserving (restrict_space M A) (restrict_space M A)" unfolding quasi_measure_preserving_def proof (auto) have "induced_map A \ A \ A" using induced_map_stabilizes_A by auto then show a: "induced_map A \ measurable (restrict_space M A) (restrict_space M A)" using measurable_restrict_space3[where ?A = A and ?B = A and ?M = M and ?N = M] induced_map_meas[OF A_meas] by auto fix B assume H: "B \ sets (restrict_space M A)" "induced_map A -`B \ space (restrict_space M A) \ null_sets (restrict_space M A)" then have "B \ K" unfolding K_def by (metis assms mem_Collect_eq sets.Int_space_eq2 sets_restrict_space_iff) then have B_meas [measurable]: "B \ sets M" and B_incl: "B \ A" unfolding K_def by auto have "induced_map A -`B \ K" using K_stable \B \ K\ by auto then have B2_meas: "induced_map A -`B \ sets M" and B2_incl: "induced_map A -`B \ A" unfolding K_def by auto have "induced_map A -` B = induced_map A -`B \ space (restrict_space M A)" using B2_incl by (simp add: Int_absorb2 assms space_restrict_space) then have "induced_map A -` B \ null_sets (restrict_space M A)" using H(2) by simp then have "induced_map A -` B \ K0" unfolding K0_def using \induced_map A -`B \ K\ * by auto { fix n have *: "((induced_map A)^^(n+1))-`B \ K0" proof (induction n) case (Suc n) have "((induced_map A)^^(Suc n+1))-`B = (induced_map A)-`(((induced_map A)^^(n+1))-` B)" by (metis Suc_eq_plus1 funpow_Suc_right vimage_comp) then show ?case by (metis Suc.IH K0_stable) qed (auto simp add: \induced_map A -` B \ K0\) have **: "((induced_map A)^^(n+1))-` B \ sets M" using * K0_def K_def by auto have "((induced_map A)^^(n+1))-` B \ B \ null_sets M" apply (rule null_sets_subset[of "((induced_map A)^^(n+1))-`B"]) using * unfolding K0_def apply simp using ** by auto } then have "((induced_map A)^^n)-` B \ B \ null_sets M" if "n>0" for n using that by (metis Suc_eq_plus1 neq0_conv not0_implies_Suc) then have "B \ null_sets M" using imp B_incl B_meas zero_less_iff_neq_zero inf.strict_order_iff by (metis null_setsD1 null_setsI) then show "B \ null_sets (restrict_space M A)" using * \B \ K\ by auto next fix B assume H: "B \ sets (restrict_space M A)" "B \ null_sets (restrict_space M A)" then have "B \ K" unfolding K_def by (metis assms mem_Collect_eq sets.Int_space_eq2 sets_restrict_space_iff) then have B_meas [measurable]: "B \ sets M" and B_incl: "B \ A" unfolding K_def by auto have "B \ null_sets M" using * H(2) \B \ K\ by simp then have "B \ K0" unfolding K0_def using \B \ K\ by simp then have inK: "(induced_map A)-`B \ K0" using K0_stable by auto then have inA: "(induced_map A)-`B \ A" unfolding K0_def K_def by auto then have "(induced_map A)-`B = (induced_map A)-`B \ space (restrict_space M A)" by (simp add: Int_absorb2 assms space_restrict_space2) then show "induced_map A -` B \ space (restrict_space M A) \ null_sets (restrict_space M A)" using * inK unfolding K0_def by auto qed fix B assume B_measA: "B \ sets (restrict_space M A)" and B_posA: "0 < emeasure (restrict_space M A) B" then have B_meas: "B \ sets M" by (metis assms sets.Int_space_eq2 sets_restrict_space_iff) have B_incl: "B \ A" by (metis B_measA assms sets.Int_space_eq2 sets_restrict_space_iff) then have B_pos: "0 < emeasure M B" using B_posA by (simp add: assms emeasure_restrict_space) obtain N where "N>0" "emeasure M (((induced_map A)^^N)-`B \ B) > 0" using imp B_meas B_incl B_pos by auto then have "emeasure (restrict_space M A) ((induced_map A ^^ N) -` B \ B) > 0" using assms emeasure_restrict_space by (metis B_incl Int_lower2 sets.Int_space_eq2 subset_trans) then show "\n>0. 0 < emeasure (restrict_space M A) ((induced_map A ^^ n) -` B \ B)" using \N > 0\ by auto qed text \Now, we want to prove that, if a map is conservative and measure preserving, then the induced map is also measure preserving. We first prove it for subsets $W$ of $A$ of finite measure, the general case will readily follow. The argument uses the fact that the preimage of the set of points with first entrance time $n$ is the union of the set of points with first entrance time $n+1$, and the points of $A$ with first return $n+1$. Following the preimage of $W$ under this process, we will get the intersection of $T_A^{-1} W$ with the different elements of the return partition, and the points in $T^{-n}W$ whose first $n-1$ iterates do not meet $A$ (and the measures of these sets add up to $\mu(W)$). To conclude, it suffices to show that the measure of points in $T^{-n}W$ whose first $n-1$ iterates do not meet $A$ tends to $0$. This follows from our local times estimates above.\ lemma (in conservative_mpt) induced_map_measure_preserving_aux: assumes A_meas [measurable]: "A \ sets M" and W_meas [measurable]: "W \ sets M" and incl: "W \ A" and fin: "emeasure M W < \" shows "emeasure M ((induced_map A)--`W) = emeasure M W" proof - have "W \ space M" using W_meas using sets.sets_into_space by blast define BW where "BW = (\n. (first_entrance_set A n) \ (T^^n)--`W)" define DW where "DW = (\n. (return_time_function A)-` {n} \ (induced_map A)--`W)" have "\n. DW n = (return_time_function A)-` {n} \ space M \ (induced_map A)--`W" using DW_def by auto then have DW_meas [measurable]: "\n. DW n \ sets M" by auto have disj_DW: "disjoint_family (\n. DW n)" using DW_def disjoint_family_on_def by blast then have disj_DW2: "disjoint_family (\n. DW (n+1))" by (simp add: disjoint_family_on_def) have "(\n. DW n) = DW 0 \ (\n. DW (n+1))" by (auto) (metis not0_implies_Suc) moreover have "(DW 0) \ (\n. DW (n+1)) = {}" by (auto) (metis IntI Suc_neq_Zero UNIV_I empty_iff disj_DW disjoint_family_on_def) ultimately have *: "emeasure M (\n. DW n) = emeasure M (DW 0) + emeasure M (\n. DW (n+1))" by (simp add: countable_Un_Int(1) plus_emeasure) have "DW 0 = (return_time_function A)-` {0} \ W" unfolding DW_def induced_map_def return_time_function_def apply (auto simp add: return_time0[of A]) using sets.sets_into_space[OF W_meas] by auto also have "... = W - recurrent_subset A" using return_time0 by blast also have "... \ A - recurrent_subset A" using incl by blast finally have "DW 0 \ null_sets M" by (metis A_meas DW_meas null_sets_subset Poincare_recurrence_thm(1)) then have "emeasure M (DW 0) = 0" by auto have "(induced_map A)--`W = (\n. DW n)" using DW_def by blast then have "emeasure M ((induced_map A)--`W) = emeasure M (\n. DW n)" by simp also have "... = emeasure M (\n. DW (n+1))" using * \emeasure M (DW 0) = 0\ by simp also have "... = (\n. emeasure M (DW (n+1)))" apply (rule suminf_emeasure[symmetric]) using disj_DW2 by auto finally have m: "emeasure M ((induced_map A)--`W) = (\n. emeasure M (DW (n+1)))" by simp moreover have "summable (\n. emeasure M (DW (n+1)))" by simp ultimately have lim: "(\N. (\ n\{.. emeasure M ((induced_map A)--`W)" by (simp add: summable_LIMSEQ) have BW_meas [measurable]: "\n. BW n \ sets M" unfolding BW_def by simp have *: "\n. T--`(BW n) - A = BW (n+1)" proof - fix n have "T--`(BW n) = T--`(first_entrance_set A n) \ (T^^(n+1))--`W" unfolding BW_def by (simp add: assms(2) T_vrestr_composed(2)) then have "T--`(BW n) - A = (T--`(first_entrance_set A n) - A) \ (T^^(n+1))--`W" by blast then have "T--`(BW n) - A = first_entrance_set A (n+1) \ (T^^(n+1))--`W" using first_entrance_rec[OF A_meas] by simp then show "T--`(BW n) - A = BW (n+1)" using BW_def by simp qed have **: "DW (n+1) = T--`(BW n) \ A" for n proof - have "T--`(BW n) = T--`(first_entrance_set A n) \ (T^^(n+1))--`W" unfolding BW_def by (simp add: assms(2) T_vrestr_composed(2)) then have "T--`(BW n) \ A = (T--`(first_entrance_set A n) \ A) \ (T^^(n+1))--`W" by blast then have *: "T--`(BW n) \ A = (return_time_function A)-`{n+1} \ (T^^(n+1))--`W" using return_time_rec[OF A_meas] by simp have "DW (n+1) = (return_time_function A)-`{n+1} \ (induced_map A)-`W" using DW_def \W \ space M\ return_time_rec by auto also have "... = (return_time_function A)-`{n+1} \ (T^^(n+1))-`W" by (auto simp add: induced_map_def) also have "... = (return_time_function A)-`{n+1} \ (T^^(n+1))--`W" using \W \ space M\ return_time_rec by auto finally show "DW (n+1) = T--`(BW n) \ A" using * by simp qed have "emeasure M W = (\ n\{.. DW (N+1)" using * ** by blast moreover have "BW (N+1) \ DW (N+1) = {}" using * ** by blast ultimately have "emeasure M (T--`(BW N)) = emeasure M (BW (N+1)) + emeasure M (DW (N+1))" using DW_meas BW_meas plus_emeasure[of "BW (N+1)"] by simp then have "emeasure M (BW N) = emeasure M (BW (N+1)) + emeasure M (DW (N+1))" using T_vrestr_same_emeasure(1) BW_meas by auto then have "(\ n\{.. n\{..N. emeasure M (BW N)) \ 0" proof (rule tendsto_sandwich[of "\_. 0"_ _ "\N. emeasure M {x \ (T^^N)--`W. local_time A N x < 1}"]) have "emeasure M (BW N) \ emeasure M {x \ (T^^N)--`W. local_time A N x < 1}" for N apply (rule emeasure_mono) unfolding BW_def local_time_def first_entrance_set_def by auto then show "\\<^sub>F n in sequentially. emeasure M (BW n) \ emeasure M {x \ (T ^^ n) --` W. local_time A n x < 1}" by auto have i: "W \ (T^^0)--`A" using incl by auto show "(\N. emeasure M {x \ (T ^^ N) --` W. local_time A N x < 1}) \ 0" apply (rule local_time_unbounded2[OF _ _ i]) using fin by auto qed (auto) then have "(\N. (\ n\{.. emeasure M (induced_map A --` W) + 0" using lim by (intro tendsto_add) auto ultimately show ?thesis by (auto intro: LIMSEQ_unique LIMSEQ_const_iff) qed lemma (in conservative_mpt) induced_map_measure_preserving: assumes A_meas [measurable]: "A \ sets M" and W_meas [measurable]: "W \ sets M" shows "emeasure M ((induced_map A)--`W) = emeasure M W" proof - define WA where "WA = W \ A" have WA_meas [measurable]: "WA \ sets M" "WA \ A" using WA_def by auto have WAi_meas [measurable]: "(induced_map A)--`WA \ sets M" by simp have a: "emeasure M WA = emeasure M ((induced_map A)--`WA)" proof (cases) assume "emeasure M WA < \" then show ?thesis using induced_map_measure_preserving_aux[OF A_meas, OF \WA \ sets M\, OF \WA \ A\] by simp next assume "\(emeasure M WA < \)" then have "emeasure M WA = \" by (simp add: less_top[symmetric]) { fix C::real obtain Z where "Z \ sets M" "Z \ WA" "emeasure M Z < \" "emeasure M Z > C" by (blast intro: \emeasure M WA = \\ WA_meas approx_PInf_emeasure_with_finite) have "Z \ A" using \Z \ WA\ WA_def by simp have "C < emeasure M Z" using \emeasure M Z > C\ by simp also have "... = emeasure M ((induced_map A)--`Z)" using induced_map_measure_preserving_aux[OF A_meas, OF \Z \ sets M\, OF \Z \ A\] \emeasure M Z < \\ by simp also have "... \ emeasure M ((induced_map A)--`WA)" apply(rule emeasure_mono) using \Z \ WA\ vrestr_inclusion by auto finally have "emeasure M ((induced_map A)--`WA) > C" by simp } then have "emeasure M ((induced_map A)--`WA) = \" by (cases "emeasure M ((induced_map A)--`WA)") auto then show ?thesis using \emeasure M WA = \\ by simp qed define WB where "WB = W - WA" have WB_meas [measurable]: "WB \ sets M" unfolding WB_def by simp have WBi_meas [measurable]: "(induced_map A)--`WB \ sets M" by simp have "WB \ A = {}" unfolding WB_def WA_def by auto moreover have id: "\x. x \ A \ (induced_map A x) = x" unfolding induced_map_def return_time_function_def apply (auto) using recurrent_subset_incl by auto ultimately have "(induced_map A)--`WB = WB" using induced_map_stabilizes_A sets.sets_into_space[OF WB_meas] apply auto by (metis disjoint_iff_not_equal) fastforce+ then have b: "emeasure M ((induced_map A)--`WB) = emeasure M WB" by simp have "W = WA \ WB" "WA \ WB = {}" using WA_def WB_def by auto have *: "emeasure M W = emeasure M WA + emeasure M WB" by (subst \W = WA \ WB\, rule plus_emeasure[symmetric], auto simp add: \WA \ WB = {}\) have W_AUB: "(induced_map A)--`W = (induced_map A)--`WA \ (induced_map A)--`WB" using \W = WA \ WB\ by auto have W_AIB: "(induced_map A)--`WA \ (induced_map A)--`WB = {}" by (metis \WA \ WB = {}\ vrestr_empty vrestr_intersec) have "emeasure M ((induced_map A)--`W) = emeasure M ((induced_map A)--`WA) + emeasure M ((induced_map A)--`WB)" unfolding W_AUB by (rule plus_emeasure[symmetric]) (auto simp add: W_AIB) then show ?thesis using a b * by simp qed text \We can now express the fact that induced maps preserve the measure.\ theorem (in conservative_mpt) induced_map_conservative_mpt: assumes "A \ sets M" shows "conservative_mpt (restrict_space M A) (induced_map A)" unfolding conservative_mpt_def proof show *: "conservative (restrict_space M A) (induced_map A)" using induced_map_conservative[OF assms] by auto show "mpt (restrict_space M A) (induced_map A)" unfolding mpt_def mpt_axioms_def proof show "qmpt (restrict_space M A) (induced_map A)" using * conservative_def by auto then have meas: "(induced_map A) \ measurable (restrict_space M A) (restrict_space M A)" unfolding qmpt_def qmpt_axioms_def quasi_measure_preserving_def by auto moreover have "\B. B \ sets (restrict_space M A) \ emeasure (restrict_space M A) ((induced_map A) -`B \ space (restrict_space M A)) = emeasure (restrict_space M A) B" proof - have s: "space (restrict_space M A) = A" using assms space_restrict_space2 by auto have i: "\D. D \ sets M \ D \ A \ emeasure (restrict_space M A) D = emeasure M D" using assms by (simp add: emeasure_restrict_space) have j: "\D. D \ sets (restrict_space M A) \ (D \ sets M \ D \ A)" using assms by (metis sets.Int_space_eq2 sets_restrict_space_iff) fix B assume a: "B \ sets (restrict_space M A)" then have B_meas: "B \ sets M" using j by auto then have first: "emeasure (restrict_space M A) B = emeasure M B" using i j a by auto have incl: "(induced_map A) -`B \ A" using j a induced_map_stabilizes_A assms by auto then have eq: "(induced_map A) -`B \ space (restrict_space M A) = (induced_map A) --`B" unfolding vimage_restr_def s using assms sets.sets_into_space by (metis a inf.orderE j meas measurable_sets s) then have "emeasure M B = emeasure M ((induced_map A) -`B \ space (restrict_space M A))" using induced_map_measure_preserving a j assms by auto also have "... = emeasure (restrict_space M A) ((induced_map A) -`B \ space (restrict_space M A))" using incl eq B_meas induced_map_meas[OF assms] assms i j by (metis emeasure_restrict_space inf.orderE s space_restrict_space) finally show "emeasure (restrict_space M A) ((induced_map A) -`B \ space (restrict_space M A)) = emeasure (restrict_space M A) B" using first by auto qed ultimately show "induced_map A \ measure_preserving (restrict_space M A) (restrict_space M A)" unfolding measure_preserving_def by auto qed qed theorem (in fmpt) induced_map_fmpt: assumes "A \ sets M" shows "fmpt (restrict_space M A) (induced_map A)" unfolding fmpt_def proof - have "conservative_mpt (restrict_space M A) (induced_map A)" using induced_map_conservative_mpt[OF assms] by simp then have "mpt (restrict_space M A) (induced_map A)" using conservative_mpt_def by auto moreover have "finite_measure (restrict_space M A)" by (simp add: assms finite_measureI finite_measure_restrict_space) ultimately show "mpt (restrict_space M A) (induced_map A) \ finite_measure (restrict_space M A)" by simp qed text \It will be useful to reformulate the fact that the recurrent subset has full measure in terms of the induced measure, to simplify the use of the induced map later on.\ lemma (in conservative) induced_map_recurrent_typical: assumes A_meas [measurable]: "A \ sets M" shows "AE z in (restrict_space M A). z \ recurrent_subset A" "AE z in (restrict_space M A). z \ recurrent_subset_infty A" proof - have "recurrent_subset A \ sets M" using recurrent_subset_meas[OF A_meas] by auto then have rsA: "recurrent_subset A \ sets (restrict_space M A)" using recurrent_subset_incl(1)[of A] by (metis (no_types, lifting) A_meas sets_restrict_space_iff space_restrict_space space_restrict_space2) have "emeasure (restrict_space M A) (space (restrict_space M A) - recurrent_subset A) = emeasure (restrict_space M A) (A - recurrent_subset A)" by (metis (no_types, lifting) A_meas space_restrict_space2) also have "... = emeasure M (A - recurrent_subset A)" by (simp add: emeasure_restrict_space) also have "... = 0" using Poincare_recurrence_thm[OF A_meas] by auto finally have "space (restrict_space M A) - recurrent_subset A \ null_sets (restrict_space M A)" using rsA by blast then show "AE z in (restrict_space M A). z \ recurrent_subset A" by (metis (no_types, lifting) DiffI eventually_ae_filter mem_Collect_eq subsetI) have "recurrent_subset_infty A \ sets M" using recurrent_subset_meas[OF A_meas] by auto then have rsiA: "recurrent_subset_infty A \ sets (restrict_space M A)" using recurrent_subset_incl(2)[of A] by (metis (no_types, lifting) A_meas sets_restrict_space_iff space_restrict_space space_restrict_space2) have "emeasure (restrict_space M A) (space (restrict_space M A) - recurrent_subset_infty A) = emeasure (restrict_space M A) (A - recurrent_subset_infty A)" by (metis (no_types, lifting) A_meas space_restrict_space2) also have "... = emeasure M (A - recurrent_subset_infty A)" apply (rule emeasure_restrict_space) using A_meas by auto also have "... = 0" using Poincare_recurrence_thm[OF A_meas] by auto finally have "space (restrict_space M A) - recurrent_subset_infty A \ null_sets (restrict_space M A)" using rsiA by blast then show "AE z in (restrict_space M A). z \ recurrent_subset_infty A" by (metis (no_types, lifting) DiffI eventually_ae_filter mem_Collect_eq subsetI) qed subsection \Kac's theorem, and variants\ text \Kac's theorem states that, for conservative maps, the integral of the return time to a subset $A$ is equal to the measure of the space if the dynamics is ergodic, or of the space seen by $A$ in the general case. This result generalizes to any induced function, not just the return time, that we define now.\ definition induced_function::"'a set \ ('a \ 'b::comm_monoid_add) \ ('a \ 'b)" where "induced_function A f = (\x. (\i\{..< return_time_function A x}. f((T^^i) x)))" text \By definition, the induced function is supported on the recurrent subset of $A$.\ lemma induced_function_support: fixes f::"'a \ ennreal" shows "induced_function A f y = induced_function A f y * indicator ((return_time_function A)-`{1..}) y" by (auto simp add: induced_function_def indicator_def not_less_eq_eq) text \Basic measurability statements.\ lemma induced_function_meas_ennreal [measurable]: fixes f::"'a \ ennreal" assumes [measurable]: "f \ borel_measurable M" "A \ sets M" shows "induced_function A f \ borel_measurable M" unfolding induced_function_def by simp lemma induced_function_meas_real [measurable]: fixes f::"'a \ real" assumes [measurable]: "f \ borel_measurable M" "A \ sets M" shows "induced_function A f \ borel_measurable M" unfolding induced_function_def by simp text \The Birkhoff sums of the induced function for the induced map form a subsequence of the original Birkhoff sums for the original map, corresponding to the return times to $A$.\ lemma (in conservative) induced_function_birkhoff_sum: fixes f::"'a \ real" assumes "A \ sets M" shows "birkhoff_sum f (qmpt.birkhoff_sum (induced_map A) (return_time_function A) n x) x = qmpt.birkhoff_sum (induced_map A) (induced_function A f) n x" proof - interpret A: conservative "restrict_space M A" "induced_map A" by (rule induced_map_conservative[OF assms]) define TA where "TA = induced_map A" define phiA where "phiA = return_time_function A" define R where "R = (\n. A.birkhoff_sum phiA n x)" show ?thesis proof (induction n) case 0 show ?case using birkhoff_sum_1(1) A.birkhoff_sum_1(1) by auto next case (Suc n) have "(T^^(R n)) x = (TA^^n) x" unfolding TA_def R_def A.birkhoff_sum_def phiA_def by (rule induced_map_iterates[symmetric]) have "R(n+1) = R n + phiA ((TA^^n) x)" unfolding R_def using A.birkhoff_sum_cocycle[where ?n = n and ?m = 1 and ?f = phiA] A.birkhoff_sum_1(2) TA_def by auto then have "birkhoff_sum f (R (n+1)) x = birkhoff_sum f (R n) x + birkhoff_sum f (phiA ((TA^^n) x)) ((T^^(R n)) x)" using birkhoff_sum_cocycle[where ?n = "R n" and ?f = f] by auto also have "... = birkhoff_sum f (R n) x + birkhoff_sum f (phiA ((TA^^n) x)) ((TA^^n) x)" using \(T^^(R n)) x = (TA^^n) x\ by simp also have "... = birkhoff_sum f (R n) x + (induced_function A f) ((TA^^n) x)" unfolding induced_function_def birkhoff_sum_def phiA_def by simp also have "... = A.birkhoff_sum (induced_function A f) n x + (induced_function A f) ((TA^^n) x)" using Suc.IH R_def phiA_def by auto also have "... = A.birkhoff_sum (induced_function A f) (n+1) x" using A.birkhoff_sum_cocycle[where ?n = n and ?m = 1 and ?f = "induced_function A f" and ?x = x, symmetric] A.birkhoff_sum_1(2)[where ?f = "induced_function A f" and ?x = "(TA^^n) x"] unfolding TA_def by auto finally show ?case unfolding R_def phiA_def by simp qed qed text \The next lemma is very simple (just a change of variables to reorder the indices in the double sum). However, the proof I give is very tedious: infinite sums on proper subsets are not handled well, hence I use integrals on products of discrete spaces instead, and go back and forth between the two notions -- maybe there are better suited tools in the library, but I could not locate them... This is the main combinatorial tool used in the proof of Kac's Formula.\ lemma kac_series_aux: fixes d:: "nat \ nat \ ennreal" shows "(\n. (\i\n. d i n)) = (\n. d 0 n) + (\n. (\i. d (i+1) (n+1+i)))" (is "_ = ?R") proof - define g where "g = (\(i,n). (i+(1::nat), n+1+i))" define U where "U = {(i,n). (i>(0::nat)) \ (n\i)}" have bij: "bij_betw g UNIV U" by (rule bij_betw_byWitness[where ?f' = "\(i, n). (i-1, n-i)"], auto simp add: g_def U_def) define e where "e = (\ (i,j). d i j)" have pos: "\x. e x \ 0" using e_def by auto have "(\n. (\i. d (i+1) (n+1+i))) = (\n. (\i. e(i+1, n+1+i)))" using e_def by simp also have "... = \\<^sup>+n. \\<^sup>+i. e (i+1, n+1+i) \count_space UNIV \count_space UNIV" using pos nn_integral_count_space_nat suminf_0_le by auto also have "... = (\\<^sup>+x. e (g x) \count_space UNIV)" unfolding g_def using nn_integral_snd_count_space[of "\(i,n). e(i+1, n+1+i)"] by (auto simp add: prod.case_distrib) also have "... = (\\<^sup>+y \ U. e y \count_space UNIV)" using nn_integral_count_compose_bij[OF bij] by simp finally have *: "(\n. (\i. d (i+1) (n+1+i))) = (\\<^sup>+y \ U. e y \count_space UNIV)" by simp define V where "V = {((i::nat),(n::nat)). i = 0}" have i: "e (i, n) * indicator {0} i = e (i, n) * indicator V (i, n)" for i n by (auto simp add: indicator_def V_def) have "d 0 n = (\\<^sup>+i \ {0}. e (i, n) \count_space UNIV)" for n proof - have "(\\<^sup>+i \ {0}. e (i, n) \count_space UNIV) = (\\<^sup>+i. e (i, n) \count_space {0})" using nn_integral_count_space_indicator[of _ "\i. e(i, n)"] by simp also have "... = e (0, n)" using nn_integral_count_space_finite[where ?f = "\i. e (i, n)"] by simp finally show ?thesis using e_def by simp qed then have "(\n. d 0 n) = (\n. (\\<^sup>+i. e (i, n) * indicator {0} i \count_space UNIV))" by simp also have "... = (\\<^sup>+n. (\\<^sup>+i. e (i, n) * indicator {0} i \count_space UNIV) \count_space UNIV)" by (simp add: nn_integral_count_space_nat) also have "... = (\\<^sup>+(i,n). e (i, n) * indicator {0} i \count_space UNIV)" using nn_integral_snd_count_space[of "\ (i,n). e(i,n) * indicator {0} i"] by auto also have "... = (\\<^sup>+(i,n). e (i, n) * indicator V (i,n) \count_space UNIV)" by (metis i) finally have "(\n. d 0 n) = (\\<^sup>+y \ V. e y \count_space UNIV)" by (simp add: split_def) then have "?R = (\\<^sup>+y \ V. e y \count_space UNIV) + (\\<^sup>+y \ U. e y \count_space UNIV)" using * by simp also have "... = (\\<^sup>+y \ V \ U. e y \count_space UNIV)" by (rule nn_integral_disjoint_pair_countspace[symmetric], auto simp add: U_def V_def) also have "... = (\\<^sup>+(i, n). e (i, n) * indicator {..n} i \count_space UNIV)" - by (rule nn_integral_cong, auto simp add: indicator_def V_def U_def pos, meson) + by (rule nn_integral_cong, auto simp add: indicator_def of_bool_def V_def U_def pos, meson) also have "... = (\\<^sup>+n. (\\<^sup>+i. e (i, n) * indicator {..n} i \count_space UNIV)\count_space UNIV)" using nn_integral_snd_count_space[of "\(i,n). e(i,n) * indicator {..n} i"] by auto also have "... = (\n. (\i. e (i, n) * indicator {..n} i))" using pos nn_integral_count_space_nat suminf_0_le by auto moreover have "(\i. e (i, n) * indicator {..n} i) = (\i\n. e (i, n))" for n proof - have "finite {..n}" by simp moreover have "\i. i \ {..n} \ e (i,n) * indicator {..n} i = 0" using indicator_def by simp then have "(\i. e (i,n) * indicator {..n} i) = (\i \ {..n} . e (i, n) * indicator {..n} i)" by (meson calculation suminf_finite) moreover have "\i. i \ {..n} \ e (i, n) * indicator {..n} i = e (i, n)" using indicator_def by auto ultimately show "(\i. e (i, n) * indicator {..n} i) = (\i\n. e (i, n))" by simp qed ultimately show ?thesis using e_def by simp qed end context conservative_mpt begin text \We prove Kac's Formula (in the general form for induced functions) first for functions taking values in ennreal (to avoid all summabilities issues). The result for real functions will follow by domination. First, we assume additionally that $f$ is bounded and has a support of finite measure, the general case will follow readily by truncation. The proof is again an instance of the fact that the preimage of the set of points with first entrance time $n$ is the union of the set of points with first entrance time $n+1$, and the points of $A$ with first return $n+1$. Keeping track of the integral of $f$ on the different parts that appear in this argument, we will see that the integral of the induced function on the set of points with return time at most $n$ is equal to the integral of the function, up to an error controlled by the measure of points in $T^{-n}(supp(f))$ with local time $0$. Local time controls ensure that this contribution vanishes asymptotically. \ lemma induced_function_nn_integral_aux: fixes f::"'a \ ennreal" assumes A_meas [measurable]: "A \ sets M" and f_meas [measurable]: "f \ borel_measurable M" and f_bound: "\x. f x \ ennreal C" "0 \ C" and f_supp: "emeasure M {x \ space M. f x > 0} < \" shows "(\\<^sup>+y. induced_function A f y \M) = (\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M)" proof - define B where "B = (\n. first_entrance_set A n)" have B_meas [measurable]: "\n. B n \ sets M" by (simp add: B_def) then have B2 [measurable]: "(\n. B (n+1)) \ sets M" by measurable have *: "B = disjointed (\i. (T^^i)--`A)" by (auto simp add: B_def disjointed_def first_entrance_set_def) then have "disjoint_family B" by (simp add: disjoint_family_disjointed) have "(\n. (T^^n)--`A) = (\n. disjointed (\i. (T^^i)--`A) n)" by (simp add: UN_disjointed_eq) then have "(\n. (T^^n)--`A) = (\n. B n)" using * by simp then have "(\n. (T^^n)--`A) = B 0 \ (\n. B (n+1))" by (auto) (metis not0_implies_Suc) then have "(\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M) = (\\<^sup>+ x \ (B 0 \ (\n. B (n+1))). f x \M)" by simp also have "... = (\\<^sup>+ x \ B 0. f x \M) + (\\<^sup>+ x \ (\n. B (n+1)). f x \M)" proof (rule nn_integral_disjoint_pair) show "B 0 \ (\n. B (n+1)) = {}" by (auto) (metis IntI Suc_neq_Zero UNIV_I empty_iff \disjoint_family B\ disjoint_family_on_def) qed auto finally have "(\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M) = (\\<^sup>+ x \ B 0. f x \M) + (\\<^sup>+ x \ (\n. B (n+1)). f x \M)" by simp moreover have "(\\<^sup>+ x \ (\n. B (n+1)). f x \M) = (\n. (\\<^sup>+ x \ B (n+1). f x \M))" apply (rule nn_integral_disjoint_family) using \disjoint_family B\ by (auto simp add: disjoint_family_on_def) ultimately have Bdec: "(\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M) = (\\<^sup>+ x \ B 0. f x \M) + (\n. \\<^sup>+ x \ B (n+1). f x \M)" by simp define D where "D = (\n. (return_time_function A)-` {n+1})" then have "disjoint_family D" by (auto simp add: disjoint_family_on_def) have *: "\n. D n = T--`(B n) \ A" using D_def B_def return_time_rec[OF assms(1)] by simp then have [measurable]: "\n. D n \ sets M" by simp have **: "\n. B (n+1) = T--`(B n) - A" using first_entrance_rec[OF assms(1)] B_def by simp have pos0: "\i x. f((T^^i)x) \ 0" using assms(3) by simp have pos:"\ i C x. f((T^^i)x) * indicator (C) x \ 0" using assms(3) by simp have mes0 [measurable]: "\ i. (\x. f((T^^i)x)) \ borel_measurable M" by simp then have [measurable]: "\ i C. C \ sets M \ (\x. f((T^^i)x) * indicator C x) \ borel_measurable M" by simp have "\y. induced_function A f y = induced_function A f y * indicator ((return_time_function A)-`{1..}) y" by (rule induced_function_support) moreover have "(return_time_function A)-`{(1::nat)..} = (\n. D n)" by (auto simp add: D_def Suc_le_D) ultimately have "\y. induced_function A f y = induced_function A f y * indicator (\n. D n) y" by simp then have "(\\<^sup>+y. induced_function A f y \M) = (\\<^sup>+y \ (\n. D n). induced_function A f y \M)" by simp also have "... = (\n. (\\<^sup>+y \ D n. induced_function A f y \M))" apply (rule nn_integral_disjoint_family) unfolding induced_function_def by (auto simp add: pos0 sum_nonneg \disjoint_family D\) finally have a: "(\\<^sup>+y. induced_function A f y \M) = (\n. (\\<^sup>+y \ D n. induced_function A f y \M))" by simp define d where "d = (\i n. (\\<^sup>+y \ D n. f((T^^i)y) \M))" have "(\\<^sup>+y \ D n. induced_function A f y \M) = (\i\{..n}. d i n)" for n proof - have "induced_function A f y * indicator (D n) y = (\i\{..\<^sup>+y \ D n. induced_function A f y \M) = (\i\{..\<^sup>+y \ D n. f((T^^i)y) \M))" using pos nn_integral_sum[of "{..i y. f((T^^i)y) * indicator (D n) y"] by simp also have "... = (\i\{..n}. (\\<^sup>+y \ D n. f((T^^i)y) \M))" using lessThan_Suc_atMost by auto finally show ?thesis using d_def by simp qed then have induced_dec: "(\\<^sup>+y. induced_function A f y \M) = (\n. (\i\{..n}. d i n))" using a by simp have "(\n\{1..}. (return_time_function A)-` {n}) = UNIV - (return_time_function A)-` {0}" by auto then have "(\n\{1..}. (return_time_function A)-` {n}) = recurrent_subset A" using return_time0 by auto moreover have "(\n. (return_time_function A)-` {n+1}) = (\n\{1..}. (return_time_function A)-` {n})" by (auto simp add: Suc_le_D) ultimately have "(\n. D n) = recurrent_subset A" using D_def by simp moreover have "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ recurrent_subset A. f x \M)" by (rule nn_integral_null_delta, auto simp add: Diff_mono Un_absorb2 recurrent_subset_incl(1)[of A] Poincare_recurrence_thm(1)[OF assms(1)]) moreover have "B 0 = A" using B_def first_entrance_set_def by simp ultimately have "(\\<^sup>+x \ B 0. f x \M) = (\\<^sup>+x \ (\n. D n). f x \M)" by simp also have "... = (\n. (\\<^sup>+x \ D n. f x \M))" by (rule nn_integral_disjoint_family, auto simp add: \disjoint_family D\) finally have B0dec: "(\\<^sup>+x \ B 0. f x \M) = (\n. d 0 n)" using d_def by simp have *: "(\\<^sup>+x \ B n. f x \M) = (\i\<^sup>+x \ D(n+i). f((T^^(i+1))x) \M)) + (\\<^sup>+x \ B(n+k). f((T^^k)x) \M)" for n k proof (induction k) case 0 show ?case by simp next case (Suc k) have "T--`(B(n+k)) = B(n+k+1) \ D(n+k)" using * ** by blast have "(\\<^sup>+x \ B(n+k). f((T^^k)x) \M) = (\\<^sup>+x. (\x. f((T^^k)x) * indicator (B (n+k)) x)(T x) \M)" by (rule measure_preserving_preserves_nn_integral[OF Tm], auto simp add: pos) also have "... = (\\<^sup>+x. f((T^^(k+1))x) * indicator (T--`(B (n+k))) x \M)" proof (rule nn_integral_cong_AE) have "(T^^k)(T x) = (T^^(k+1))x" for x using comp_eq_dest_lhs by (metis Suc_eq_plus1 funpow.simps(2) funpow_swap1) moreover have "AE x in M. f((T^^k)(T x)) * indicator (B (n+k)) (T x) = f((T^^k)(T x)) * indicator (T--`(B (n+k))) x" by (simp add: indicator_def \\n. B n \ sets M\) ultimately show "AE x in M. f((T^^k)(T x)) * indicator (B (n+k)) (T x) = f((T^^(k+1))x) * indicator (T--`(B (n+k))) x" by simp qed also have "... = (\\<^sup>+x \ B(n+k+1) \ D(n+k). f((T^^(k+1))x) \M)" using \T--`(B(n+k)) = B(n+k+1) \ D(n+k)\ by simp also have "... = (\\<^sup>+x \ B(n+k+1). f((T^^(k+1))x) \M) + (\\<^sup>+x \ D(n+k). f((T^^(k+1))x) \M)" proof (rule nn_integral_disjoint_pair[OF mes0[of "k+1"]]) show "B(n+k+1) \ D(n+k) = {}" using * ** by blast qed (auto) finally have "(\\<^sup>+x \ B(n+k). f((T^^k)x) \M) = (\\<^sup>+x \ B(n+k+1). f((T^^(k+1))x) \M) + (\\<^sup>+x \ D(n+k). f((T^^(k+1))x) \M)" by simp then show ?case by (simp add: Suc.IH add.commute add.left_commute) qed have a: "(\k. (\\<^sup>+x \ B(n+k). f((T^^k)x) \M)) \ 0" for n proof - define W where "W = {x \ space M. f x > 0} \ (T^^n)--`A" have "emeasure M W \ emeasure M {x \ space M. f x > 0}" by (intro emeasure_mono, auto simp add: W_def) then have W_fin: "emeasure M W < \" using f_supp by auto have W_meas [measurable]: "W \ sets M" unfolding W_def by simp have W_incl: "W \ (T^^n)--`A" unfolding W_def by simp define V where "V = (\k. {x \ (T^^k)--`W. local_time A k x = 0})" have V_meas [measurable]: "V k \ sets M" for k unfolding V_def by simp have a: "(\\<^sup>+x \ B(n+k). f((T^^k)x) \M) \ C * emeasure M (V k)" for k proof - have "f((T^^k)x) * indicator (B(n+k)) x = f((T^^k)x) * indicator (B(n+k) \ (T^^k)--`W) x" for x proof (cases) assume "f((T^^k)x) * indicator (B(n+k)) x = 0" then show ?thesis by (simp add: indicator_def) next assume "\(f((T^^k)x) * indicator (B(n+k)) x = 0)" then have H: "f((T^^k)x) * indicator (B(n+k)) x \ 0" by simp then have inB: "x \ B(n+k)" using H using indicator_simps(2) by fastforce then have s: "x \ space M" using B_meas[of "n+k"] sets.sets_into_space by blast then have a: "(T^^k)x \ space M" by (metis measurable_space Tn_meas[of k]) have "f((T^^k)x) > 0" using H by (simp add: le_neq_trans) then have *: "(T^^k)x \ {y \ space M. f y > 0}" using a by simp have "(T^^(n+k))x \ A" using inB B_def first_entrance_set_def by auto then have "(T^^n)((T^^k)x) \ A" by (simp add: funpow_add) then have "(T^^k)x \ (T^^n)--`A" using a by auto then have "(T^^k)x \ W" using * W_def by simp then have "x \ (T^^k)--`W" using s a by simp then have "x \ (B(n+k) \ (T^^k)--`W)" using inB by simp then show ?thesis by auto qed then have *: "(\\<^sup>+x \ B(n+k). f((T^^k)x) \M) = (\\<^sup>+x \ B(n+k) \ (T^^k)--`W. f((T^^k)x) \M)" by simp have "B(n+k) \ {x \ space M. local_time A k x = 0}" unfolding local_time_def B_def first_entrance_set_def by auto then have "B(n+k) \ (T^^k)--`W \ V k" unfolding V_def by blast then have "f((T^^k)x) * indicator (B(n+k) \ (T^^k)--`W) x \ ennreal C * indicator (V k) x" for x using f_bound by (auto split: split_indicator) then have "(\\<^sup>+x \ B(n+k) \ (T^^k)--`W. f((T^^k)x) \M) \ (\\<^sup>+x. ennreal C * indicator (V k) x \M)" by (simp add: nn_integral_mono) also have "... = ennreal C * emeasure M (V k)" by (simp add: \0 \ C\ nn_integral_cmult) finally show "(\\<^sup>+x \ B(n+k). f((T^^k)x) \M) \ C * emeasure M (V k)" using * by simp qed have "(\k. emeasure M (V k)) \ 0" unfolding V_def using local_time_unbounded2[OF W_meas, OF W_fin, OF W_incl, of 1] by auto from ennreal_tendsto_cmult[OF _ this, of C] have t0: "(\k. C * emeasure M (V k)) \ 0" by simp from a show "(\k. (\\<^sup>+x \ B(n+k). f((T^^k)x) \M)) \ 0" by (intro tendsto_sandwich[OF _ _ tendsto_const t0]) auto qed have b: "(\k. (\i\<^sup>+x \ D(n+i). f((T^^(i+1))x) \M))) \ (\i. d (i+1) (n+i))" for n proof - define e where "e = (\i. d (i+1) (n+i))" then have "(\k. (\i (\i. e i)" by (intro summable_LIMSEQ) simp then show "(\k. (\i\<^sup>+x \ D(n+i). f((T^^(i+1))x) \M))) \ (\i. d (i+1) (n+i))" using e_def d_def by simp qed have "(\k. (\i\<^sup>+x \ D(n+i). f((T^^(i+1))x) \M)) + (\\<^sup>+x \ B(n+k). f((T^^k)x) \M)) \ (\i. d (i+1) (n+i))" for n using tendsto_add[OF b a] by simp moreover have "(\k. (\i\<^sup>+x \ D(n+i). f((T^^(i+1))x) \M)) + (\\<^sup>+x \ B(n+k). f((T^^k)x) \M)) \ (\\<^sup>+x \ B n. f x \M)" for n using * by simp ultimately have "(\\<^sup>+x \ B n. f x \M) = (\i. d (i+1) (n+i))" for n using LIMSEQ_unique by blast then have "(\n. (\\<^sup>+x \ B (n+1). f x \M)) = (\n. (\i. d (i+1) (n+1+i)))" by simp then have "(\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M) = (\n. d 0 n) + (\n. (\i. d (i+1) (n+1+i)))" using Bdec B0dec by simp then show ?thesis using induced_dec kac_series_aux by simp qed text \We remove the boundedness assumption on $f$ and the finiteness assumption on its support by truncation (both in space and on the values of $f$).\ theorem induced_function_nn_integral: fixes f::"'a \ ennreal" assumes A_meas [measurable]: "A \ sets M" and f_meas [measurable]: "f \ borel_measurable M" shows "(\\<^sup>+y. induced_function A f y \M) = (\\<^sup>+x \ (\n. (T^^n)--`A). f x \M)" proof - obtain Y::"nat \ 'a set" where Y_meas: "\ n. Y n \ sets M" and Y_fin: "\n. emeasure M (Y n) \ \" and Y_full: "(\n. Y n) = space M" and Y_inc: "incseq Y" by (meson range_subsetD sigma_finite_incseq) define F where "F = (\(n::nat) x. min (f x) n * indicator (Y n) x)" have mes [measurable]: "\n. (F n) \ borel_measurable M" unfolding F_def using assms(2) Y_meas by measurable then have mes_rel [measurable]: "(\x. F n x * indicator (\n. (T^^n)--`A) x) \ borel_measurable M" for n by measurable have bound: "\n x. F n x \ ennreal n" by (simp add: F_def indicator_def ennreal_of_nat_eq_real_of_nat) have "\n. {x \ space M. F n x > 0} \ Y n" unfolding F_def using not_le by fastforce then have le: "emeasure M {x \ space M. F n x > 0} \ emeasure M (Y n)" for n by (metis emeasure_mono Y_meas) have fin: "emeasure M {x \ space M. F n x > 0} < \" for n using Y_fin[of n] le[of n] by (simp add: less_top) have *: "(\\<^sup>+y. induced_function A (F n) y \M) = (\\<^sup>+ x \ (\n. (T^^n)--`A). (F n) x \M)" for n by (rule induced_function_nn_integral_aux[OF A_meas mes bound _ fin]) simp have inc_Fx: "\x. incseq (\n. F n x)" unfolding F_def incseq_def proof (auto simp add: incseq_def) fix x::"'a" and m n::"nat" assume "m \ n" then have "min (f x) m \ min (f x) n" using linear by fastforce moreover have "(indicator (Y m) x::ennreal) \ (indicator (Y n) x::ennreal)" using Y_inc apply (auto simp add: incseq_def) using \m \ n\ by blast ultimately show "min (f x) m * (indicator (Y m) x::ennreal) \ min (f x) n * (indicator (Y n) x::ennreal)" by (auto split: split_indicator) qed then have "incseq (\n. F n x * indicator (\n. (T^^n)--`A) x)" for x by (auto simp add: indicator_def incseq_def) then have inc_rel: "incseq (\n x. F n x * indicator (\n. (T^^n)--`A) x)" by (auto simp add: incseq_def le_fun_def) then have a: "(SUP n. (\\<^sup>+ x \ (\n. (T^^n)--`A). F n x \M)) = (\\<^sup>+ x. (SUP n. F n x * indicator (\n. (T^^n)--`A) x) \M)" using nn_integral_monotone_convergence_SUP[OF inc_rel, OF mes_rel] by simp have SUP_Fx: "(SUP n. F n x) = f x" if "x \ space M" for x proof - obtain N where "x \ Y N" using Y_full \x \ space M\ by auto have "(SUP n. F n x) = (SUP n. inf (f x) (of_nat n))" proof (rule SUP_eq) show "\j\UNIV. F i x \ inf (f x) (of_nat j)" for i by (auto simp: F_def intro!: exI[of _ i] split: split_indicator) show "\i\UNIV. inf (f x) (of_nat j) \ F i x" for j using \x \ Y N\ \incseq Y\[THEN monoD, of N "max N j"] by (intro bexI[of _ "max N j"]) (auto simp: F_def subset_eq not_le inf_min intro: min.coboundedI2 less_imp_le split: split_indicator split_max) qed then show ?thesis by (simp add: inf_SUP[symmetric] ennreal_SUP_of_nat_eq_top) qed then have "(SUP n. F n x * indicator (\n. (T^^n)--`A) x) = f x * indicator (\n. (T^^n)--`A) x" if "x \ space M" for x by (auto simp add: indicator_def SUP_Fx that) then have **: "(SUP n. (\\<^sup>+ x \ (\n. (T^^n)--`A). F n x \M)) = (\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M)" by (simp add: a cong: nn_integral_cong) have "incseq (\n. induced_function A (F n) x)" for x unfolding induced_function_def using incseq_sumI2[of "{..i n. F n ((T^^i)x)"] inc_Fx by simp then have "incseq (\n. induced_function A (F n))" by (auto simp add: incseq_def le_fun_def) then have b: "(SUP n. (\\<^sup>+ x. induced_function A (F n) x \M)) = (\\<^sup>+ x. (SUP n. induced_function A (F n) x) \M)" by (rule nn_integral_monotone_convergence_SUP[symmetric]) (measurable) have "(SUP n. induced_function A (F n) x) = induced_function A f x" if [simp]: "x \ space M" for x proof - have "(SUP n. (\ i \{.. i \ {..\<^sup>+ x. induced_function A (F n) x \M)) = (\\<^sup>+ x. induced_function A f x \M)" by (simp add: b cong: nn_integral_cong) then show ?thesis using * ** by simp qed text \Taking the constant function equal to $1$ in the previous statement, we obtain the usual Kac Formula.\ theorem kac_formula_nonergodic: assumes A_meas [measurable]: "A \ sets M" shows "(\\<^sup>+y. return_time_function A y \M) = emeasure M (\n. (T^^n)--`A)" proof - define f where "f = (\(x::'a). 1::ennreal)" have "\x. induced_function A f x = return_time_function A x" unfolding induced_function_def f_def by (simp add:) then have "(\\<^sup>+y. return_time_function A y \M) = (\\<^sup>+y. induced_function A f y \M)" by auto also have "... = (\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M)" by (rule induced_function_nn_integral) (auto simp add: f_def) also have "... = emeasure M (\n. (T^^n)--`A)" using f_def by auto finally show ?thesis by simp qed lemma (in fmpt) return_time_integrable: assumes A_meas [measurable]: "A \ sets M" shows "integrable M (return_time_function A)" by (rule integrableI_nonneg) (auto simp add: kac_formula_nonergodic[OF assms] ennreal_of_nat_eq_real_of_nat[symmetric] less_top[symmetric]) text \Now, we want to prove the same result but for real-valued integrable function. We first prove the statement for nonnegative functions by reducing to the nonnegative extended reals, and then for general functions by difference.\ lemma induced_function_integral_aux: fixes f::"'a \ real" assumes A_meas [measurable]: "A \ sets M" and f_int [measurable]: "integrable M f" and f_pos: "\x. f x \ 0" shows "integrable M (induced_function A f)" "(\y. induced_function A f y \M) = (\x \ (\n. (T^^n)--`A). f x \M)" proof - show "integrable M (induced_function A f)" proof (rule integrableI_nonneg) show "AE x in M. induced_function A f x \ 0" unfolding induced_function_def by (simp add: f_pos sum_nonneg) have "(\\<^sup>+x. ennreal (induced_function A f x) \M) = (\\<^sup>+ x. induced_function A (\x. ennreal(f x)) x \M)" unfolding induced_function_def by (auto simp: f_pos) also have "... = (\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M)" by (rule induced_function_nn_integral, auto simp add: assms) also have "... \ (\\<^sup>+ x. f x \M)" using nn_set_integral_set_mono[where ?A = "(\n. (T^^n)--`A)" and ?B = UNIV and ?f = "\x. ennreal(f x)"] by auto also have "... < \" using assms by (auto simp: less_top) finally show "(\\<^sup>+ x. induced_function A f x \M) < \" by simp qed (simp) have "(\\<^sup>+ x. (f x * indicator (\n. (T^^n)--`A) x) \M) = (\\<^sup>+ x \ (\n. (T^^n)--`A). f x \M)" by (auto split: split_indicator intro!: nn_integral_cong) also have "... = (\\<^sup>+ x. induced_function A (\x. ennreal(f x)) x \M)" by (rule induced_function_nn_integral[symmetric], auto simp add: assms) also have "... = (\\<^sup>+x. ennreal (induced_function A f x) \M)" unfolding induced_function_def by (auto simp: f_pos) finally have *: "(\\<^sup>+ x. (f x * indicator (\n. (T^^n)--`A) x) \M) = (\\<^sup>+x. ennreal (induced_function A f x) \M)" by simp have "(\ x \ (\n. (T^^n)--`A). f x \M) = (\ x. f x * indicator (\n. (T^^n)--`A) x \M)" by (simp add: mult.commute set_lebesgue_integral_def) also have "... = enn2real (\\<^sup>+ x. (f x * indicator (\n. (T^^n)--`A) x) \M)" by (rule integral_eq_nn_integral, auto simp add: f_pos) also have "... = enn2real (\\<^sup>+x. ennreal (induced_function A f x) \M)" using * by simp also have "... = (\ x. induced_function A f x \M)" apply (rule integral_eq_nn_integral[symmetric]) unfolding induced_function_def by (auto simp add: f_pos sum_nonneg) finally show "(\ x. induced_function A f x \M) = (\ x \ (\n. (T^^n)--`A). f x \M)" by simp qed text \Here is the general version of Kac's Formula (for a general induced function, starting from a real-valued integrable function).\ theorem induced_function_integral_nonergodic: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" shows "integrable M (induced_function A f)" "(\y. induced_function A f y \M) = (\x \ (\n. (T^^n)--`A). f x \M)" proof - have U_meas [measurable]: "(\n. (T^^n)--`A) \ sets M" by measurable define g where "g = (\x. max (f x) 0)" have g_int [measurable]: "integrable M g" unfolding g_def using assms by auto then have g_int2: "integrable M (induced_function A g)" using induced_function_integral_aux(1) g_def by auto define h where "h = (\x. max (-f x) 0)" have h_int [measurable]: "integrable M h" unfolding h_def using assms by auto then have h_int2: "integrable M (induced_function A h)" using induced_function_integral_aux(1) h_def by auto have D1: "f = (\x. g x - h x)" unfolding g_def h_def by auto have D2: "induced_function A f = (\x. induced_function A g x - induced_function A h x)" unfolding induced_function_def using D1 by (simp add: sum_subtractf) then show "integrable M (induced_function A f)" using g_int2 h_int2 by auto have "(\x. induced_function A f x \M) = (\x. induced_function A g x - induced_function A h x \M)" using D2 by simp also have "... = (\x. induced_function A g x \M) - (\x. induced_function A h x \M)" using g_int2 h_int2 by auto also have "... = (\x \ (\n. (T^^n)--`A). g x \M) - (\x \ (\n. (T^^n)--`A). h x \M)" using induced_function_integral_aux(2) g_def h_def g_int h_int by auto also have "... = (\x \ (\n. (T^^n)--`A). (g x - h x) \M)" apply (rule set_integral_diff(2)[symmetric]) unfolding set_integrable_def using g_int h_int integrable_mult_indicator[OF U_meas] by blast+ also have "... = (\x \ (\n. (T^^n)--`A). f x \M)" using D1 by simp finally show "(\x. induced_function A f x \M) = (\x \ (\n. (T^^n)--`A). f x \M)" by simp qed text \We can reformulate the previous statement in terms of induced measure.\ lemma induced_function_integral_restr_nonergodic: fixes f::"'a \ real" assumes [measurable]: "A \ sets M" "integrable M f" shows "integrable (restrict_space M A) (induced_function A f)" "(\y. induced_function A f y \(restrict_space M A)) = (\ x \ (\n. (T^^n)--`A). f x \M)" proof - have [measurable]: "integrable M (induced_function A f)" by (rule induced_function_integral_nonergodic(1)[OF assms]) then show "integrable (restrict_space M A) (induced_function A f)" by (metis assms(1) integrable_mult_indicator integrable_restrict_space sets.Int_space_eq2) have "(\y. induced_function A f y \(restrict_space M A)) = (\y \ A. induced_function A f y \M)" by (simp add: integral_restrict_space set_lebesgue_integral_def) also have "... = (\y. induced_function A f y \M)" unfolding real_scaleR_def set_lebesgue_integral_def proof (rule Bochner_Integration.integral_cong [OF refl]) have "induced_function A f y = 0" if "y \ A" for y unfolding induced_function_def using that return_time0[of A] recurrent_subset_incl(1)[of A] return_time_function_def by auto then show "\x. indicator A x * induced_function A f x = induced_function A f x" unfolding indicator_def by auto qed also have "... = (\ x \ (\n. (T^^n)--`A). f x \M)" by (rule induced_function_integral_nonergodic(2)[OF assms]) finally show "(\y. induced_function A f y \(restrict_space M A)) = (\ x \ (\n. (T^^n)--`A). f x \M)" by simp qed end end (*of Recurrence.thy*) diff --git a/thys/Ergodic_Theory/SG_Library_Complement.thy b/thys/Ergodic_Theory/SG_Library_Complement.thy --- a/thys/Ergodic_Theory/SG_Library_Complement.thy +++ b/thys/Ergodic_Theory/SG_Library_Complement.thy @@ -1,1319 +1,1322 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \SG Libary complements\ theory SG_Library_Complement imports "HOL-Probability.Probability" begin text \In this file are included many statements that were useful to me, but belong rather naturally to existing theories. In a perfect world, some of these statements would get included into these files. I tried to indicate to which of these classical theories the statements could be added. \ subsection \Basic logic\ text \This one is certainly available, but I could not locate it...\ lemma equiv_neg: "\ P \ Q; \P \ \Q \ \ (P\Q)" by blast subsection \Basic set theory\ lemma compl_compl_eq_id [simp]: "UNIV - (UNIV - s) = s" by auto abbreviation sym_diff :: "'a set \ 'a set \ 'a set" (infixl "\" 70) where "sym_diff A B \ ((A - B) \ (B-A))" text \Not sure the next lemmas are useful, as they are proved solely by auto, so they could be reproved automatically whenever necessary.\ lemma sym_diff_inc: "A \ C \ A \ B \ B \ C" by auto lemma sym_diff_vimage [simp]: "f-`(A \ B) = (f-`A) \ (f-`B)" by auto subsection \Set-Interval.thy\ text \The next two lemmas belong naturally to \verb+Set_Interval.thy+, next to \verb+UN_le_add_shift+. They are not trivially equivalent to the corresponding lemmas with large inequalities, due to the difference when $n = 0$.\ lemma UN_le_eq_Un0_strict: "(\ii\{1.. M 0" (is "?A = ?B") proof show "?A \ ?B" proof fix x assume "x \ ?A" then obtain i where i: "i M i" by auto show "x \ ?B" proof(cases i) case 0 with i show ?thesis by simp next case (Suc j) with i show ?thesis by auto qed qed qed (auto) text \I use repeatedly this one, but I could not find it directly\ lemma union_insert_0: "(\n::nat. A n) = A 0 \ (\n\{1..}. A n)" by (metis UN_insert Un_insert_left sup_bot.left_neutral One_nat_def atLeast_0 atLeast_Suc_greaterThan ivl_disj_un_singleton(1)) text \Next one could be close to \verb+sum.nat_group+\ lemma sum_arith_progression: "(\r<(N::nat). (\ijr j \ {i*N..riiri j \ {i*N..jMiscellanous basic results\ lemma ind_from_1 [case_names 1 Suc, consumes 1]: assumes "n > 0" assumes "P 1" and "\n. n > 0 \ P n \ P (Suc n)" shows "P n" proof - have "(n = 0) \ P n" proof (induction n) case 0 then show ?case by auto next case (Suc k) consider "Suc k = 1" | "Suc k > 1" by linarith then show ?case apply (cases) using assms Suc.IH by auto qed then show ?thesis using \n > 0\ by auto qed text \This lemma is certainly available somewhere, but I couldn't locate it\ lemma tends_to_real_e: fixes u::"nat \ real" assumes "u \ l" "e>0" shows "\N. \n>N. abs(u n -l) < e" by (metis assms dist_real_def le_less lim_sequentially) lemma nat_mod_cong: assumes "a = b+(c::nat)" "a mod n = b mod n" shows "c mod n = 0" proof - let ?k = "a mod n" obtain a1 where "a = a1*n + ?k" by (metis div_mult_mod_eq) moreover obtain b1 where "b = b1*n + ?k" using assms(2) by (metis div_mult_mod_eq) ultimately have "a1 * n + ?k = b1 * n + ?k + c" using assms(1) by arith then have "c = (a1 - b1) * n" by (simp add: diff_mult_distrib) then show ?thesis by simp qed lemma funpow_add': "(f ^^ (m + n)) x = (f ^^ m) ((f ^^ n) x)" by (simp add: funpow_add) text \The next two lemmas are not directly equivalent, since $f$ might not be injective.\ lemma abs_Max_sum: fixes A::"real set" assumes "finite A" "A \ {}" shows "abs(Max A) \ (\a\A. abs(a))" by (simp add: assms member_le_sum) lemma abs_Max_sum2: fixes f::"_ \ real" assumes "finite A" "A \ {}" shows "abs(Max (f`A)) \ (\a\A. abs(f a))" using assms by (induct rule: finite_ne_induct, auto) subsection \Conditionally-Complete-Lattices.thy\ lemma mono_cInf: fixes f :: "'a::conditionally_complete_lattice \ 'b::conditionally_complete_lattice" assumes "mono f" "A \ {}" "bdd_below A" shows "f(Inf A) \ Inf (f`A)" using assms by (simp add: cINF_greatest cInf_lower monoD) lemma mono_bij_cInf: fixes f :: "'a::conditionally_complete_linorder \ 'b::conditionally_complete_linorder" assumes "mono f" "bij f" "A \ {}" "bdd_below A" shows "f (Inf A) = Inf (f`A)" proof - have "(inv f) (Inf (f`A)) \ Inf ((inv f)`(f`A))" apply (rule cInf_greatest, auto simp add: assms(3)) using mono_inv[OF assms(1) assms(2)] assms by (simp add: mono_def bdd_below_image_mono cInf_lower) then have "Inf (f`A) \ f (Inf ((inv f)`(f`A)))" by (metis (no_types, lifting) assms(1) assms(2) mono_def bij_inv_eq_iff) also have "... = f(Inf A)" using assms by (simp add: bij_is_inj) finally show ?thesis using mono_cInf[OF assms(1) assms(3) assms(4)] by auto qed subsection \Topological-spaces.thy\ lemma open_less_abs [simp]: "open {x. (C::real) < abs x}" proof - have *: "{x. C < abs x} = abs-`{C<..}" by auto show ?thesis unfolding * by (auto intro!: continuous_intros) qed lemma closed_le_abs [simp]: "closed {x. (C::real) \ abs x}" proof - have *: "{x. C \ \x\} = abs-`{C..}" by auto show ?thesis unfolding * by (auto intro!: continuous_intros) qed text \The next statements come from the same statements for true subsequences\ lemma eventually_weak_subseq: fixes u::"nat \ nat" assumes "(\n. real(u n)) \ \" "eventually P sequentially" shows "eventually (\n. P (u n)) sequentially" proof - obtain N where *: "\n\N. P n" using assms(2) unfolding eventually_sequentially by auto obtain M where "\m\M. ereal(u m) \ N" using assms(1) by (meson Lim_PInfty) then have "\m. m \ M \ u m \ N" by auto then have "\m. m \ M \ P(u m)" using \\n\N. P n\ by simp then show ?thesis unfolding eventually_sequentially by auto qed lemma filterlim_weak_subseq: fixes u::"nat \ nat" assumes "(\n. real(u n)) \ \" shows "LIM n sequentially. u n:> at_top" unfolding filterlim_iff by (metis assms eventually_weak_subseq) lemma limit_along_weak_subseq: fixes u::"nat \ nat" and v::"nat \ _" assumes "(\n. real(u n)) \ \" "v \ l" shows "(\ n. v(u n)) \ l" using filterlim_compose[of v, OF _ filterlim_weak_subseq] assms by auto lemma frontier_indist_le: assumes "x \ frontier {y. infdist y S \ r}" shows "infdist x S = r" proof - have "infdist x S = r" if H: "\e>0. (\y. infdist y S \ r \ dist x y < e) \ (\z. \ infdist z S \ r \ dist x z < e)" proof - have "infdist x S < r + e" if "e > 0" for e proof - obtain y where "infdist y S \ r" "dist x y < e" using H \e > 0\ by blast then show ?thesis by (metis add.commute add_mono_thms_linordered_field(3) infdist_triangle le_less_trans) qed then have A: "infdist x S \ r" by (meson field_le_epsilon order.order_iff_strict) have "r < infdist x S + e" if "e > 0" for e proof - obtain y where "\(infdist y S \ r)" "dist x y < e" using H \e > 0\ by blast then have "r < infdist y S" by auto also have "... \ infdist x S + dist y x" by (rule infdist_triangle) finally show ?thesis using \dist x y < e\ by (simp add: dist_commute) qed then have B: "r \ infdist x S" by (meson field_le_epsilon order.order_iff_strict) show ?thesis using A B by auto qed then show ?thesis using assms unfolding frontier_straddle by auto qed subsection \Limits\ text \The next lemmas are not very natural, but I needed them several times\ lemma tendsto_shift_1_over_n [tendsto_intros]: fixes f::"nat \ real" assumes "(\n. f n / n) \ l" shows "(\n. f (n+k) / n) \ l" proof - have "(1+k*(1/n))* (f(n+k)/(n+k)) = f(n+k)/n" if "n>0" for n using that by (auto simp add: divide_simps) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n.(1+k*(1/n))* (f(n+k)/(n+k)) = f(n+k)/n) sequentially" by auto moreover have "(\n. (1+k*(1/n))* (f(n+k)/(n+k))) \ (1+real k*0) * l" by (intro tendsto_intros LIMSEQ_ignore_initial_segment assms) ultimately show ?thesis using Lim_transform_eventually by auto qed lemma tendsto_shift_1_over_n' [tendsto_intros]: fixes f::"nat \ real" assumes "(\n. f n / n) \ l" shows "(\n. f (n-k) / n) \ l" proof - have "(1-k*(1/(n+k)))* (f n/ n) = f n/(n+k)" if "n>0" for n using that by (auto simp add: divide_simps) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n. (1-k*(1/(n+k)))* (f n/ n) = f n/(n+k)) sequentially" by auto moreover have "(\n. (1-k*(1/(n+k)))* (f n/ n)) \ (1-real k*0) * l" by (intro tendsto_intros assms LIMSEQ_ignore_initial_segment) ultimately have "(\n. f n / (n+k)) \ l" using Lim_transform_eventually by auto then have a: "(\n. f(n-k)/(n-k+k)) \ l" using seq_offset_neg by auto have "f(n-k)/(n-k+k) = f(n-k)/n" if "n>k" for n using that by auto with eventually_mono[OF eventually_gt_at_top[of k] this] have "eventually (\n. f(n-k)/(n-k+k) = f(n-k)/n) sequentially" by auto with Lim_transform_eventually[OF a this] show ?thesis by auto qed declare LIMSEQ_realpow_zero [tendsto_intros] subsection \Topology-Euclidean-Space\ text \A (more usable) variation around \verb+continuous_on_closure_sequentially+. The assumption that the spaces are metric spaces is definitely too strong, but sufficient for most applications.\ lemma continuous_on_closure_sequentially': fixes f::"'a::metric_space \ 'b::metric_space" assumes "continuous_on (closure C) f" "\(n::nat). u n \ C" "u \ l" shows "(\n. f (u n)) \ f l" proof - have "l \ closure C" unfolding closure_sequential using assms by auto then show ?thesis using \continuous_on (closure C) f\ unfolding comp_def continuous_on_closure_sequentially using assms by auto qed subsection \Convexity\ lemma convex_on_mean_ineq: fixes f::"real \ real" assumes "convex_on A f" "x \ A" "y \ A" shows "f ((x+y)/2) \ (f x + f y) / 2" using convex_onD[OF assms(1), of "1/2" x y] using assms by (auto simp add: divide_simps) lemma convex_on_closure: assumes "convex (C::'a::real_normed_vector set)" "convex_on C f" "continuous_on (closure C) f" shows "convex_on (closure C) f" proof (rule convex_onI) fix x y::'a and t::real assume "x \ closure C" "y \ closure C" "0 < t" "t < 1" obtain u v::"nat \ 'a" where *: "\n. u n \ C" "u \ x" "\n. v n \ C" "v \ y" using \x \ closure C\ \y \ closure C\ unfolding closure_sequential by blast define w where "w = (\n. (1-t) *\<^sub>R (u n) + t *\<^sub>R (v n))" have "w n \ C" for n using \0 < t\ \t< 1\ convexD[OF \convex C\ *(1)[of n] *(3)[of n]] unfolding w_def by auto have "w \ ((1-t) *\<^sub>R x + t *\<^sub>R y)" unfolding w_def using *(2) *(4) by (intro tendsto_intros) have *: "f(w n) \ (1-t) * f(u n) + t * f (v n)" for n using *(1) *(3) \convex_on C f\ \0 \t<1\ less_imp_le unfolding w_def convex_on_alt by (simp add: add.commute) have i: "(\n. f (w n)) \ f ((1-t) *\<^sub>R x + t *\<^sub>R y)" by (rule continuous_on_closure_sequentially'[OF assms(3) \\n. w n \ C\ \w \ ((1-t) *\<^sub>R x + t *\<^sub>R y)\]) have ii: "(\n. (1-t) * f(u n) + t * f (v n)) \ (1-t) * f x + t * f y" apply (intro tendsto_intros) apply (rule continuous_on_closure_sequentially'[OF assms(3) \\n. u n \ C\ \u \ x\]) apply (rule continuous_on_closure_sequentially'[OF assms(3) \\n. v n \ C\ \v \ y\]) done show "f ((1 - t) *\<^sub>R x + t *\<^sub>R y) \ (1 - t) * f x + t * f y" apply (rule LIMSEQ_le[OF i ii]) using * by auto qed lemma convex_on_norm [simp]: "convex_on UNIV (\(x::'a::real_normed_vector). norm x)" using convex_on_dist[of UNIV "0::'a"] by auto lemma continuous_abs_powr [continuous_intros]: assumes "p > 0" shows "continuous_on UNIV (\(x::real). \x\ powr p)" apply (rule continuous_on_powr') using assms by (auto intro: continuous_intros) lemma continuous_mult_sgn [continuous_intros]: fixes f::"real \ real" assumes "continuous_on UNIV f" "f 0 = 0" shows "continuous_on UNIV (\x. sgn x * f x)" proof - have *: "continuous_on {0..} (\x. sgn x * f x)" apply (subst continuous_on_cong[of "{0..}" "{0..}" _ f], auto simp add: sgn_real_def assms(2)) by (rule continuous_on_subset[OF assms(1)], auto) have **: "continuous_on {..0} (\x. sgn x * f x)" apply (subst continuous_on_cong[of "{..0}" "{..0}" _ "\x. -f x"], auto simp add: sgn_real_def assms(2)) by (rule continuous_on_subset[of UNIV], auto simp add: assms intro!: continuous_intros) show ?thesis using continuous_on_closed_Un[OF _ _ * **] apply (auto intro: continuous_intros) using continuous_on_subset by fastforce qed lemma DERIV_abs_powr [derivative_intros]: assumes "p > (1::real)" shows "DERIV (\x. \x\ powr p) x :> p * sgn x * \x\ powr (p - 1)" proof - consider "x = 0" | "x>0" | "x < 0" by linarith then show ?thesis proof (cases) case 1 have "continuous_on UNIV (\x. sgn x * \x\ powr (p - 1))" by (auto simp add: assms intro!:continuous_intros) then have "(\h. sgn h * \h\ powr (p-1)) \0\ (\h. sgn h * \h\ powr (p-1)) 0" using continuous_on_def by blast moreover have "\h\ powr p / h = sgn h * \h\ powr (p-1)" for h proof - have "\h\ powr p / h = sgn h * \h\ powr p / \h\" by (auto simp add: algebra_simps divide_simps sgn_real_def) also have "... = sgn h * \h\ powr (p-1)" using assms apply (cases "h = 0") apply (auto) by (metis abs_ge_zero powr_diff [symmetric] powr_one_gt_zero_iff times_divide_eq_right) finally show ?thesis by simp qed ultimately have "(\h. \h\ powr p / h) \0\ 0" by auto then show ?thesis unfolding DERIV_def by (auto simp add: \x = 0\) next case 2 have *: "\\<^sub>F y in nhds x. \y\ powr p = y powr p" unfolding eventually_nhds apply (rule exI[of _ "{0<..}"]) using \x > 0\ by auto show ?thesis apply (subst DERIV_cong_ev[of _ x _ "(\x. x powr p)" _ "p * x powr (p-1)"]) using \x > 0\ by (auto simp add: * has_real_derivative_powr) next case 3 have *: "\\<^sub>F y in nhds x. \y\ powr p = (-y) powr p" unfolding eventually_nhds apply (rule exI[of _ "{..<0}"]) using \x < 0\ by auto show ?thesis apply (subst DERIV_cong_ev[of _ x _ "(\x. (-x) powr p)" _ "p * (- x) powr (p - real 1) * - 1"]) using \x < 0\ apply (simp, simp add: *, simp) apply (rule DERIV_fun_powr[of "\y. -y" "-1" "x" p]) using \x < 0\ by (auto simp add: derivative_intros) qed qed lemma convex_abs_powr: assumes "p \ 1" shows "convex_on UNIV (\x::real. \x\ powr p)" proof (cases "p = 1") case True have "convex_on UNIV (\x::real. norm x)" by (rule convex_on_norm) moreover have "\x\ powr p = norm x" for x using True by auto ultimately show ?thesis by simp next case False then have "p > 1" using assms by auto define g where "g = (\x::real. p * sgn x * \x\ powr (p - 1))" have *: "DERIV (\x. \x\ powr p) x :> g x" for x unfolding g_def using \p>1\ by (intro derivative_intros) have **: "g x \ g y" if "x \ y" for x y proof - consider "x \ 0 \ y \ 0" | "x \ 0 \ y \ 0" | "x < 0 \ y > 0" using \x \ y\ by linarith then show ?thesis proof (cases) case 1 then show ?thesis unfolding g_def sgn_real_def using \p>1\ \x \ y\ by (auto simp add: powr_mono2) next case 2 then show ?thesis unfolding g_def sgn_real_def using \p>1\ \x \ y\ by (auto simp add: powr_mono2) next case 3 then have "g x \ 0" "0 \ g y" unfolding g_def using \p > 1\ by auto then show ?thesis by simp qed qed show ?thesis apply (rule convex_on_realI[of _ _ g]) using * ** by auto qed lemma convex_powr: assumes "p \ 1" shows "convex_on {0..} (\x::real. x powr p)" proof - have "convex_on {0..} (\x::real. \x\ powr p)" using convex_abs_powr[OF \p \ 1\] convex_on_subset by auto moreover have "\x\ powr p = x powr p" if "x \ {0..}" for x using that by auto ultimately show ?thesis by (simp add: convex_on_def) qed lemma convex_powr': assumes "p > 0" "p \ 1" shows "convex_on {0..} (\x::real. - (x powr p))" proof - have "convex_on {0<..} (\x::real. - (x powr p))" apply (rule convex_on_realI[of _ _ "\x. -p * x powr (p-1)"]) apply (auto intro!:derivative_intros simp add: has_real_derivative_powr) using \p > 0\ \p \ 1\ by (auto simp add: algebra_simps divide_simps powr_mono2') moreover have "continuous_on {0..} (\x::real. - (x powr p))" by (rule continuous_on_minus, rule continuous_on_powr', auto simp add: \p > 0\ intro!: continuous_intros) moreover have "{(0::real)..} = closure {0<..}" "convex {(0::real)<..}" by auto ultimately show ?thesis using convex_on_closure by metis qed lemma convex_fx_plus_fy_ineq: fixes f::"real \ real" assumes "convex_on {0..} f" "x \ 0" "y \ 0" "f 0 = 0" shows "f x + f y \ f (x+y)" proof - have *: "f a + f b \ f (a+b)" if "a \ 0" "b \ a" for a b proof (cases "a = 0") case False then have "a > 0" "b > 0" using \b \ a\ \a \ 0\ by auto have "(f 0 - f a) / (0 - a) \ (f 0 - f (a+b))/ (0 - (a+b))" apply (rule convex_on_diff[OF \convex_on {0..} f\]) using \a > 0\ \b > 0\ by auto also have "... \ (f b - f (a+b)) / (b - (a+b))" apply (rule convex_on_diff[OF \convex_on {0..} f\]) using \a > 0\ \b > 0\ by auto finally show ?thesis using \a > 0\ \b > 0\ \f 0 = 0\ by (auto simp add: divide_simps algebra_simps) qed (simp add: \f 0 = 0\) then show ?thesis using \x \ 0\ \y \ 0\ by (metis add.commute le_less not_le) qed lemma x_plus_y_p_le_xp_plus_yp: fixes p x y::real assumes "p > 0" "p \ 1" "x \ 0" "y \ 0" shows "(x + y) powr p \ x powr p + y powr p" using convex_fx_plus_fy_ineq[OF convex_powr'[OF \p > 0\ \p \ 1\] \x \ 0\ \y \ 0\] by auto subsection \Nonnegative-extended-real.thy\ lemma x_plus_top_ennreal [simp]: "x + \ = (\::ennreal)" by simp lemma ennreal_ge_nat_imp_PInf: fixes x::ennreal assumes "\N. x \ of_nat N" shows "x = \" using assms apply (cases x, auto) by (meson not_less reals_Archimedean2) lemma ennreal_archimedean: assumes "x \ (\::ennreal)" shows "\n::nat. x \ n" using assms ennreal_ge_nat_imp_PInf linear by blast lemma e2ennreal_mult: fixes a b::ereal assumes "a \ 0" shows "e2ennreal(a * b) = e2ennreal a * e2ennreal b" by (metis assms e2ennreal_neg eq_onp_same_args ereal_mult_le_0_iff linear times_ennreal.abs_eq) lemma e2ennreal_mult': fixes a b::ereal assumes "b \ 0" shows "e2ennreal(a * b) = e2ennreal a * e2ennreal b" using e2ennreal_mult[OF assms, of a] by (simp add: mult.commute) lemma SUP_real_ennreal: assumes "A \ {}" "bdd_above (f`A)" shows "(SUP a\A. ennreal (f a)) = ennreal(SUP a\A. f a)" apply (rule antisym, simp add: SUP_least assms(2) cSUP_upper ennreal_leI) by (metis assms(1) ennreal_SUP ennreal_less_top le_less) lemma e2ennreal_Liminf: "F \ bot \ e2ennreal (Liminf F f) = Liminf F (\n. e2ennreal (f n))" by (rule Liminf_compose_continuous_mono[symmetric]) (auto simp: mono_def e2ennreal_mono continuous_on_e2ennreal) lemma e2ennreal_eq_infty[simp]: "0 \ x \ e2ennreal x = top \ x = \" by (cases x) (auto) lemma ennreal_Inf_cmult: assumes "c>(0::real)" shows "Inf {ennreal c * x |x. P x} = ennreal c * Inf {x. P x}" proof - have "(\x::ennreal. c * x) (Inf {x::ennreal. P x}) = Inf ((\x::ennreal. c * x)`{x::ennreal. P x})" apply (rule mono_bij_Inf) apply (simp add: monoI mult_left_mono) apply (rule bij_betw_byWitness[of _ "\x. (x::ennreal) / c"], auto simp add: assms) apply (metis assms ennreal_lessI ennreal_neq_top mult.commute mult_divide_eq_ennreal not_less_zero) apply (metis assms divide_ennreal_def ennreal_less_zero_iff ennreal_neq_top less_irrefl mult.assoc mult.left_commute mult_divide_eq_ennreal) done then show ?thesis by (simp only: setcompr_eq_image[symmetric]) qed lemma continuous_on_const_minus_ennreal: fixes f :: "'a :: topological_space \ ennreal" shows "continuous_on A f \ continuous_on A (\x. a - f x)" including ennreal.lifting proof (transfer fixing: A; clarsimp) fix f :: "'a \ ereal" and a :: "ereal" assume "0 \ a" "\x. 0 \ f x" and f: "continuous_on A f" then show "continuous_on A (\x. max 0 (a - f x))" proof cases assume "\r. a = ereal r" with f show ?thesis by (auto simp: continuous_on_def minus_ereal_def ereal_Lim_uminus[symmetric] intro!: tendsto_add_ereal_general tendsto_max) next assume "\r. a = ereal r" with \0 \ a\ have "a = \" by (cases a) auto then show ?thesis by (simp add: continuous_on_const) qed qed lemma const_minus_Liminf_ennreal: fixes a :: ennreal shows "F \ bot \ a - Liminf F f = Limsup F (\x. a - f x)" by (intro Limsup_compose_continuous_antimono[symmetric]) (auto simp: antimono_def ennreal_mono_minus continuous_on_id continuous_on_const_minus_ennreal) lemma tendsto_cmult_ennreal [tendsto_intros]: fixes c l::ennreal assumes "\(c = \ \ l = 0)" "(f \ l) F" shows "((\x. c * f x) \ c * l) F" by (cases "c = 0", insert assms, auto intro!: tendsto_intros) subsection \Indicator-Function.thy\ text \There is something weird with \verb+sum_mult_indicator+: it is defined both in Indicator.thy and BochnerIntegration.thy, with a different meaning. I am surprised there is no name collision... Here, I am using the version from BochnerIntegration.\ lemma sum_indicator_eq_card2: assumes "finite I" shows "(\i\I. (indicator (P i) x)::nat) = card {i\I. x \ P i}" using sum_mult_indicator [OF assms, of "\y. 1::nat" P "\y. x"] unfolding card_eq_sum by auto +subclass (in zero_less_one) zero_neq_one + by standard (use zero_less_one in blast) + lemma disjoint_family_indicator_le_1: assumes "disjoint_family_on A I" shows "(\ i\ I. indicator (A i) x) \ (1::'a:: {comm_monoid_add,zero_less_one})" proof (cases "finite I") case True then have *: "(\ i\ I. indicator (A i) x) = ((indicator (\i\I. A i) x)::'a)" by (simp add: indicator_UN_disjoint[OF True assms(1), of x]) show ?thesis unfolding * unfolding indicator_def by (simp add: order_less_imp_le) next case False then show ?thesis by (simp add: order_less_imp_le) qed subsection \sigma-algebra.thy\ lemma algebra_intersection: assumes "algebra \ A" "algebra \ B" shows "algebra \ (A \ B)" apply (subst algebra_iff_Un) using assms by (auto simp add: algebra_iff_Un) lemma sigma_algebra_intersection: assumes "sigma_algebra \ A" "sigma_algebra \ B" shows "sigma_algebra \ (A \ B)" apply (subst sigma_algebra_iff) using assms by (auto simp add: sigma_algebra_iff algebra_intersection) lemma subalgebra_M_M [simp]: "subalgebra M M" by (simp add: subalgebra_def) text \The next one is \verb+disjoint_family_Suc+ with inclusions reversed.\ lemma disjoint_family_Suc2: assumes Suc: "\n. A (Suc n) \ A n" shows "disjoint_family (\i. A i - A (Suc i))" proof - have "A (m+n) \ A n" for m n proof (induct m) case 0 show ?case by simp next case (Suc m) then show ?case by (metis Suc_eq_plus1 assms add.commute add.left_commute subset_trans) qed then have "A m \ A n" if "m > n" for m n by (metis that add.commute le_add_diff_inverse nat_less_le) then show ?thesis by (auto simp add: disjoint_family_on_def) (metis insert_absorb insert_subset le_SucE le_antisym not_le_imp_less) qed subsection \Measure-Space.thy\ lemma AE_equal_sum: assumes "\i. AE x in M. f i x = g i x" shows "AE x in M. (\i\I. f i x) = (\i\I. g i x)" proof (cases) assume "finite I" have "\A. A \ null_sets M \ (\x\ (space M - A). f i x = g i x)" for i using assms(1)[of i] by (metis (mono_tags, lifting) AE_E3) then obtain A where A: "\i. A i \ null_sets M \ (\x\ (space M -A i). f i x = g i x)" by metis define B where "B = (\i\I. A i)" have "B \ null_sets M" using \finite I\ A B_def by blast then have "AE x in M. x \ space M - B" by (simp add: AE_not_in) moreover { fix x assume "x \ space M - B" then have "\i. i \ I \ f i x = g i x" unfolding B_def using A by auto then have "(\i\I. f i x) = (\i\I. g i x)" by auto } ultimately show ?thesis by auto qed (simp) lemma emeasure_pos_unionE: assumes "\ (N::nat). A N \ sets M" "emeasure M (\N. A N) > 0" shows "\N. emeasure M (A N) > 0" proof (rule ccontr) assume "\(\N. emeasure M (A N) > 0)" then have "\N. A N \ null_sets M" using assms(1) by auto then have "(\N. A N) \ null_sets M" by auto then show False using assms(2) by auto qed lemma (in prob_space) emeasure_intersection: fixes e::"nat \ real" assumes [measurable]: "\n. U n \ sets M" and [simp]: "\n. 0 \ e n" "summable e" and ge: "\n. emeasure M (U n) \ 1 - (e n)" shows "emeasure M (\n. U n) \ 1 - (\n. e n)" proof - define V where "V = (\n. space M - (U n))" have [measurable]: "V n \ sets M" for n unfolding V_def by auto have *: "emeasure M (V n) \ e n" for n unfolding V_def using ge[of n] by (simp add: emeasure_eq_measure prob_compl ennreal_leI) have "emeasure M (\n. V n) \ (\n. emeasure M (V n))" by (rule emeasure_subadditive_countably, auto) also have "... \ (\n. ennreal (e n))" using * by (intro suminf_le) auto also have "... = ennreal (\n. e n)" by (intro suminf_ennreal_eq) auto finally have "emeasure M (\n. V n) \ suminf e" by simp then have "1 - suminf e \ emeasure M (space M - (\n. V n))" by (simp add: emeasure_eq_measure prob_compl suminf_nonneg) also have "... \ emeasure M (\n. U n)" by (rule emeasure_mono) (auto simp: V_def) finally show ?thesis by simp qed lemma null_sym_diff_transitive: assumes "A \ B \ null_sets M" "B \ C \ null_sets M" and [measurable]: "A \ sets M" "C \ sets M" shows "A \ C \ null_sets M" proof - have "A \ B \ B \ C \ null_sets M" using assms(1) assms(2) by auto moreover have "A \ C \ A \ B \ B \ C" by auto ultimately show ?thesis by (meson null_sets_subset assms(3) assms(4) sets.Diff sets.Un) qed lemma Delta_null_of_null_is_null: assumes "B \ sets M" "A \ B \ null_sets M" "A \ null_sets M" shows "B \ null_sets M" proof - have "B \ A \ (A \ B)" by auto then show ?thesis using assms by (meson null_sets.Un null_sets_subset) qed lemma Delta_null_same_emeasure: assumes "A \ B \ null_sets M" and [measurable]: "A \ sets M" "B \ sets M" shows "emeasure M A = emeasure M B" proof - have "A = (A \ B) \ (A-B)" by blast moreover have "A-B \ null_sets M" using assms null_sets_subset by blast ultimately have a: "emeasure M A = emeasure M (A \ B)" using emeasure_Un_null_set by (metis assms(2) assms(3) sets.Int) have "B = (A \ B) \ (B-A)" by blast moreover have "B-A \ null_sets M" using assms null_sets_subset by blast ultimately have "emeasure M B = emeasure M (A \ B)" using emeasure_Un_null_set by (metis assms(2) assms(3) sets.Int) then show ?thesis using a by auto qed lemma AE_upper_bound_inf_ereal: fixes F G::"'a \ ereal" assumes "\e. (e::real) > 0 \ AE x in M. F x \ G x + e" shows "AE x in M. F x \ G x" proof - have "AE x in M. \n::nat. F x \ G x + ereal (1 / Suc n)" using assms by (auto simp: AE_all_countable) then show ?thesis proof (eventually_elim) fix x assume x: "\n::nat. F x \ G x + ereal (1 / Suc n)" show "F x \ G x" proof (intro ereal_le_epsilon2[of _ "G x"] allI impI) fix e :: real assume "0 < e" then obtain n where n: "1 / Suc n < e" by (blast elim: nat_approx_posE) have "F x \ G x + 1 / Suc n" using x by simp also have "\ \ G x + e" using n by (intro add_mono ennreal_leI) auto finally show "F x \ G x + ereal e" . qed qed qed text \Egorov theorem asserts that, if a sequence of functions converges almost everywhere to a limit, then the convergence is uniform on a subset of close to full measure. The first step in the proof is the following lemma, often useful by itself, asserting the same result for predicates: if a property $P_n x$ is eventually true for almost every $x$, then there exists $N$ such that $P_n x$ is true for all $n\geq N$ and all $x$ in a set of close to full measure. \ lemma (in finite_measure) Egorov_lemma: assumes [measurable]: "\n. (P n) \ measurable M (count_space UNIV)" and "AE x in M. eventually (\n. P n x) sequentially" "epsilon > 0" shows "\U N. U \ sets M \ (\n \ N. \x \ U. P n x) \ emeasure M (space M - U) < epsilon" proof - define K where "K = (\n. {x \ space M. \k\n. \(P k x)})" have [measurable]: "K n \ sets M" for n unfolding K_def by auto have "x \ (\n. K n)" if "eventually (\n. P n x) sequentially" for x unfolding K_def using that unfolding K_def eventually_sequentially by auto then have "AE x in M. x \ (\n. K n)" using assms by auto then have Z: "0 = emeasure M (\n. K n)" using AE_iff_measurable[of "(\n. K n)" M "\x. x \ (\n. K n)"] unfolding K_def by auto have *: "(\n. emeasure M (K n)) \ 0" unfolding Z apply (rule Lim_emeasure_decseq) using order_trans by (auto simp add: K_def decseq_def) have "eventually (\n. emeasure M (K n) < epsilon) sequentially" by (rule order_tendstoD(2)[OF * \epsilon > 0\]) then obtain N where N: "\n. n \ N \ emeasure M (K n) < epsilon" unfolding eventually_sequentially by auto define U where "U = space M - K N" have A [measurable]: "U \ sets M" unfolding U_def by auto have "space M - U = K N" unfolding U_def K_def by auto then have B: "emeasure M (space M - U) < epsilon" using N by auto have "\n \ N. \x \ U. P n x" unfolding U_def K_def by auto then show ?thesis using A B by blast qed text \The next lemma asserts that, in an uncountable family of disjoint sets, then there is one set with zero measure (and in fact uncountably many). It is often applied to the boundaries of $r$-neighborhoods of a given set, to show that one could choose $r$ for which this boundary has zero measure (this shows up often in relation with weak convergence).\ lemma (in finite_measure) uncountable_disjoint_family_then_exists_zero_measure: assumes [measurable]: "\i. i \ I \ A i \ sets M" and "uncountable I" "disjoint_family_on A I" shows "\i\I. measure M (A i) = 0" proof - define f where "f = (\(r::real). {i \ I. measure M (A i) > r})" have *: "finite (f r)" if "r > 0" for r proof - obtain N::nat where N: "measure M (space M)/r \ N" using real_arch_simple by blast have "finite (f r) \ card (f r) \ N" proof (rule finite_if_finite_subsets_card_bdd) fix G assume G: "G \ f r" "finite G" then have "G \ I" unfolding f_def by auto have "card G * r = (\i \ G. r)" by auto also have "... \ (\i \ G. measure M (A i))" apply (rule sum_mono) using G unfolding f_def by auto also have "... = measure M (\i\G. A i)" apply (rule finite_measure_finite_Union[symmetric]) using \finite G\ \G \ I\ \disjoint_family_on A I\ disjoint_family_on_mono by auto also have "... \ measure M (space M)" by (simp add: bounded_measure) finally have "card G \ measure M (space M)/r" using \r > 0\ by (simp add: divide_simps) then show "card G \ N" using N by auto qed then show ?thesis by simp qed have "countable (\n. f (((1::real)/2)^n))" by (rule countable_UN, auto intro!: countable_finite *) then have "I - (\n. f (((1::real)/2)^n)) \ {}" using assms(2) by (metis countable_empty uncountable_minus_countable) then obtain i where "i \ I" "i \ (\n. f ((1/2)^n))" by auto then have "measure M (A i) \ (1 / 2) ^ n" for n unfolding f_def using linorder_not_le by auto moreover have "(\n. ((1::real) / 2) ^ n) \ 0" by (intro tendsto_intros, auto) ultimately have "measure M (A i) \ 0" using LIMSEQ_le_const by force then have "measure M (A i) = 0" by (simp add: measure_le_0_iff) then show ?thesis using \i \ I\ by auto qed text \The next statements are useful measurability statements.\ lemma measurable_Inf [measurable]: assumes [measurable]: "\(n::nat). P n \ measurable M (count_space UNIV)" shows "(\x. Inf {n. P n x}) \ measurable M (count_space UNIV)" (is "?f \ _") proof - define A where "A = (\n. (P n)-`{True} \ space M - (\m space M))" have A_meas [measurable]: "A n \ sets M" for n unfolding A_def by measurable define B where "B = (\n. if n = 0 then (space M - (\n. A n)) else A (n-1))" show ?thesis proof (rule measurable_piecewise_restrict2[of B]) show "B n \ sets M" for n unfolding B_def by simp show "space M = (\n. B n)" unfolding B_def using sets.sets_into_space [OF A_meas] by auto have *: "?f x = n" if "x \ A n" for x n apply (rule cInf_eq_minimum) using that unfolding A_def by auto moreover have **: "?f x = (Inf ({}::nat set))" if "x \ space M - (\n. A n)" for x proof - have "\(P n x)" for n apply (induction n rule: nat_less_induct) using that unfolding A_def by auto then show ?thesis by simp qed ultimately have "\c. \x \ B n. ?f x = c" for n apply (cases "n = 0") unfolding B_def by auto then show "\h \ measurable M (count_space UNIV). \x \ B n. ?f x = h x" for n by fastforce qed qed lemma measurable_T_iter [measurable]: fixes f::"'a \ nat" assumes [measurable]: "T \ measurable M M" "f \ measurable M (count_space UNIV)" shows "(\x. (T^^(f x)) x) \ measurable M M" proof - have [measurable]: "(T^^n) \ measurable M M" for n::nat by (induction n, auto) show ?thesis by (rule measurable_compose_countable, auto) qed lemma measurable_infdist [measurable]: "(\x. infdist x S) \ borel_measurable borel" by (rule borel_measurable_continuous_onI, intro continuous_intros) text \The next lemma shows that, in a sigma finite measure space, sets with large measure can be approximated by sets with large but finite measure.\ lemma (in sigma_finite_measure) approx_with_finite_emeasure: assumes W_meas: "W \ sets M" and W_inf: "emeasure M W > C" obtains Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" proof (cases "emeasure M W = \") case True obtain r where r: "C = ennreal r" using W_inf by (cases C, auto) obtain Z where "Z \ sets M" "Z \ W" "emeasure M Z < \" "emeasure M Z > C" unfolding r using approx_PInf_emeasure_with_finite[OF W_meas True, of r] by auto then show ?thesis using that by blast next case False then have "W \ sets M" "W \ W" "emeasure M W < \" "emeasure M W > C" using assms apply auto using top.not_eq_extremum by blast then show ?thesis using that by blast qed subsection \Nonnegative-Lebesgue-Integration.thy\ text \The next lemma is a variant of \verb+nn_integral_density+, with the density on the right instead of the left, as seems more common.\ lemma nn_integral_densityR: assumes [measurable]: "f \ borel_measurable F" "g \ borel_measurable F" shows "(\\<^sup>+ x. f x * g x \F) = (\\<^sup>+ x. f x \(density F g))" proof - have "(\\<^sup>+ x. f x * g x \F) = (\\<^sup>+ x. g x * f x \F)" by (simp add: mult.commute) also have "... = (\\<^sup>+ x. f x \(density F g))" by (rule nn_integral_density[symmetric], simp_all add: assms) finally show ?thesis by simp qed lemma not_AE_zero_int_ennreal_E: fixes f::"'a \ ennreal" assumes "(\\<^sup>+x. f x \M) > 0" and [measurable]: "f \ borel_measurable M" shows "\A\sets M. \e::real>0. emeasure M A > 0 \ (\x \ A. f x \ e)" proof (rule not_AE_zero_ennreal_E, auto simp add: assms) assume *: "AE x in M. f x = 0" have "(\\<^sup>+x. f x \M) = (\\<^sup>+x. 0 \M)" by (rule nn_integral_cong_AE, simp add: *) then have "(\\<^sup>+x. f x \M) = 0" by simp then show False using assms by simp qed lemma (in finite_measure) nn_integral_bounded_eq_bound_then_AE: assumes "AE x in M. f x \ ennreal c" "(\\<^sup>+x. f x \M) = c * emeasure M (space M)" and [measurable]: "f \ borel_measurable M" shows "AE x in M. f x = c" proof (cases) assume "emeasure M (space M) = 0" then show ?thesis by (rule emeasure_0_AE) next assume "emeasure M (space M) \ 0" have fin: "AE x in M. f x \ top" using assms by (auto simp: top_unique) define g where "g = (\x. c - f x)" have [measurable]: "g \ borel_measurable M" unfolding g_def by auto have "(\\<^sup>+x. g x \M) = (\\<^sup>+x. c \M) - (\\<^sup>+x. f x \M)" unfolding g_def by (rule nn_integral_diff, auto simp add: assms ennreal_mult_eq_top_iff) also have "\ = 0" using assms(2) by (auto simp: ennreal_mult_eq_top_iff) finally have "AE x in M. g x = 0" by (subst nn_integral_0_iff_AE[symmetric]) auto then have "AE x in M. c \ f x" unfolding g_def using fin by (auto simp: ennreal_minus_eq_0) then show ?thesis using assms(1) by auto qed lemma null_sets_density: assumes [measurable]: "h \ borel_measurable M" and "AE x in M. h x \ 0" shows "null_sets (density M h) = null_sets M" proof - have *: "A \ sets M \ (AE x\A in M. h x = 0) \ A \ null_sets M" for A proof (auto) assume "A \ sets M" "AE x\A in M. h x = 0" then show "A \ null_sets M" unfolding AE_iff_null_sets[OF \A \ sets M\] using assms(2) by auto next assume "A \ null_sets M" then show "AE x\A in M. h x = 0" by (metis (mono_tags, lifting) AE_not_in eventually_mono) qed show ?thesis apply (rule set_eqI) unfolding null_sets_density_iff[OF \h \ borel_measurable M\] using * by auto qed text \The next proposition asserts that, if a function $h$ is integrable, then its integral on any set with small enough measure is small. The good conceptual proof is by considering the distribution of the function $h$ on $\mathbb{R}$ and looking at its tails. However, there is a less conceptual but more direct proof, based on dominated convergence and a proof by contradiction. This is the proof we give below.\ proposition integrable_small_integral_on_small_sets: fixes h::"'a \ real" assumes [measurable]: "integrable M h" and "delta > 0" shows "\epsilon>(0::real). \U \ sets M. emeasure M U < epsilon \ abs (\x\U. h x \M) < delta" proof (rule ccontr) assume H: "\ (\epsilon>0. \U\sets M. emeasure M U < ennreal epsilon \ abs(set_lebesgue_integral M U h) < delta)" have "\f. \epsilon\{0<..}. f epsilon \sets M \ emeasure M (f epsilon) < ennreal epsilon \ \(abs(set_lebesgue_integral M (f epsilon) h) < delta)" apply (rule bchoice) using H by auto then obtain f::"real \ 'a set" where f: "\epsilon. epsilon > 0 \ f epsilon \sets M" "\epsilon. epsilon > 0 \ emeasure M (f epsilon) < ennreal epsilon" "\epsilon. epsilon > 0 \ \(abs(set_lebesgue_integral M (f epsilon) h) < delta)" by blast define A where "A = (\n::nat. f ((1/2)^n))" have [measurable]: "A n \ sets M" for n unfolding A_def using f(1) by auto have *: "emeasure M (A n) < ennreal ((1/2)^n)" for n unfolding A_def using f(2) by auto have Large: "\(abs(set_lebesgue_integral M (A n) h) < delta)" for n unfolding A_def using f(3) by auto have S: "summable (\n. Sigma_Algebra.measure M (A n))" apply (rule summable_comparison_test'[of "\n. (1/2)^n" 0]) apply (rule summable_geometric, auto) apply (subst ennreal_le_iff[symmetric], simp) using less_imp_le[OF *] by (metis * emeasure_eq_ennreal_measure top.extremum_strict) have "AE x in M. eventually (\n. x \ space M - A n) sequentially" apply (rule borel_cantelli_AE1, auto simp add: S) by (metis * top.extremum_strict top.not_eq_extremum) moreover have "(\n. indicator (A n) x * h x) \ 0" if "eventually (\n. x \ space M - A n) sequentially" for x proof - have "eventually (\n. indicator (A n) x * h x = 0) sequentially" apply (rule eventually_mono[OF that]) unfolding indicator_def by auto then show ?thesis unfolding eventually_sequentially using lim_explicit by force qed ultimately have A: "AE x in M. ((\n. indicator (A n) x * h x) \ 0)" by auto have I: "integrable M (\x. abs(h x))" using \integrable M h\ by auto have L: "(\n. abs (\x. indicator (A n) x * h x \M)) \ abs (\x. 0 \M)" apply (intro tendsto_intros) apply (rule integral_dominated_convergence[OF _ _ I A]) unfolding indicator_def by auto have "eventually (\n. abs (\x. indicator (A n) x * h x \M) < delta) sequentially" apply (rule order_tendstoD[OF L]) using \delta > 0\ by auto then show False using Large by (auto simp: set_lebesgue_integral_def) qed text \We also give the version for nonnegative ennreal valued functions. It follows from the previous one.\ proposition small_nn_integral_on_small_sets: fixes h::"'a \ ennreal" assumes [measurable]: "h \ borel_measurable M" and "delta > (0::real)" "(\\<^sup>+x. h x \M) \ \" shows "\epsilon>(0::real). \U \ sets M. emeasure M U < epsilon \ (\\<^sup>+x\U. h x \M) < delta" proof - define f where "f = (\x. enn2real(h x))" have "AE x in M. h x \ \" using assms by (metis nn_integral_PInf_AE) then have *: "AE x in M. ennreal (f x) = h x" unfolding f_def using ennreal_enn2real_if by auto have **: "(\\<^sup>+x. ennreal (f x) \M) \ \" using nn_integral_cong_AE[OF *] assms by auto have [measurable]: "f \ borel_measurable M" unfolding f_def by auto have "integrable M f" apply (rule integrableI_nonneg) using assms * f_def ** apply auto using top.not_eq_extremum by blast obtain epsilon::real where H: "epsilon > 0" "\U. U \ sets M \ emeasure M U < epsilon \ abs(\x\U. f x \M) < delta" using integrable_small_integral_on_small_sets[OF \integrable M f\ \delta > 0\] by blast have "(\\<^sup>+x\U. h x \M) < delta" if [measurable]: "U \ sets M" "emeasure M U < epsilon" for U proof - have "(\\<^sup>+x. indicator U x * h x \M) = (\\<^sup>+x. ennreal(indicator U x * f x) \M)" apply (rule nn_integral_cong_AE) using * unfolding indicator_def by auto also have "... = ennreal (\x. indicator U x * f x \M)" apply (rule nn_integral_eq_integral) apply (rule Bochner_Integration.integrable_bound[OF \integrable M f\]) unfolding indicator_def f_def by auto also have "... < ennreal delta" apply (rule ennreal_lessI) using H(2)[OF that] by (auto simp: set_lebesgue_integral_def) finally show ?thesis by (auto simp add: mult.commute) qed then show ?thesis using \epsilon > 0\ by auto qed subsection \Probability-measure.thy\ text \The next lemmas ensure that, if sets have a probability close to $1$, then their intersection also does.\ lemma (in prob_space) sum_measure_le_measure_inter: assumes "A \ sets M" "B \ sets M" shows "prob A + prob B \ 1 + prob (A \ B)" proof - have "prob A + prob B = prob (A \ B) + prob (A \ B)" by (simp add: assms fmeasurable_eq_sets measure_Un3) also have "... \ 1 + prob (A \ B)" by auto finally show ?thesis by simp qed lemma (in prob_space) sum_measure_le_measure_inter3: assumes [measurable]: "A \ sets M" "B \ sets M" "C \ sets M" shows "prob A + prob B + prob C \ 2 + prob (A \ B \ C)" using sum_measure_le_measure_inter[of B C] sum_measure_le_measure_inter[of A "B \ C"] by (auto simp add: inf_assoc) lemma (in prob_space) sum_measure_le_measure_Inter: assumes [measurable]: "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" shows "(\i\I. prob (A i)) \ real(card I) - 1 + prob (\i\I. A i)" using assms proof (induct I rule: finite_ne_induct) fix x F assume H: "finite F" "F \ {}" "x \ F" "((\i. i \ F \ A i \ events) \ (\i\F. prob (A i)) \ real (card F) - 1 + prob (\(A ` F)))" and [measurable]: "(\i. i \ insert x F \ A i \ events)" have "(\x\F. A x) \ events" using \finite F\ \F \ {}\ by auto have "(\i\insert x F. prob (A i)) = (\i\F. prob (A i)) + prob (A x)" using H(1) H(3) by auto also have "... \ real (card F)-1 + prob (\(A ` F)) + prob (A x)" using H(4) by auto also have "... \ real (card F) + prob ((\(A ` F)) \ A x)" using sum_measure_le_measure_inter[OF \(\x\F. A x) \ events\, of "A x"] by auto also have "... = real (card (insert x F)) - 1 + prob (\(A ` (insert x F)))" using H(1) H(2) unfolding card_insert_disjoint[OF \finite F\ \x \ F\] by (simp add: inf_commute) finally show "(\i\insert x F. prob (A i)) \ real (card (insert x F)) - 1 + prob (\(A ` (insert x F)))" by simp qed (auto) text \A random variable gives a small mass to small neighborhoods of infinity.\ lemma (in prob_space) random_variable_small_tails: assumes "alpha > 0" and [measurable]: "f \ borel_measurable M" shows "\(C::real). prob {x \ space M. abs(f x) \ C} < alpha \ C \ K" proof - have *: "(\(n::nat). {x\space M. abs(f x) \ n}) = {}" apply auto by (metis real_arch_simple add.right_neutral add_mono_thms_linordered_field(4) not_less zero_less_one) have **: "(\n. prob {x \ space M. abs(f x) \ n}) \ prob (\(n::nat). {x \ space M. abs(f x) \ n})" by (rule finite_Lim_measure_decseq, auto simp add: decseq_def) have "eventually (\n. prob {x \ space M. abs(f x) \ n} < alpha) sequentially" apply (rule order_tendstoD[OF _ \alpha > 0\]) using ** unfolding * by auto then obtain N::nat where N: "\n::nat. n \ N \ prob {x \ space M. abs(f x) \ n} < alpha" unfolding eventually_sequentially by blast have "\n::nat. n \ N \ n \ K" by (meson le_cases of_nat_le_iff order.trans real_arch_simple) then obtain n::nat where n: "n \ N" "n \ K" by blast show ?thesis apply (rule exI[of _ "of_nat n"]) using N n by auto qed subsection \Distribution-functions.thy\ text \There is a locale called \verb+finite_borel_measure+ in \verb+distribution-functions.thy+. However, it only deals with real measures, and real weak convergence. I will not need the weak convergence in more general settings, but still it seems more natural to me to do the proofs in the natural settings. Let me introduce the locale \verb+finite_borel_measure'+ for this, although it would be better to rename the locale in the library file.\ locale finite_borel_measure' = finite_measure M for M :: "('a::metric_space) measure" + assumes M_is_borel [simp, measurable_cong]: "sets M = sets borel" begin lemma space_eq_univ [simp]: "space M = UNIV" using M_is_borel[THEN sets_eq_imp_space_eq] by simp lemma measurable_finite_borel [simp]: "f \ borel_measurable borel \ f \ borel_measurable M" by (rule borel_measurable_subalgebra[where N = borel]) auto text \Any closed set can be slightly enlarged to obtain a set whose boundary has $0$ measure.\ lemma approx_closed_set_with_set_zero_measure_boundary: assumes "closed S" "epsilon > 0" "S \ {}" shows "\r. r < epsilon \ r > 0 \ measure M {x. infdist x S = r} = 0 \ measure M {x. infdist x S \ r} < measure M S + epsilon" proof - have [measurable]: "S \ sets M" using \closed S\ by auto define T where "T = (\r. {x. infdist x S \ r})" have [measurable]: "T r \ sets borel" for r unfolding T_def by measurable have *: "(\n. T ((1/2)^n)) = S" unfolding T_def proof (auto) fix x assume *: "\n. infdist x S \ (1 / 2) ^n" have "infdist x S \ 0" apply (rule LIMSEQ_le_const[of "\n. (1/2)^n"], intro tendsto_intros) using * by auto then show "x \ S" using assms infdist_pos_not_in_closed by fastforce qed have A: "((1::real)/2)^n \ (1/2)^m" if "m \ n" for m n::nat using that by (simp add: power_decreasing) have "(\n. measure M (T ((1/2)^n))) \ measure M S" unfolding *[symmetric] apply (rule finite_Lim_measure_decseq, auto simp add: T_def decseq_def) using A order.trans by blast then have B: "eventually (\n. measure M (T ((1/2)^n)) < measure M S + epsilon) sequentially" apply (rule order_tendstoD) using \epsilon > 0\ by simp have C: "eventually (\n. (1/2)^n < epsilon) sequentially" by (rule order_tendstoD[OF _ \epsilon > 0\], intro tendsto_intros, auto) obtain n where n: "(1/2)^n < epsilon" "measure M (T ((1/2)^n)) < measure M S + epsilon" using eventually_conj[OF B C] unfolding eventually_sequentially by auto have "\r\{0<..<(1/2)^n}. measure M {x. infdist x S = r} = 0" apply (rule uncountable_disjoint_family_then_exists_zero_measure, auto simp add: disjoint_family_on_def) using uncountable_open_interval by fastforce then obtain r where r: "r\{0<..<(1/2)^n}" "measure M {x. infdist x S = r} = 0" by blast then have r2: "r > 0" "r < epsilon" using n by auto have "measure M {x. infdist x S \ r} \ measure M {x. infdist x S \ (1/2)^n}" apply (rule finite_measure_mono) using r by auto then have "measure M {x. infdist x S \ r} < measure M S + epsilon" using n(2) unfolding T_def by auto then show ?thesis using r(2) r2 by auto qed end (* of locale finite_borel_measure'*) sublocale finite_borel_measure \ finite_borel_measure' by (standard, simp add: M_is_borel) subsection \Weak-convergence.thy\ text \Since weak convergence is not implemented as a topology, the fact that the convergence of a sequence implies the convergence of a subsequence is not automatic. We prove it in the lemma below..\ lemma weak_conv_m_subseq: assumes "weak_conv_m M_seq M" "strict_mono r" shows "weak_conv_m (\n. M_seq (r n)) M" using assms LIMSEQ_subseq_LIMSEQ unfolding weak_conv_m_def weak_conv_def comp_def by auto context fixes \ :: "nat \ real measure" and M :: "real measure" assumes \: "\n. real_distribution (\ n)" assumes M: "real_distribution M" assumes \_to_M: "weak_conv_m \ M" begin text \The measure of a closed set behaves upper semicontinuously with respect to weak convergence: if $\mu_n \to \mu$, then $\limsup \mu_n(F) \leq \mu(F)$ (and the inequality can be strict, think of the situation where $\mu$ is a Dirac mass at $0$ and $F = \{0\}$, but $\mu_n$ has a density so that $\mu_n(\{0\}) = 0$).\ lemma closed_set_weak_conv_usc: assumes "closed S" "measure M S < l" shows "eventually (\n. measure (\ n) S < l) sequentially" proof (cases "S = {}") case True then show ?thesis using \measure M S < l\ by auto next case False interpret real_distribution M using M by simp define epsilon where "epsilon = l - measure M S" have "epsilon > 0" unfolding epsilon_def using assms(2) by auto obtain r where r: "r > 0" "r < epsilon" "measure M {x. infdist x S = r} = 0" "measure M {x. infdist x S \ r} < measure M S + epsilon" using approx_closed_set_with_set_zero_measure_boundary[OF \closed S\ \epsilon > 0\ \S \ {}\] by blast define T where "T = {x. infdist x S \ r}" have [measurable]: "T \ sets borel" unfolding T_def by auto have "S \ T" unfolding T_def using \closed S\ \r > 0\ by auto have "measure M T < l" using r(4) unfolding T_def epsilon_def by auto have "measure M (frontier T) \ measure M {x. infdist x S = r}" apply (rule finite_measure_mono) unfolding T_def using frontier_indist_le by auto then have "measure M (frontier T) = 0" using \measure M {x. infdist x S = r} = 0\ by (auto simp add: measure_le_0_iff) then have "(\n. measure (\ n) T) \ measure M T" using \_to_M by (simp add: \ emeasure_eq_measure real_distribution_axioms weak_conv_imp_continuity_set_conv) then have *: "eventually (\n. measure (\ n) T < l) sequentially" apply (rule order_tendstoD) using \measure M T < l\ by simp have **: "measure (\ n) S \ measure (\ n) T" for n apply (rule finite_measure.finite_measure_mono) using \ apply (simp add: finite_borel_measure.axioms(1) real_distribution.finite_borel_measure_M) using \S \ T\ apply simp by (simp add: \ real_distribution.events_eq_borel) show ?thesis apply (rule eventually_mono[OF *]) using ** le_less_trans by auto qed text \In the same way, the measure of an open set behaves lower semicontinuously with respect to weak convergence: if $\mu_n \to \mu$, then $\liminf \mu_n(U) \geq \mu(U)$ (and the inequality can be strict). This follows from the same statement for closed sets by passing to the complement.\ lemma open_set_weak_conv_lsc: assumes "open S" "measure M S > l" shows "eventually (\n. measure (\ n) S > l) sequentially" proof - interpret real_distribution M using M by auto have [measurable]: "S \ events" using assms(1) by auto have "eventually (\n. measure (\ n) (UNIV - S) < 1 - l) sequentially" apply (rule closed_set_weak_conv_usc) using assms prob_compl[of S] by auto moreover have "measure (\ n) (UNIV - S) = 1 - measure (\ n) S" for n proof - interpret mu: real_distribution "\ n" using \ by auto have "S \ mu.events" using assms(1) by auto then show ?thesis using mu.prob_compl[of S] by auto qed ultimately show ?thesis by auto qed end (*of context weak_conv_m*) end (*of SG_Library_Complement.thy*) diff --git a/thys/Green/Green.thy b/thys/Green/Green.thy --- a/thys/Green/Green.thy +++ b/thys/Green/Green.thy @@ -1,3117 +1,3117 @@ theory Green imports Paths Derivs Integrals General_Utils begin lemma frontier_Un_subset_Un_frontier: "frontier (s \ t) \ (frontier s) \ (frontier t)" by (simp add: frontier_def Un_Diff) (auto simp add: closure_def interior_def islimpt_def) definition has_partial_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ('a \ 'b) \ ('a) \ bool" where "has_partial_derivative F base_vec F' a \ ((\x::'a::euclidean_space. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + (x \ base_vec) *\<^sub>R base_vec )) has_derivative F') (at a)" definition has_partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ ( 'b) \ ('a) \ bool" where "has_partial_vector_derivative F base_vec F' a \ ((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" definition partially_vector_differentiable where "partially_vector_differentiable F base_vec p \ (\F'. has_partial_vector_derivative F base_vec F' p)" definition partial_vector_derivative:: "(('a::euclidean_space) \ 'b::euclidean_space) \ 'a \ 'a \ 'b" where "partial_vector_derivative F base_vec a \ (vector_derivative (\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec)) (at (a \ base_vec)))" lemma partial_vector_derivative_works: assumes "partially_vector_differentiable F base_vec a" shows "has_partial_vector_derivative F base_vec (partial_vector_derivative F base_vec a) a" proof - obtain F' where F'_prop: "((\x. F( (a - ((a \ base_vec) *\<^sub>R base_vec)) + x *\<^sub>R base_vec )) has_vector_derivative F') (at (a \ base_vec))" using assms partially_vector_differentiable_def has_partial_vector_derivative_def by blast show ?thesis using Derivative.differentiableI_vector[OF F'_prop] by(simp add: vector_derivative_works partial_vector_derivative_def[symmetric] has_partial_vector_derivative_def[symmetric]) qed lemma fundamental_theorem_of_calculus_partial_vector: fixes a b:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and j:: "'b" and F_j_i:: "('a::euclidean_space \ real)" assumes a_leq_b: "a \ b" and Base_vecs: "i \ Basis" "j \ Basis" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i p) p" and domain_subset_of_D: "{x *\<^sub>R i + c |x. a \ x \ x \ b} \ D" shows "((\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" proof - let ?domain = "{v. \x. a \ x \ x \ b \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative (\x. (F x) \ j) i (F_j_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at x)" proof(clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "a \ x" "x \ b" assume ass2: "\x. (\z\a. z \ b \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" by (simp add: assms(2) inner_left_distrib no_i_component) have 1: "(\x. F(x *\<^sub>R i + c) \ j) = (\x. F(c + x *\<^sub>R i) \ j)" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative F_j_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox a b). ((\x. F(x *\<^sub>R i + c) \ j) has_vector_derivative (F_j_i( x *\<^sub>R i + c))) (at_within x (cbox a b))" using has_vector_derivative_at_within by blast then show "( (\x. F_j_i( x *\<^sub>R i + c)) has_integral F(b *\<^sub>R i + c) \ j - F(a *\<^sub>R i + c) \ j) (cbox a b)" using fundamental_theorem_of_calculus[of "a" "b" "(\x::real. F(x *\<^sub>R i + c) \ j)" "(\x::real. F_j_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma fundamental_theorem_of_calculus_partial_vector_gen: fixes k1 k2:: "real" and F:: "('a::euclidean_space \ 'b::euclidean_space)" and i:: "'a" and F_i:: "('a::euclidean_space \ 'b)" assumes a_leq_b: "k1 \ k2" and unit_len: "i \ i = 1" and no_i_component: "c \ i = 0 " and has_partial_deriv: "\p \ D. has_partial_vector_derivative F i (F_i p) p" and domain_subset_of_D: "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c} \ D" shows "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" proof - let ?domain = "{v. \x. k1 \ x \ x \ k2 \ v = x *\<^sub>R i + c}" have "\x\ ?domain. has_partial_vector_derivative F i (F_i x) x" using has_partial_deriv domain_subset_of_D by auto then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c)) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at x)" proof (clarsimp simp add: has_partial_vector_derivative_def) fix x::real assume range_of_x: "k1 \ x" "x \ k2" assume ass2: "\x. (\z\k1. z \ k2 \ x = z *\<^sub>R i + c) \ ((\z. F(x - (x \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i x) (at (x \ i))" have "((\z. F((x *\<^sub>R i + c) - ((x *\<^sub>R i + c) \ i) *\<^sub>R i + z *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at ((x *\<^sub>R i + c) \ i)) " using range_of_x ass2 by auto then have 0: "((\x. F( c + x *\<^sub>R i)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" by (simp add: inner_commute inner_right_distrib no_i_component unit_len) have 1: "(\x. F(x *\<^sub>R i + c)) = (\x. F(c + x *\<^sub>R i))" by (simp add: add.commute) then show "((\x. F(x *\<^sub>R i + c)) has_vector_derivative F_i (x *\<^sub>R i + c)) (at x)" using 0 and 1 by auto qed then have "\x\ (cbox k1 k2). ((\x. F(x *\<^sub>R i + c) ) has_vector_derivative (F_i( x *\<^sub>R i + c))) (at_within x (cbox k1 k2))" using has_vector_derivative_at_within by blast then show "( (\x. F_i( x *\<^sub>R i + c)) has_integral F(k2 *\<^sub>R i + c) - F(k1 *\<^sub>R i + c)) (cbox k1 k2)" using fundamental_theorem_of_calculus[of "k1" "k2" "(\x::real. F(x *\<^sub>R i + c))" "(\x::real. F_i( x *\<^sub>R i + c))"] and a_leq_b by auto qed lemma add_scale_img: assumes "a < b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using assms apply (auto simp: algebra_simps affine_ineq image_iff) using less_eq_real_def apply force apply (rule_tac x="(x-a)/(b-a)" in bexI) apply (auto simp: field_simps) done lemma add_scale_img': assumes "a \ b" shows "(\x::real. a + (b - a) * x) ` {0 .. 1} = {a .. b}" proof (cases "a = b") case True then show ?thesis by force next case False then show ?thesis using add_scale_img assms by auto qed -definition analytically_valid:: "'a::euclidean_space set \ ('a \ 'b::{euclidean_space,times,one}) \ 'a \ bool" where +definition analytically_valid:: "'a::euclidean_space set \ ('a \ 'b::{euclidean_space,times,zero_neq_one}) \ 'a \ bool" where "analytically_valid s F i \ (\a \ s. partially_vector_differentiable F i a) \ continuous_on s F \ \ \TODO: should we replace this with saying that \F\ is partially diffrerentiable on \Dy\,\ \ \i.e. there is a partial derivative on every dimension\ integrable lborel (\p. (partial_vector_derivative F i) p * indicator s p) \ (\x. integral UNIV (\y. (partial_vector_derivative F i (y *\<^sub>R i + x *\<^sub>R (\ b \(Basis - {i}). b))) * (indicator s (y *\<^sub>R i + x *\<^sub>R (\b \ Basis - {i}. b)) ))) \ borel_measurable lborel" (*(\x. integral UNIV (\y. ((partial_vector_derivative F i) (y, x)) * (indicator s (y, x)))) \ borel_measurable lborel)"*) lemma analytically_valid_imp_part_deriv_integrable_on: assumes "analytically_valid (s::(real*real) set) (f::(real*real)\ real) i" shows "(partial_vector_derivative f i) integrable_on s" proof - have "integrable lborel (\p. partial_vector_derivative f i p * indicator s p)" using assms by(simp add: analytically_valid_def indic_ident) then have "integrable lborel (\p::(real*real). if p \ s then partial_vector_derivative f i p else 0)" using indic_ident[of "partial_vector_derivative f i"] by (simp add: indic_ident) then have "(\x. if x \ s then partial_vector_derivative f i x else 0) integrable_on UNIV" using Equivalence_Lebesgue_Henstock_Integration.integrable_on_lborel by auto then show "(partial_vector_derivative f i) integrable_on s" using integrable_restrict_UNIV by auto qed (*******************************************************************************) definition typeII_twoCube :: "((real * real) \ (real * real)) \ bool" where "typeII_twoCube twoC \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)), (1-x)*a + x*b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b}" abbreviation unit_cube where "unit_cube \ cbox (0,0) (1::real,1::real)" definition cubeImage:: "two_cube \ ((real*real) set)" where "cubeImage twoC \ (twoC ` unit_cube)" lemma typeII_twoCubeImg: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y,x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " using assms proof (simp add: typeII_twoCube_def cubeImage_def image_def) assume " \a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" then obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ab1: "a - x * a + x * b \ b" if a1: "0 \ x" "x \ 1" for x using that apply (simp add: algebra_simps) by (metis affine_ineq less_eq_real_def mult.commute twoCisTypeII(1)) have ex: "\z\Green.unit_cube. (d, c) = (case z of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have b_minus_a_nzero: "b - a \ 0" using twoCisTypeII(1) by auto have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" apply (rule_tac x="(c - a)/(b - a)" in exI) using c_bounds \a < b\ apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by blast have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by (fastforce simp add: algebra_simps) next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeII(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False sum_sqs_eq by (fastforce simp add: divide_simps algebra_simps) with k2_in_bounds show ?thesis by fastforce qed show "\x\unit_cube. (d, c) = (case x of (y, x) \ (g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b), a - x * a + x * b))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeII ab1 left_diff_distrib ex) using order.order_iff_strict twoCisTypeII(1) apply fastforce apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeII)+ done then show "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ {y. \x\unit_cube. y = twoC x} = {(y, x). a \ x \ x \ b \ g2 x \ y \ y \ g1 x} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b})" using twoCisTypeII by blast qed definition horizontal_boundary :: "two_cube \ one_chain" where "horizontal_boundary twoC \ {(1, (\x. twoC(x,0))), (-1, (\x. twoC(x,1)))}" definition vertical_boundary :: "two_cube \ one_chain" where "vertical_boundary twoC \ {(-1, (\y. twoC(0,y))), (1, (\y. twoC(1,y)))}" definition boundary :: "two_cube \ one_chain" where "boundary twoC \ horizontal_boundary twoC \ vertical_boundary twoC" definition valid_two_cube where "valid_two_cube twoC \ card (boundary twoC) = 4" definition two_chain_integral:: "two_chain \ ((real*real)\(real)) \ real" where "two_chain_integral twoChain F \ \C\twoChain. (integral (cubeImage C) F)" definition valid_two_chain where "valid_two_chain twoChain \ (\twoCube \ twoChain. valid_two_cube twoCube) \ pairwise (\c1 c2. ((boundary c1) \ (boundary c2)) = {}) twoChain \ inj_on cubeImage twoChain" definition two_chain_boundary:: "two_chain \ one_chain" where "two_chain_boundary twoChain == \(boundary ` twoChain)" definition gen_division where "gen_division s S \ (finite S \ (\S = s) \ pairwise (\X Y. negligible (X \ Y)) S)" definition two_chain_horizontal_boundary:: "two_chain \ one_chain" where "two_chain_horizontal_boundary twoChain \ \(horizontal_boundary ` twoChain)" definition two_chain_vertical_boundary:: "two_chain \ one_chain" where "two_chain_vertical_boundary twoChain \ \(vertical_boundary ` twoChain)" definition only_horizontal_division where "only_horizontal_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" lemma sum_zero_set: assumes "\x \ s. f x = 0" "finite s" "finite t" shows "sum f (s \ t) = sum f t" using assms by (simp add: IntE sum.union_inter_neutral sup_commute) abbreviation "valid_typeII_division s twoChain \ ((\twoCube \ twoChain. typeII_twoCube twoCube) \ (gen_division s (cubeImage ` twoChain)) \ (valid_two_chain twoChain))" lemma two_chain_vertical_boundary_is_boundary_chain: shows "boundary_chain (two_chain_vertical_boundary twoChain)" by(simp add: boundary_chain_def two_chain_vertical_boundary_def vertical_boundary_def) lemma two_chain_horizontal_boundary_is_boundary_chain: shows "boundary_chain (two_chain_horizontal_boundary twoChain)" by(simp add: boundary_chain_def two_chain_horizontal_boundary_def horizontal_boundary_def) definition typeI_twoCube :: "two_cube \ bool" where "typeI_twoCube (twoC::two_cube) \ \a b g1 g2. a < b \ (\x \ {a..b}. g2 x \ g1 x) \ twoC = (\(x,y). ((1-x)*a + x*b, (1 - y) * (g2 ((1-x)*a + x*b)) + y * (g1 ((1-x)*a + x*b)))) \ g1 piecewise_C1_differentiable_on {a..b} \ g2 piecewise_C1_differentiable_on {a..b}" lemma typeI_twoCubeImg: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} " proof - have "\a b. a < b \ (\g1 g2. (\x\{a..b}. g2 x \ g1 x) \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))\ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b})" using assms by (simp add: typeI_twoCube_def) then obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x\{a..b}. g2 x \ g1 x)" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" by auto have ex: "\z\Green.unit_cube. (c, d) = (case z of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" if c_bounds: "a \ c" "c \ b" and d_bounds: "g2 c \ d" "d \ g1 c" for c d proof - have x_witness: "\k1. c = (1 - k1)*a + k1 * b \ 0 \ k1 \ k1 \ 1" proof - let ?k1 = "(c - a)/(b - a)" have k1_in_bounds: "0 \ (c - a)/(b - a) \ (c - a)/(b - a) \ 1" using twoCisTypeI(1) c_bounds by simp have "c = (1 - ?k1)*a + ?k1 * b" using twoCisTypeI(1) sum_sqs_eq by (auto simp add: divide_simps algebra_simps) then show ?thesis using twoCisTypeI k1_in_bounds by fastforce qed have y_witness: "\k2. d = (1 - k2)*(g2 c) + k2 * (g1 c) \ 0 \ k2 \ k2 \ 1" proof (cases "g1 c - g2 c = 0") case True with d_bounds show ?thesis by force next case False let ?k2 = "(d - g2 c)/(g1 c - g2 c)" have k2_in_bounds: "0 \ ?k2 \ ?k2 \ 1" using twoCisTypeI(2) c_bounds d_bounds False by simp have "d = (1 - ?k2) * g2 c + ?k2 * g1 c" using False apply (simp add: divide_simps algebra_simps) using sum_sqs_eq by fastforce then show ?thesis using k2_in_bounds by fastforce qed show "\x\unit_cube. (c, d) = (case x of (x, y) \ (a - x * a + x * b, g2 (a - x * a + x * b) - y * g2 (a - x * a + x * b) + y * g1 (a - x * a + x * b)))" using x_witness y_witness by (force simp add: left_diff_distrib) qed have "{y. \x\unit_cube. y = twoC x} = {(x, y). a \ x \ x \ b \ g2 x \ y \ y \ g1 x}" apply (auto simp add: twoCisTypeI left_diff_distrib ex) using less_eq_real_def twoCisTypeI(1) apply auto[1] apply (smt affine_ineq twoCisTypeI) apply (smt affine_ineq atLeastAtMost_iff mult_left_mono twoCisTypeI)+ done then show ?thesis unfolding cubeImage_def image_def using twoCisTypeI by auto qed lemma typeI_cube_explicit_spec: assumes "typeI_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(x,y). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b))) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x))) \ (\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) \ (\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x))) \ (\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeI_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" by (simp add: twoCisTypeI(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" by (simp add: twoCisTypeI(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" by (simp add: twoCisTypeI(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeI by auto qed lemma typeI_twoCube_smooth_edges: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using assms and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (a + (b - a) * x, g2 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeI(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeI(1)] twoCisTypeI(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x. (a + (b - a) * x, g1 (a + (b - a) * x))) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x. (a + (b - a) * x, g1 (a + (b - a) * x)))"] apply (simp only: real_pair_basis) by fastforce then show ?thesis using twoCisTypeI(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by auto then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce then have "(\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a))) piecewise_C1_differentiable_on {0..1}" using pair_prod_smooth_pw_smooth by auto then show ?thesis using twoCisTypeI(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma two_chain_integral_eq_integral_divisable: assumes f_integrable: "\twoCube \ twoChain. F integrable_on cubeImage twoCube" and gen_division: "gen_division s (cubeImage ` twoChain)" and valid_two_chain: "valid_two_chain twoChain" shows "integral s F = two_chain_integral twoChain F" proof - show "integral s F = two_chain_integral twoChain F" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "\twoCube \ twoChain. ((F) has_integral (integral (cubeImage twoCube) (F))) (cubeImage twoCube)" using f_integrable by auto then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` twoChain \ (F has_integral (integral (twoCubeImg) F)) (twoCubeImg)" using Henstock_Kurzweil_Integration.integrable_neg by force have finite_images: "finite (cubeImage ` twoChain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` twoChain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have inj: "inj_on cubeImage twoChain" using valid_two_chain by (simp add: inj_on_def valid_two_chain_def) have "integral s F = (\twoCubeImg\cubeImage ` twoChain. integral twoCubeImg F)" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) also have "\ = (\C\twoChain. integral (cubeImage C) F)" using sum.reindex inj by auto finally show "integral s F = (\C\twoChain. integral (cubeImage C) F)" . qed qed definition only_vertical_division where "only_vertical_division one_chain two_chain \ \\ \. finite \ \ finite \ \ (\(k,\) \ \. (\(k',\') \ two_chain_vertical_boundary two_chain. (\a \ {0..1}. \b \ {0..1}. a \ b \ subpath a b \' = \))) \ (common_sudiv_exists (two_chain_horizontal_boundary two_chain) \ \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)) \ boundary_chain \ \ one_chain = \ \ \ \ (\(k,\)\\. valid_path \)" abbreviation "valid_typeI_division s twoChain \ (\twoCube \ twoChain. typeI_twoCube twoCube) \ gen_division s (cubeImage ` twoChain) \ valid_two_chain twoChain" lemma field_cont_on_typeI_region_cont_on_edges: assumes typeI_twoC: "typeI_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using typeI_twoC and typeI_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x))) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeI(2) twoCisTypeI(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(x, 1)) = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(x, 1)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(2) twoCisTypeI(3) by auto have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(x, 0)) = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using twoCisTypeI by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(x, 0)) = ?Dg2" using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp path_image_def add_scale_img and twoCisTypeI(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (a, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 a \ g1 a" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 a + (g1 a - g2 a) * x) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by blast then have left_eq: "?D_left_edge = ?left_edge ` {0..1}" unfolding twoCisTypeI(10) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (b, y)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeI(1) twoCisTypeI(3) by auto have "g2 b \ g1 b" using twoCisTypeI(1) twoCisTypeI(2) by auto then have "(\x. g2 b + (g1 b - g2 b) * x) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by blast then have right_eq: "?D_right_edge = ?right_edge ` {0..1}" unfolding twoCisTypeI(8) by(auto simp add: subset_iff image_def set_eq_iff Semiring_Normalization.comm_semiring_1_class.semiring_normalization_rules(7)) then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma typeII_cube_explicit_spec: assumes "typeII_twoCube twoC" shows "\a b g1 g2. a < b \ (\x \ {a .. b}. g2 x \ g1 x) \ cubeImage twoC = {(y, x). x \ {a..b} \ y \ {g2 x .. g1 x}} \ twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b)) \ g1 piecewise_C1_differentiable_on {a .. b} \ g2 piecewise_C1_differentiable_on {a .. b} \ (\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) \ (\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x)) \ (\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using assms and typeII_twoCubeImg[of"twoC"] by auto have bottom_edge_explicit: "?bottom_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have right_edge_explicit: "?right_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" by (simp add: twoCisTypeII(4) algebra_simps) have top_edge_explicit: "?top_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_explicit: "?left_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show ?thesis using bottom_edge_explicit right_edge_explicit top_edge_explicit left_edge_explicit twoCisTypeII by auto qed lemma typeII_twoCube_smooth_edges: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "\ piecewise_C1_differentiable_on {0..1}" proof - let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using assms and typeII_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "?bottom_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x)) -` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(7) by auto qed have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by auto then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using scale_shift_smooth C1_differentiable_imp_piecewise by blast have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using twoCisTypeII(9) by auto qed have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b))) piecewise_C1_differentiable_on {0..1}" by (simp add: C1_differentiable_imp_piecewise) then have "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[of "(1,0)" "(\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))"] by (auto simp add: real_pair_basis) then show ?thesis using twoCisTypeII(8) by auto qed have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have 0: "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a))) C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by fastforce have "(\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise[OF C1_differentiable_on_pair[OF 0 C1_differentiable_on_const[of "a" "{0..1}"]]] by force then show ?thesis using twoCisTypeII(10) by auto qed have "\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge" using assms by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show ?thesis using left_edge_smooth right_edge_smooth top_edge_smooth bottom_edge_smooth by auto qed lemma field_cont_on_typeII_region_cont_on_edges: assumes typeII_twoC: "typeII_twoCube twoC" and field_cont: "continuous_on (cubeImage twoC) F" and member_of_boundary: "(k,\) \ boundary twoC" shows "continuous_on (\ ` {0 .. 1}) F" proof - obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(0, x)) = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 1)) = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" "(\x. twoC(1, x)) = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" "(\y. twoC(y, 0)) = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" using typeII_twoC and typeII_cube_explicit_spec[of"twoC"] by auto let ?bottom_edge = "(\x. twoC(0, x))" let ?right_edge = "(\y. twoC(y, 1))" let ?top_edge = "(\x. twoC(1, x))" let ?left_edge = "(\y. twoC(y, 0))" let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_top_edge_image: "continuous_on ?Dg1 F" by (rule continuous_on_subset [OF field_cont]) (auto simp: twoCisTypeII(2) twoCisTypeII(3)) have top_edge_is_compos_of_scal_and_g1: "(\x. twoC(1, x)) = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg1_is_bot_edge_pathimg: "path_image (\x. twoC(1, x)) = ?Dg1" using line_is_pair_img and top_edge_is_compos_of_scal_and_g1 image_comp path_image_def add_scale_img and twoCisTypeII(1) by (metis (no_types, lifting) cbox_interval) then have cont_on_top: "continuous_on (path_image ?top_edge) F" using field_cont_on_top_edge_image by auto let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x)) ` (cbox a b)" using image_def by auto have field_cont_on_bot_edge_image: "continuous_on ?Dg2 F" by (rule continuous_on_subset [OF field_cont]) (auto simp add: twoCisTypeII(2) twoCisTypeII(3)) have bot_edge_is_compos_of_scal_and_g2: "(\x. twoC(0, x)) = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using twoCisTypeII by auto have Dg2_is_bot_edge_pathimg: "path_image (\x. twoC(0, x)) = ?Dg2" unfolding path_image_def using line_is_pair_img and bot_edge_is_compos_of_scal_and_g2 image_comp add_scale_img [OF \a < b\] by (metis (no_types, lifting) box_real(2)) then have cont_on_bot: "continuous_on (path_image ?bottom_edge) F" using field_cont_on_bot_edge_image by auto let ?D_left_edge = "{p. \y. y \ cbox (g2 a) (g1 a) \ p = (y, a)}" have field_cont_on_left_edge_image: "continuous_on ?D_left_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 a \ g1 a" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 a + x * (g1 a - g2 a)) ` {(0::real)..1} = {g2 a .. g1 a}" using add_scale_img'[of "g2 a" "g1 a"] by (auto simp add: ac_simps) with \g2 a \ g1 a\ have left_eq: "?D_left_edge = ?left_edge ` {0..1}" by (simp only: twoCisTypeII(10)) auto then have cont_on_left: "continuous_on (path_image ?left_edge) F" using field_cont_on_left_edge_image path_image_def by (metis left_eq field_cont_on_left_edge_image path_image_def) let ?D_right_edge = "{p. \y. y \ cbox (g2 b) (g1 b) \ p = (y, b)}" have field_cont_on_left_edge_image: "continuous_on ?D_right_edge F" apply (rule continuous_on_subset [OF field_cont]) using twoCisTypeII(1) twoCisTypeII(3) by auto have "g2 b \ g1 b" using twoCisTypeII(1) twoCisTypeII(2) by auto then have "(\x. g2 b + x * (g1 b - g2 b)) ` {(0::real)..1} = {g2 b .. g1 b}" using add_scale_img'[of "g2 b" "g1 b"] by (auto simp add: ac_simps) with \g2 b \ g1 b\ have right_eq: "?D_right_edge = ?right_edge ` {0..1}" by (simp only: twoCisTypeII(8)) auto then have cont_on_right: "continuous_on (path_image ?right_edge) F" using field_cont_on_left_edge_image path_image_def by (metis right_eq field_cont_on_left_edge_image path_image_def) have all_edge_cases: "(\ = ?bottom_edge \ \ = ?right_edge \ \ = ?top_edge \ \ = ?left_edge)" using assms unfolding boundary_def horizontal_boundary_def vertical_boundary_def by blast show ?thesis apply (simp add: path_image_def[symmetric]) using cont_on_top cont_on_bot cont_on_right cont_on_left all_edge_cases by blast qed lemma two_cube_boundary_is_boundary: "boundary_chain (boundary C)" by (auto simp add: boundary_chain_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma common_boundary_subdiv_exists_refl: assumes "\(k,\)\boundary twoC. valid_path \" shows "common_boundary_sudivision_exists (boundary twoC) (boundary twoC)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def two_cube_boundary_is_boundary by blast lemma common_boundary_subdiv_exists_refl': assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_boundary_sudivision_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def by blast lemma gen_common_boundary_subdiv_exists_refl_twochain_boundary: assumes "\(k,\)\C. valid_path \" "boundary_chain (C::(int \ (real \ real \ real)) set)" shows "common_sudiv_exists (C) (C)" using assms chain_subdiv_chain_refl common_boundary_sudivision_exists_def common_subdiv_imp_gen_common_subdiv by blast lemma two_chain_boundary_is_boundary_chain: shows "boundary_chain (two_chain_boundary twoChain)" by (simp add: boundary_chain_def two_chain_boundary_def boundary_def horizontal_boundary_def vertical_boundary_def) lemma typeI_edges_are_valid_paths: assumes "typeI_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeI_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma typeII_edges_are_valid_paths: assumes "typeII_twoCube twoC" "(k,\) \ boundary twoC" shows "valid_path \" using typeII_twoCube_smooth_edges[OF assms] C1_differentiable_imp_piecewise by (auto simp: valid_path_def) lemma finite_two_chain_vertical_boundary: assumes "finite two_chain" shows "finite (two_chain_vertical_boundary two_chain)" using assms by (simp add: two_chain_vertical_boundary_def vertical_boundary_def) lemma finite_two_chain_horizontal_boundary: assumes "finite two_chain" shows "finite (two_chain_horizontal_boundary two_chain)" using assms by (simp add: two_chain_horizontal_boundary_def horizontal_boundary_def) locale R2 = fixes i j assumes i_is_x_axis: "i = (1::real,0::real)" and j_is_y_axis: "j = (0::real, 1::real)" begin lemma analytically_valid_y: assumes "analytically_valid s F i" shows "(\x. integral UNIV (\y. (partial_vector_derivative F i) (y, x) * (indicator s (y, x)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(1, 0)} = {(0::real,1::real)}" by force with assms show ?thesis using assms by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma analytically_valid_x: assumes "analytically_valid s F j" shows "(\x. integral UNIV (\y. ((partial_vector_derivative F j) (x, y)) * (indicator s (x, y)))) \ borel_measurable lborel" proof - have "{(1::real, 0::real), (0, 1)} - {(0, 1)} = {(1::real, 0::real)}" by force with assms show ?thesis by(simp add: real_pair_basis analytically_valid_def i_is_x_axis j_is_y_axis) qed lemma Greens_thm_type_I: fixes F:: "((real*real) \ (real * real))" and gamma1 gamma2 gamma3 gamma4 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dy_def: "Dy_pair = {(x::real,y) . x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" and gamma1_def: "gamma1 = (\x. (a + (b - a) * x, g2(a + (b - a) * x)))" and gamma1_smooth: "gamma1 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma2_def: "gamma2 = (\x. (b, g2(b) + x *\<^sub>R (g1(b) - g2(b))))" and gamma3_def: "gamma3 = (\x. (a + (b - a) * x, g1(a + (b - a) * x)))" and gamma3_smooth: "gamma3 piecewise_C1_differentiable_on {0..1}" and gamma4_def: "gamma4 = (\x. (a, g2(a) + x *\<^sub>R (g1(a) - g2(a))))" and F_i_analytically_valid: "analytically_valid Dy_pair (\p. F(p) \ i) j" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (partial_vector_derivative (\p. F(p) \ i) j a)))" "line_integral_exists F {i} gamma4" "line_integral_exists F {i} gamma3" "line_integral_exists F {i} gamma2" "line_integral_exists F {i} gamma1" proof - let ?F_b' = "partial_vector_derivative (\a. (F a) \ i) j" have F_first_is_continuous: "continuous_on Dy_pair (\a. F(a) \ i)" using F_i_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dy_pair) then (partial_vector_derivative (\a. (F a) \ i) j) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_i_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dy_pair. partially_vector_differentiable (\a. (F a) \ i) j a" using F_i_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have x_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(x, y))) \ borel_measurable lborel" proof - have "(\p. (?F_b' p) * indicator (Dy_pair) p) = (\x. if x \ (Dy_pair) then (?F_b') x else 0)" using indic_ident[of "?F_b'"] by auto then have "\x y. ?F_b' (x,y) * indicator (Dy_pair) (x,y) = (\x. if x \ (Dy_pair) then (?F_b') x else 0) (x,y)" by auto then show ?thesis using analytically_valid_x[OF F_i_analytically_valid] by (auto simp add: indicator_def) qed have F_partially_differentiable: "\a\Dy_pair. has_partial_vector_derivative (\x. (F x) \ i) j (?F_b' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" using a_lt_b apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using le_diff_eq by fastforce+ have gamma1_y_component: "(\x. g2(a + (b - a) * x)) = g2 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma1_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma1_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g2 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma1_y_component by auto then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have gamma3_y_component: "(\x. g1(a + (b - a) * x)) = g1 \ (\x.(a + (b - a) * x))" by auto have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma3_smooth], of "(\x. j)", OF continuous_on_const] by (simp add: gamma3_def j_is_y_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then have "continuous_on {a..b} (g1 \ (\x.(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using gamma3_y_component by auto then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g2 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (x, g2(x))}" have line_is_pair_img: "?Dg2 = (\x. (x, g2(x))) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (x, g2(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma1_image: "continuous_on ?Dg2 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma1_is_compos_of_scal_and_g2: "gamma1 = (\x. (x, g2(x))) \ (\x. a + (b - a) * x)" using gamma1_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma1_pathimg: "path_image gamma1 = ?Dg2" by (metis (no_types, lifting) box_real(2) gamma1_is_compos_of_scal_and_g2 image_comp line_is_pair_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma1_as_euclid_space_fun: "gamma1 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g2 (a + (b - a) * x)))" using i_is_x_axis gamma1_def by auto have 0: "line_integral F {i} gamma1 = integral (cbox a b) (\x. F(x, g2(x)) \ i )" "line_integral_exists F {i} gamma1" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma1_as_euclid_space_fun, of "F"] gamma1_def gamma1_smooth g2_scale_j_contin a_lt_b add_scale_img Dg2_is_gamma1_pathimg and field_cont_on_gamma1_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma1)" by metis have gamma2_x_const: "\x. gamma2 x \ i = b" by (simp add: i_is_x_axis gamma2_def) have 1: "(line_integral F {i} gamma2) = 0" "(line_integral_exists F {i} gamma2)" using line_integral_on_pair_straight_path[OF gamma2_x_const] straight_path_diffrentiable_x gamma2_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma2)" by metis have "continuous_on (cbox a b) (\x. F(x, g2(x)) \ i)" using line_is_pair_img and g2_path_continuous and field_cont_on_gamma1_image Topological_Spaces.continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 6: "(\x. F(x, g2(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g2(x)) \ i)"] by auto have g1_scale_j_contin: "continuous_on (cbox a b) (\x. (0, g1 x))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (x, g1(x))}" have line_is_pair_img: "?Dg1 = (\x. (x, g1(x)) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (x, g1(x)))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma3_image: "continuous_on ?Dg1 (\a. F(a) \ i)" apply (rule continuous_on_subset [OF F_first_is_continuous]) by (auto simp add: Dy_def g2_leq_g1) have gamma3_is_compos_of_scal_and_g1: "gamma3 = (\x. (x, g1(x))) \ (\x. a + (b - a) * x)" using gamma3_def by auto then have Dg1_is_gamma3_pathimg: "path_image gamma3 = ?Dg1" by (metis (no_types, lifting) box_real(2) image_comp line_is_pair_img local.add_scale_img path_image_def) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma3_as_euclid_space_fun: "gamma3 = (\x. (a + (b - a) * x) *\<^sub>R i + (0, g1 (a + (b - a) * x)))" using i_is_x_axis gamma3_def by auto have 2: "line_integral F {i} gamma3 = integral (cbox a b) (\x. F(x, g1(x)) \ i )" "line_integral_exists F {i} gamma3" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(1)] _ gamma3_as_euclid_space_fun, of "F"] gamma3_def and gamma3_smooth and g1_scale_j_contin and a_lt_b and add_scale_img Dg1_is_gamma3_pathimg and field_cont_on_gamma3_image by (auto simp: pathstart_def pathfinish_def i_is_x_axis) then show "(line_integral_exists F {i} gamma3)" by metis have gamma4_x_const: "\x. gamma4 x \ i = a" using gamma4_def by (auto simp add: real_inner_class.inner_add_left inner_not_same_Basis i_is_x_axis) have 3: "(line_integral F {i} gamma4) = 0" "(line_integral_exists F {i} gamma4)" using line_integral_on_pair_straight_path[OF gamma4_x_const] straight_path_diffrentiable_x gamma4_def by (auto simp add: mult.commute) then show "(line_integral_exists F {i} gamma4)" by metis have "continuous_on (cbox a b) (\x. F(x, g1(x)) \ i)" using line_is_pair_img and g1_path_continuous and field_cont_on_gamma3_image continuous_on_compose i_is_x_axis j_is_y_axis by auto then have 7: "(\x. F(x, g1(x)) \ i) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(x, g1(x)) \ i)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_b'(xc, y)) has_integral F(xc,g1(xc)) \ i - F(xc, g2(xc)) \ i) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc proof - have "{(xc', y). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dy_pair" using that by (auto simp add: Dy_def) then show "((\y. ?F_b' (xc, y)) has_integral F(xc, g1 xc) \ i - F(xc, g2 xc) \ i) (cbox (g2 xc) (g1 xc))" using that and Base_vecs and F_partially_differentiable and Dy_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "j" "i" "xc *\<^sub>R i" "Dy_pair" "F" "?F_b'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have partial_deriv_integrable: "(?F_b') integrable_on Dy_pair" by (simp add: F_i_analytically_valid analytically_valid_imp_part_deriv_integrable_on) have 4: "integral Dy_pair ?F_b' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_b' (x, y)))" proof - have x_axis_gauge_integrable: "\x. (\y. ?f(x,y)) integrable_on UNIV" proof - fix x::real have "\x. x \ cbox a b \ (\y. ?f (x, y)) = (\y. 0)" by (auto simp add: Dy_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" by (simp add: integrable_0) let ?F_b'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_b' (x,y) else 0))" have f_value_x_in_range: "\x \ cbox a b. ?F_b'_oneD x = (\y. ?f(x,y))" by (auto simp add: Dy_def) have "\x \ cbox a b. ?F_b'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast then have f_integrable_x_in_range: "\x. x \ cbox a b \ (\y. ?f (x, y)) integrable_on UNIV" using f_value_x_in_range by auto show "(\y. ?f (x, y)) integrable_on UNIV" using f_integrable_x_not_in_range and f_integrable_x_in_range by auto qed have arg: "(\a. if a \ Dy_pair then partial_vector_derivative (\a. F a \ i) j a else 0) = (\x. if x \ Dy_pair then if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0 else 0)" by auto have arg2: "Dy_pair = {(x, y). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dy_def by auto have arg3: "\ x. x \ Dy_pair \ (\x. if x \ Dy_pair then partial_vector_derivative (\a. F a \ i) j x else 0) x = (\x. partial_vector_derivative (\a. F a \ i) j x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (x, y) \ Dy_pair then partial_vector_derivative (\a. F a \ i) j (x, y) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ i) j (x, y))) x" apply (simp add: Dy_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_x [OF f_lesbegue_integrable x_axis_gauge_integrable x_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dy_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have 5: "(integral Dy_pair (\a. (?F_b' a)) = integral (cbox a b) (\x. F(x, g1(x)) \ i - F(x, g2(x)) \ i))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) show "(line_integral F {i} gamma1) + (line_integral F {i} gamma2) - (line_integral F {i} gamma3) - (line_integral F {i} gamma4) = (integral Dy_pair (\a. - (?F_b' a)))" using "0"(1) "1"(1) "2"(1) "3"(1) 5 "6" "7" by (simp add: Henstock_Kurzweil_Integration.integral_diff) qed theorem Greens_thm_type_II: fixes F:: "((real*real) \ (real * real))" and gamma4 gamma3 gamma2 gamma1 :: "(real \ (real * real))" and a:: "real" and b:: "real" and g1:: "(real \ real)" and g2:: "(real \ real)" assumes Dx_def: "Dx_pair = {(x::real,y) . y \ cbox a b \ x \ cbox (g2 y) (g1 y)}" and gamma4_def: "gamma4 = (\x. (g2(a + (b - a) * x), a + (b - a) * x))" and gamma4_smooth: "gamma4 piecewise_C1_differentiable_on {0..1}" and (*TODO: This should be piecewise smooth*) gamma3_def: "gamma3 = (\x. (g2(b) + x *\<^sub>R (g1(b) - g2(b)), b))" and gamma2_def: "gamma2 = (\x. (g1(a + (b - a) * x), a + (b - a) * x))" and gamma2_smooth: "gamma2 piecewise_C1_differentiable_on {0..1}" and gamma1_def: "gamma1 = (\x. (g2(a) + x *\<^sub>R (g1(a) - g2(a)), a))" and F_j_analytically_valid: "analytically_valid Dx_pair (\p. F(p) \ j) i" and g2_leq_g1: "\x \ cbox a b. (g2 x) \ (g1 x)" and (*This is needed otherwise what would Dy be?*) a_lt_b: "a < b" shows "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (partial_vector_derivative (\a. (F a) \ j) i a)))" "line_integral_exists F {j} gamma4" "line_integral_exists F {j} gamma3" "line_integral_exists F {j} gamma2" "line_integral_exists F {j} gamma1" proof - let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have F_first_is_continuous: "continuous_on Dx_pair (\a. F(a) \ j)" using F_j_analytically_valid by (auto simp add: analytically_valid_def) let ?f = "(\x. if x \ (Dx_pair) then (partial_vector_derivative (\a. (F a) \ j) i) x else 0)" have f_lesbegue_integrable: "integrable lborel ?f" using F_j_analytically_valid by (auto simp add: analytically_valid_def indic_ident) have partially_vec_diff: "\a \ Dx_pair. partially_vector_differentiable (\a. (F a) \ j) i a" using F_j_analytically_valid by (auto simp add: analytically_valid_def indicator_def) have "\x y. ?F_a' (x,y) * indicator (Dx_pair) (x,y) = (\x. if x \ (Dx_pair) then ?F_a' x else 0) (x,y)" using indic_ident[of "?F_a'"] by auto then have y_axis_integral_measurable: "(\x. integral UNIV (\y. ?f(y, x))) \ borel_measurable lborel" using analytically_valid_y[OF F_j_analytically_valid] by (auto simp add: indicator_def) have F_partially_differentiable: "\a\Dx_pair. has_partial_vector_derivative (\x. (F x) \ j) i (?F_a' a) a" using partial_vector_derivative_works partially_vec_diff by fastforce have g1_g2_continuous: "continuous_on (cbox a b) g1" "continuous_on (cbox a b) g2" proof - have shift_scale_cont: "continuous_on {a..b} (\x. (x - a)*(1/(b-a)))" by (intro continuous_intros) have shift_scale_inv: "(\x. a + (b - a) * x) \ (\x. (x - a)*(1/(b-a))) = id" using a_lt_b by (auto simp add: o_def) have img_shift_scale: "(\x. (x - a)*(1/(b-a)))`{a..b} = {0..1}" apply (auto simp: divide_simps image_iff) apply (rule_tac x="x * (b - a) + a" in bexI) using a_lt_b by (auto simp: algebra_simps mult_le_cancel_left affine_ineq) have "continuous_on {0..1} (\x. g2(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma4_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma4_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g2(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g2" using a_lt_b by (simp add: shift_scale_inv) have "continuous_on {0..1} (\x. g1(a + (b - a) * x))" using continuous_on_inner[OF piecewise_C1_differentiable_on_imp_continuous_on[OF gamma2_smooth], of "(\x. i)", OF continuous_on_const] by (simp add: gamma2_def i_is_x_axis) then have "continuous_on {a..b} ((\x. g1(a + (b - a) * x)) \ (\x. (x - a)*(1/(b-a))))" using img_shift_scale continuous_on_compose shift_scale_cont by force then show "continuous_on (cbox a b) g1" using a_lt_b by (simp add: shift_scale_inv) qed have g2_scale_i_contin: "continuous_on (cbox a b) (\x. (g2 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg2 = "{p. \x. x \ cbox a b \ p = (g2(x), x)}" have line_is_pair_img: "?Dg2 = (\x. (g2(x), x) ) ` (cbox a b)" using image_def by auto have g2_path_continuous: "continuous_on (cbox a b) (\x. (g2(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma4_image: "continuous_on ?Dg2 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have gamma4_is_compos_of_scal_and_g2: "gamma4 = (\x. (g2(x), x)) \ (\x. a + (b - a) * x)" using gamma4_def by auto have add_scale_img: "(\x. a + (b - a) * x) ` {0 .. 1} = {a .. b}" using add_scale_img and a_lt_b by auto then have Dg2_is_gamma4_pathimg: "path_image gamma4 = ?Dg2" using line_is_pair_img and gamma4_is_compos_of_scal_and_g2 image_comp path_image_def by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma4_as_euclid_space_fun: "gamma4 = (\x. (a + (b - a) * x) *\<^sub>R j + (g2 (a + (b - a) * x), 0))" using j_is_y_axis gamma4_def by auto have 0: "(line_integral F {j} gamma4) = integral (cbox a b) (\x. F(g2(x), x) \ j)" "line_integral_exists F {j} gamma4" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma4_as_euclid_space_fun] gamma4_def gamma4_smooth g2_scale_i_contin a_lt_b add_scale_img Dg2_is_gamma4_pathimg and field_cont_on_gamma4_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma4" by metis have gamma3_y_const: "\x. gamma3 x \ j = b" by (simp add: gamma3_def j_is_y_axis) have 1: "(line_integral F {j} gamma3) = 0" "(line_integral_exists F {j} gamma3)" using line_integral_on_pair_straight_path[OF gamma3_y_const] straight_path_diffrentiable_y gamma3_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma3" by auto have "continuous_on (cbox a b) (\x. F(g2(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma4_image g2_path_continuous line_is_pair_img) then have 6: "(\x. F(g2(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous by blast have g1_scale_i_contin: "continuous_on (cbox a b) (\x. (g1 x, 0))" by (intro continuous_intros g1_g2_continuous) let ?Dg1 = "{p. \x. x \ cbox a b \ p = (g1(x), x)}" have line_is_pair_img: "?Dg1 = (\x. (g1(x), x) ) ` (cbox a b)" using image_def by auto have g1_path_continuous: "continuous_on (cbox a b) (\x. (g1(x), x))" by (intro continuous_intros g1_g2_continuous) have field_cont_on_gamma2_image: "continuous_on ?Dg1 (\a. F(a) \ j)" by (rule continuous_on_subset [OF F_first_is_continuous]) (auto simp: Dx_def g2_leq_g1) have "gamma2 = (\x. (g1(x), x)) \ (\x. a + (b - a) * x)" using gamma2_def by auto then have Dg1_is_gamma2_pathimg: "path_image gamma2 = ?Dg1" using line_is_pair_img image_comp path_image_def add_scale_img by (metis (no_types, lifting) cbox_interval) have Base_vecs: "i \ Basis" "j \ Basis" "i \ j" using real_pair_basis and i_is_x_axis and j_is_y_axis by auto have gamma2_as_euclid_space_fun: "gamma2 = (\x. (a + (b - a) * x) *\<^sub>R j + (g1 (a + (b - a) * x), 0))" using j_is_y_axis gamma2_def by auto have 2: "(line_integral F {j} gamma2) = integral (cbox a b) (\x. F(g1(x), x) \ j)" "(line_integral_exists F {j} gamma2)" using line_integral_on_pair_path_strong [OF norm_Basis[OF Base_vecs(2)] _ gamma2_as_euclid_space_fun] gamma2_def and gamma2_smooth and g1_scale_i_contin and a_lt_b and add_scale_img Dg1_is_gamma2_pathimg and field_cont_on_gamma2_image by (auto simp: pathstart_def pathfinish_def j_is_y_axis) then show "line_integral_exists F {j} gamma2" by metis have gamma1_y_const: "\x. gamma1 x \ j = a" using gamma1_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have 3: "(line_integral F {j} gamma1) = 0" "(line_integral_exists F {j} gamma1)" using line_integral_on_pair_straight_path[OF gamma1_y_const] straight_path_diffrentiable_y gamma1_def by (auto simp add: mult.commute) then show "line_integral_exists F {j} gamma1" by auto have "continuous_on (cbox a b) (\x. F(g1(x), x) \ j)" by (smt Collect_mono_iff continuous_on_compose2 continuous_on_eq field_cont_on_gamma2_image g1_path_continuous line_is_pair_img) then have 7: "(\x. F(g1(x), x) \ j) integrable_on (cbox a b)" using integrable_continuous [of "a" "b" "(\x. F(g1(x), x) \ j)"] by auto have partial_deriv_one_d_integrable: "((\y. ?F_a'(y, xc)) has_integral F(g1(xc), xc) \ j - F(g2(xc), xc) \ j) (cbox (g2 xc) (g1 xc))" if "xc \ cbox a b" for xc::real proof - have "{(y, xc'). y \ cbox (g2 xc) (g1 xc) \ xc' = xc} \ Dx_pair" using that by (auto simp add: Dx_def) then show ?thesis using that and Base_vecs and F_partially_differentiable and Dx_def [symmetric] and g2_leq_g1 and fundamental_theorem_of_calculus_partial_vector [of "g2 xc" "g1 xc" "i" "j" "xc *\<^sub>R j" "Dx_pair" "F" "?F_a'" ] by (auto simp add: Groups.ab_semigroup_add_class.add.commute i_is_x_axis j_is_y_axis) qed have "?f integrable_on UNIV" by (simp add: f_lesbegue_integrable integrable_on_lborel) then have partial_deriv_integrable: "?F_a' integrable_on Dx_pair" using integrable_restrict_UNIV by auto have 4: "integral Dx_pair ?F_a' = integral (cbox a b) (\x. integral (cbox (g2 x) (g1 x)) (\y. ?F_a' (y, x)))" proof - have y_axis_gauge_integrable: "(\y. ?f(y, x)) integrable_on UNIV" for x proof - let ?F_a'_oneD = "(\x. (\y. if y \ (cbox (g2 x) ( g1 x)) then ?F_a' (y, x) else 0))" have "\x. x \ cbox a b \ (\y. ?f (y, x)) = (\y. 0)" by (auto simp add: Dx_def) then have f_integrable_x_not_in_range: "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by (simp add: integrable_0) have "\x \ cbox a b. ?F_a'_oneD x = (\y. ?f(y, x))" using g2_leq_g1 by (auto simp add: Dx_def) moreover have "\x \ cbox a b. ?F_a'_oneD x integrable_on UNIV" using has_integral_integrable integrable_restrict_UNIV partial_deriv_one_d_integrable by blast ultimately have "\x. x \ cbox a b \ (\y. ?f (y, x)) integrable_on UNIV" by auto then show "(\y. ?f (y, x)) integrable_on UNIV" using f_integrable_x_not_in_range by auto qed have arg: "(\a. if a \ Dx_pair then partial_vector_derivative (\a. F a \ j) i a else 0) = (\x. if x \ Dx_pair then if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0 else 0)" by auto have arg2: "Dx_pair = {(y, x). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. g2 x \ i \ y \ i \ y \ i \ g1 x \ i)}" using Dx_def by auto have arg3: "\ x. x \ Dx_pair \ (\x. if x \ Dx_pair then partial_vector_derivative (\a. F a \ j) i x else 0) x = (\x. partial_vector_derivative (\a. F a \ j) i x) x" by auto have arg4: "\x. x \ (cbox a b) \ (\x. integral (cbox (g2 x) (g1 x)) (\y. if (y, x) \ Dx_pair then partial_vector_derivative (\a. F a \ j) i (y, x) else 0)) x = (\x. integral (cbox (g2 x) (g1 x)) (\y. partial_vector_derivative (\a. F a \ j) i (y, x))) x" apply (clarsimp simp: Dx_def) by (smt Henstock_Kurzweil_Integration.integral_cong atLeastAtMost_iff) show ?thesis using gauge_integral_Fubini_curve_bounded_region_y [OF f_lesbegue_integrable y_axis_gauge_integrable y_axis_integral_measurable arg arg2] Henstock_Kurzweil_Integration.integral_cong[OF arg3, of "Dx_pair" "(\x. x)"] Henstock_Kurzweil_Integration.integral_cong[OF arg4, of "cbox a b" "(\x. x)"] by auto qed have "((integral Dx_pair (\a. (?F_a' a))) = integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j))" using 4 Henstock_Kurzweil_Integration.integral_cong partial_deriv_one_d_integrable integrable_def by (smt integral_unique) then have "integral Dx_pair (\a. - (?F_a' a)) = - integral (cbox a b) (\x. F(g1(x), x) \ j - F(g2(x), x) \ j)" using partial_deriv_integrable and integral_neg by auto then have 5: "integral Dx_pair (\a. - (?F_a' a)) = integral (cbox a b) (\x. ( F(g2(x), x) \ j - F(g1(x), x) \ j))" using 6 7 and integral_neg [of _ "(\x. F(g1 x, x) \ j - F(g2 x, x) \ j)"] by auto have "(line_integral F {j} gamma4) + (line_integral F {j} gamma3) - (line_integral F {j} gamma2) - (line_integral F {j} gamma1) = (integral Dx_pair (\a. -(?F_a' a)))" using 0 1 2 3 5 6 7 Henstock_Kurzweil_Integration.integral_diff[of "(\x. F(g2(x), x) \ j)" "cbox a b" "(\x. F(g1(x), x) \ j)"] by (auto) then show "-(line_integral F {j} gamma4) - (line_integral F {j} gamma3) + (line_integral F {j} gamma2) + (line_integral F {j} gamma1) = (integral Dx_pair (\a. (?F_a' a)))" by (simp add: \integral Dx_pair (\a. - ?F_a' a) = - integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\ \integral Dx_pair ?F_a' = integral (cbox a b) (\x. F(g1 x, x) \ j - F(g2 x, x) \ j)\) qed end locale green_typeII_cube = R2 + fixes twoC F assumes two_cube: "typeII_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma GreenThm_typeII_twoCube: shows "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {j} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {j} (boundary twoC) = line_integral F {j} ?bottom_edge + line_integral F {j} ?right_edge - line_integral F {j} ?top_edge - line_integral F {j} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto then have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite1] using valid_two_cube apply (simp only: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto then have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {j} g) = (- line_integral F {j} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto then have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {j} g) = ( line_integral F {j} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {j} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {j} g) = line_integral F {j} (\x. twoC (x, 0)) + line_integral F {j} (\y. twoC (1, y)) - line_integral F {j} (\x. twoC (x, 1)) - line_integral F {j} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeII: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(y, x). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(y, x). ((1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b), (1 - x) * a + x * b))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" using two_cube and typeII_twoCubeImg[of"twoC"] by auto have left_edge_explicit: "?left_edge = (\x. (g2 (a + (b - a) * x), a + (b - a) * x))" by (simp add: twoCisTypeII(4) algebra_simps) have left_edge_smooth: "?left_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto have scale_shif_smth: "(\x. (a + (b - a) * x)) C1_differentiable_on {0..1}" using scale_shift_smooth by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise by blast have g2_smooth: "g2 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(6) by auto have "(\x. g2 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g2_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g2 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using left_edge_explicit by auto qed have top_edge_explicit: "?top_edge = (\x. (g2 b + x *\<^sub>R (g1 b - g2 b), b))" and right_edge_explicit: "?right_edge = (\x. (g1 (a + (b - a) * x), a + (b - a) * x))" by (simp_all add: twoCisTypeII(4) algebra_simps) have right_edge_smooth: "?right_edge piecewise_C1_differentiable_on {0..1}" proof - have "\x. (\x. (a + (b - a) * x))-` {x} = {(x - a)/(b -a)}" using twoCisTypeII(1) by(auto simp add: Set.vimage_def) then have finite_vimg: "\x. finite({0..1} \ (\x. (a + (b - a) * x))-` {x})" by auto then have scale_shif_pw_smth: "(\x. (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using C1_differentiable_imp_piecewise [OF scale_shift_smooth] by auto have g1_smooth: "g1 piecewise_C1_differentiable_on (\x. a + (b - a) * x) ` {0..1}" using add_scale_img[OF twoCisTypeII(1)] twoCisTypeII(5) by auto have "(\x. g1 (a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using piecewise_C1_differentiable_compose[OF scale_shif_pw_smth g1_smooth finite_vimg] by (auto simp add: o_def) then have "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x)) piecewise_C1_differentiable_on {0..1}" using all_components_smooth_one_pw_smooth_is_pw_smooth[where f = "(\x::real. (g1 (a + (b - a) * x), a + (b - a) * x))"] by (fastforce simp add: real_pair_basis) then show ?thesis using right_edge_explicit by auto qed have bottom_edge_explicit: "?bottom_edge = (\x. (g2 a + x *\<^sub>R (g1 a - g2 a), a))" by (simp add: twoCisTypeII(4) algebra_simps) show "integral (cubeImage twoC) (\a. partial_vector_derivative (\x. (F x) \ j) i a) = one_chain_line_integral F {j} (boundary twoC)" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto have "line_integral_exists F {j} \" if "(k,\) \ boundary twoC" for k \ proof - have "line_integral_exists F {j} (\y. twoC (0, y))" "line_integral_exists F {j} (\x. twoC (x, 1))" "line_integral_exists F {j} (\y. twoC (1, y))" "line_integral_exists F {j} (\x. twoC (x, 0))" using Greens_thm_type_II[OF twoCisTypeII(3) left_edge_explicit left_edge_smooth top_edge_explicit right_edge_explicit right_edge_smooth bottom_edge_explicit f_analytically_valid twoCisTypeII(2) twoCisTypeII(1)] line_integral_around_boundary by auto then show "line_integral_exists F {j} \" using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then show "\(k,\) \ boundary twoC. line_integral_exists F {j} \" by auto qed lemma line_integral_exists_on_typeII_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {j} \" using assms GreenThm_typeII_twoCube(2) by blast end locale green_typeII_chain = R2 + fixes F two_chain s assumes valid_typeII_div: "valid_typeII_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ j) i" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeII_div by (auto simp add: valid_two_chain_def) lemma typeII_chain_line_integral_exists_boundary': shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] F_anal_valid valid_typeII_div apply(auto simp add: two_chain_boundary_def) using R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by blast then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeII_chain_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries'[of i j] valid_typeII_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {j} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeII_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using valid_typeII_div typeII_chain_line_integral_exists_boundary' typeII_chain_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) using boundary_def by auto lemma type_II_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using valid_typeII_div typeII_edges_are_valid_paths by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_II_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_horiz_div_line_integrable': assumes "only_horizontal_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" "\two_cube \ two_chain. valid_two_cube two_cube" shows "line_integral_exists F {j} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by auto have integ_exis_horiz: "(\k \. (\(k', \') \ two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis type_II_chain_horiz_bound_valid using line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "((common_sudiv_exists (two_chain_vertical_boundary two_chain) \) \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_horizontal_division_def by blast have finite_j: "finite {j}" by auto show "line_integral_exists F {j} \" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) integ_exis_vert finite_two_chain_vertical_boundary[OF assms(4)] hv_props(5) finite_j] integ_exis_horiz[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using F_anal_valid field_cont_on_typeII_region_cont_on_edges valid_typeII_div by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "line_integral_exists F {j} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_j hv_props(5) finite_two_chain_vertical_boundary[OF assms(4)] hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] integ_exis_horiz[of "\"] assms(3) hv_props False by fastforce qed qed lemma GreenThm_typeII_twoChain: shows "two_chain_integral two_chain (partial_vector_derivative (\a. (F a) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_a' = "partial_vector_derivative (\a. (F a) \ j) i" have "(\(k,g)\ boundary twoCube. k * line_integral F {j} g) = integral (cubeImage twoCube) (\a. ?F_a' a)" if "twoCube \ two_chain" for twoCube using green_typeII_cube.GreenThm_typeII_twoCube(1) valid_typeII_div F_anal_valid one_chain_line_integral_def valid_two_chain_def by (simp add: R2_axioms green_typeII_cube_axioms_def green_typeII_cube_def that) then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. ?F_a' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using valid_typeII_div by (fastforce simp add: image_def valid_two_chain_def pairwise_def) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using valid_typeII_div image_iff by (fastforce simp add: valid_two_cube_def valid_two_chain_def) have boundary_inj: "inj_on boundary two_chain" using valid_typeII_div by (force simp add: valid_two_cube_def valid_two_chain_def pairwise_def inj_on_def) have "(\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g) = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {j} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {j} g)"] using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries, of "\x. case x of (k, g) \ (k::int) * line_integral F {j} g"] by auto then show "(\C\two_chain. integral (cubeImage C) (\a. ?F_a' a)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ k * line_integral F {j} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeII_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" (*This should follow from the assumption that images are not negligible*) shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} (two_chain_boundary two_chain)" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" have "integral s (\x. ?F_a' x) = two_chain_integral two_chain (\a. ?F_a' a)" proof (simp add: two_chain_integral_def) have partial_deriv_integrable: "(?F_a' has_integral (integral (cubeImage twoCube) ?F_a')) (cubeImage twoCube)" if "twoCube \ two_chain" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on F_anal_valid integrable_integral that) then have partial_deriv_integrable: "\twoCubeImg. twoCubeImg \ cubeImage ` two_chain \ ((\x. ?F_a' x) has_integral (integral (twoCubeImg) (\x. ?F_a' x))) (twoCubeImg)" using integrable_neg by force have finite_images: "finite (cubeImage ` two_chain)" using gen_division gen_division_def by auto have negligible_images: "pairwise (\S S'. negligible (S \ S')) (cubeImage ` two_chain)" using gen_division by (auto simp add: gen_division_def pairwise_def) have "inj_on cubeImage two_chain" using valid_typeII_div valid_two_chain_def by auto then have "(\twoCubeImg\cubeImage ` two_chain. integral twoCubeImg (\x. ?F_a' x)) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using sum.reindex by auto then show "integral s (\x. ?F_a' x) = (\C\two_chain. integral (cubeImage C) (\a. ?F_a' a))" using has_integral_Union[OF finite_images partial_deriv_integrable negligible_images] gen_division by (auto simp add: gen_division_def) qed then show ?thesis using GreenThm_typeII_twoChain F_anal_valid by auto qed lemma GreenThm_typeII_divisible_region_boundary_gen: assumes only_horizontal_division: "only_horizontal_division \ two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - from that obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (auto simp add: horiz_edge_def real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (auto simp: horiz_edge_def algebra_simps) have "\x. ?horiz_edge differentiable at x" using horiz_edge_is_straight_path straight_path_diffrentiable_y by (metis mult_commute_abs) then show "line_integral F {j} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) j_is_y_axis real_pair_basis horiz_edge_y_const by blast qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (force intro: sum.neutral) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have card_4: "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) show ?thesis using card_4 by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (fastforce simp add: Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (simp add: hor_vert_finite sum.union_disjoint one_chain_line_integral_def) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_horizontal_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_vertical_boundary two_chain) \ \ common_reparam_exists \ (two_chain_vertical_boundary two_chain))" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_horizontal_division by(fastforce simp add: only_horizontal_division_def) have "finite {j}" by auto then have eq_integrals: " one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_vertical_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_vertical_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(1) hv_props(5)] typeII_chain_line_integral_exists_boundary' by force next case False have integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {j} \" using typeII_chain_line_integral_exists_boundary' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {j} \" using typeII_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeII_edges_are_valid_paths valid_typeII_div by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) have integ_exis_horiz: "(\k \. (\(k', \')\two_chain_horizontal_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {j} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{j}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) have finite_j: "finite {j}" by auto have i: "{j} \ Basis" using j_is_y_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_vertical_boundary two_chain. \b\{j}. continuous_on (path_image \2) (\x. F x \ b)" using valid_typeII_div field_cont_on_typeII_region_cont_on_edges F_anal_valid by (fastforce simp add: analytically_valid_def two_chain_vertical_boundary_def boundary_def path_image_def) show "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using hv_props(4) False common_reparam_exists_imp_eq_line_integral(1)[OF finite_j hv_props(5) hor_vert_finite(1) hv_props(6) two_chain_vertical_boundary_is_boundary_chain ii _ hv_props(7) type_II_chain_vert_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" proof (simp only: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" using prod.collapse by blast then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_horizontal_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where horiz_edge_def: "(k',\') = (k', (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div kp_gammap by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) have horiz_edge_y_const: "\t. \ (t) \ j = x" using horiz_edge_def kp_gammap(4) by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis subpath_def) have horiz_edge_is_straight_path: "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" proof - fix t::real have "(1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1 = (1 - (b - a)*t + a) * (y2) + ((b-a)*t + a) * y1" by auto then have "\ = (\t::real. ((1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1, x::real))" using horiz_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "\ = (\t::real. ((1*y2 - (b - a)*y2*t - a*y2) + ((b-a)*y1*t + a*y1), x::real))" by(auto simp add: ring_class.ring_distribs(2) Groups.diff_diff_eq left_diff_distrib) also have "... = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" by (force simp add: left_diff_distrib) finally show "\ = (\t::real. ((1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t, x::real))" . qed show "line_integral F {j} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: horiz_edge_is_straight_path straight_path_diffrentiable_y) then have "line_integral F {j} \ = 0" by (simp add: horiz_edge_y_const line_integral_on_pair_straight_path(1)) then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {j} g) = (\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ real_of_int k * line_integral F {j} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(1) eq_integrals by (clarsimp simp add: one_chain_line_integral_def sum_zero_set) qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed lemma GreenThm_typeII_divisible_region_boundary: assumes two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (partial_vector_derivative (\x. (F x) \ j) i) = one_chain_line_integral F {j} \" proof - let ?F_a' = "(partial_vector_derivative (\x. (F x) \ j) i)" (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have horiz_line_integral_zero: "one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if one: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div one by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" using horiz_edge_def by (auto simp add: real_inner_class.inner_add_left euclidean_space_class.inner_not_same_Basis j_is_y_axis) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (metis horiz_edge_is_straight_path horiz_edge_y_const line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then show "(\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ k * line_integral F {j} g) = 0" by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using valid_typeII_div finite_image_iff gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using valid_typeII_div valid_two_chain_def that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def valid_two_cube_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False show ?thesis using valid_typeII_div valid_two_chain_def by (simp add: boundary_def vertical_boundary_def horizontal_boundary_def pairwise_def \twoCube \ twoCube2\ that(1) that(2)) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0::real, y)), (1::real, \y. twoCube (1::real, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \z. twoCube (z, 0::real)), (- 1, \z. twoCube (z, 1::real))}) ` two_chain) = {}" by (force simp add: Complete_Lattices.Union_disjoint) then show ?thesis by force qed qed have "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {j} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have y_axis_line_integral_is_only_vertical: "one_chain_line_integral F {j} (two_chain_boundary two_chain) = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using horiz_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>H(2-\) such that \ = \ \ \\<^sub>v(2-\)*) have "\\. \ \ (two_chain_horizontal_boundary two_chain) \ \ = \ \ (two_chain_vertical_boundary two_chain)" proof let ?\ = "\ - (two_chain_vertical_boundary two_chain)" show "?\ \ two_chain_horizontal_boundary two_chain \ \ = ?\ \ two_chain_vertical_boundary two_chain" using two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where h_props: "\ \ (two_chain_horizontal_boundary two_chain)" "\ = \ \ (two_chain_vertical_boundary two_chain)" by auto have h_vert_disj: "\ \ (two_chain_vertical_boundary two_chain) = {}" using horiz_verti_disjoint h_props(1) by auto have h_finite: "finite \" using hor_vert_finite h_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} \ + one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" by (auto simp add: one_chain_line_integral_def h_props sum.union_disjoint[OF h_finite hor_vert_finite(1) h_vert_disj]) (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {j} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {j} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_horizontal_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where horiz_edge_def: "oneCube = (k, (\t::real. ((1 - t) * (y2) + t * y1, x::real)))" using valid_typeII_div oneCube by (auto simp add: typeII_twoCube_def two_chain_horizontal_boundary_def horizontal_boundary_def) let ?horiz_edge = "(snd oneCube)" have horiz_edge_y_const: "\t. (?horiz_edge t) \ j = x" by (simp add: j_is_y_axis horiz_edge_def) have horiz_edge_is_straight_path: "?horiz_edge = (\t. (y2 + t * (y1 - y2), x))" by (simp add: add_diff_eq diff_add_eq mult.commute right_diff_distrib horiz_edge_def) show "line_integral F {j} (snd oneCube) = 0" by (simp add: horiz_edge_is_straight_path j_is_y_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_y) qed then have "\oneCube \ \. line_integral F {j} (snd oneCube) = 0" using h_props by auto then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {j} g) = 0" by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {j} g) = 0" by simp qed then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_vertical_boundary two_chain)" using line_integral_on_path by auto then have "one_chain_line_integral F {j} \ = one_chain_line_integral F {j} (two_chain_boundary two_chain)" using y_axis_line_integral_is_only_vertical by auto then show ?thesis using valid_typeII_div GreenThm_typeII_divisible by auto qed end locale green_typeI_cube = R2 + fixes twoC F assumes two_cube: "typeI_twoCube twoC" and valid_two_cube: "valid_two_cube twoC" and f_analytically_valid: "analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma GreenThm_typeI_twoCube: shows "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" "\(k,\) \ boundary twoC. line_integral_exists F {i} \" proof - let ?bottom_edge = "(\x. twoC(x, 0))" let ?right_edge = "(\y. twoC(1, y))" let ?top_edge = "(\x. twoC(x, 1))" let ?left_edge = "(\y. twoC(0, y))" have line_integral_around_boundary: "one_chain_line_integral F {i} (boundary twoC) = line_integral F {i} ?bottom_edge + line_integral F {i} ?right_edge - line_integral F {i} ?top_edge - line_integral F {i} ?left_edge" proof (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite1: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}" by auto have sum_step1: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite1] valid_two_cube by (auto simp: horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def card_insert_if split: if_split_asm) have three_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (- 1, \x. twoC (x, 1))} = 3" using valid_two_cube apply (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def valid_two_cube_def) by (auto simp: card_insert_if split: if_split_asm) have finite2: "finite {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))}" by auto have sum_step2: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y)), (-1, \x. twoC (x, 1))}. k * line_integral F {i} g) = (- line_integral F {i} (\x. twoC (x, 1))) + (\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) have two_distinct_edges: "card {(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y))} = 2" using three_distinct_edges by (simp add: one_chain_line_integral_def horizontal_boundary_def vertical_boundary_def boundary_def) have finite3: "finite {(- 1::int, \y. twoC (0, y))}" by auto have sum_step3: "(\(k, g)\{(- (1::int), \y. twoC (0, y)), (1, \y. twoC (1, y))}. k * line_integral F {i} g) = (line_integral F {i} (\y. twoC (1, y))) + (\(k, g)\{(- (1::real), \y. twoC (0, y))}. k * line_integral F {i} g)" using sum.insert_remove [OF finite2] three_distinct_edges by (auto simp: card_insert_if split: if_split_asm) show "(\x\{(- 1::int, \y. twoC (0, y)), (1, \y. twoC (1, y)), (1, \x. twoC (x, 0)), (- 1, \x. twoC (x, 1))}. case x of (k, g) \ k * line_integral F {i} g) = line_integral F {i} (\x. twoC (x, 0)) + line_integral F {i} (\y. twoC (1, y)) - line_integral F {i} (\x. twoC (x, 1)) - line_integral F {i} (\y. twoC (0, y))" using sum_step1 sum_step2 sum_step3 by auto qed obtain a b g1 g2 where twoCisTypeI: "a < b" "(\x \ cbox a b. g2 x \ g1 x)" "cubeImage twoC = {(x,y). x \ cbox a b \ y \ cbox (g2 x) (g1 x)}" "twoC = (\(x, y). ((1 - x) * a + x * b, (1 - y) * g2 ((1 - x) * a + x * b) + y * g1 ((1 - x) * a + x * b)))" "g1 piecewise_C1_differentiable_on {a .. b}" "g2 piecewise_C1_differentiable_on {a .. b}" "(\x. twoC(x, 0)) = (\x. (a + (b - a) * x, g2 (a + (b - a) * x)))" "(\y. twoC(1, y)) = (\x. (b, g2 b + x *\<^sub>R (g1 b - g2 b)))" "(\x. twoC(x, 1)) = (\x. (a + (b - a) * x, g1 (a + (b - a) * x)))" "(\y. twoC(0, y)) = (\x. (a, g2 a + x *\<^sub>R (g1 a - g2 a)))" using two_cube and typeI_cube_explicit_spec[of"twoC"] by auto have bottom_edge_smooth: "(\x. twoC (x, 0)) piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto have top_edge_smooth: "?top_edge piecewise_C1_differentiable_on {0..1}" using typeI_twoCube_smooth_edges two_cube boundary_def vertical_boundary_def horizontal_boundary_def by auto show "integral (cubeImage twoC) (\a. - partial_vector_derivative (\p. F p \ i) j a) = one_chain_line_integral F {i} (boundary twoC)" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto have "line_integral_exists F {i} (\y. twoC (0, y))" "line_integral_exists F {i} (\x. twoC (x, 1))" "line_integral_exists F {i} (\y. twoC (1, y))" "line_integral_exists F {i} (\x. twoC (x, 0))" using Greens_thm_type_I[OF twoCisTypeI(3) twoCisTypeI(7) bottom_edge_smooth twoCisTypeI(8) twoCisTypeI(9) top_edge_smooth twoCisTypeI(10) f_analytically_valid twoCisTypeI(2) twoCisTypeI(1)] line_integral_around_boundary by auto then have "line_integral_exists F {i} \" if "(k,\) \ boundary twoC" for k \ using that by (auto simp add: boundary_def horizontal_boundary_def vertical_boundary_def) then show "\(k,\) \ boundary twoC. line_integral_exists F {i} \" by auto qed lemma line_integral_exists_on_typeI_Cube_boundaries': assumes "(k,\) \ boundary twoC" shows "line_integral_exists F {i} \" using assms two_cube valid_two_cube f_analytically_valid GreenThm_typeI_twoCube(2) by blast end locale green_typeI_chain = R2 + fixes F two_chain s assumes valid_typeI_div: "valid_typeI_division s two_chain" and F_anal_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" begin lemma two_chain_valid_valid_cubes: "\two_cube \ two_chain. valid_two_cube two_cube" using valid_typeI_div by (auto simp add: valid_two_chain_def) lemma typeI_cube_line_integral_exists_boundary': assumes "\two_cube \ two_chain. typeI_twoCube two_cube" assumes "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\x. (F x) \ i) j" assumes "\two_cube \ two_chain. valid_two_cube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] assms using R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_boundary_def by fastforce then show integ_exis_vert: "\(k,\) \ two_chain_vertical_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary'': "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries'[of i j] valid_typeI_div apply (simp add: two_chain_boundary_def boundary_def) using F_anal_valid R2_axioms green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes by fastforce then show integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" by (simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) qed lemma typeI_cube_line_integral_exists_boundary: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary' typeI_cube_line_integral_exists_boundary'' apply (auto simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def) by (meson R2_axioms green_typeI_chain.F_anal_valid green_typeI_chain_axioms green_typeI_cube.line_integral_exists_on_typeI_Cube_boundaries' green_typeI_cube_axioms_def green_typeI_cube_def two_chain_valid_valid_cubes valid_typeI_div) lemma type_I_chain_horiz_bound_valid: "\(k,\) \ two_chain_horizontal_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths valid_typeI_div by (force simp add: two_chain_boundary_def two_chain_horizontal_boundary_def boundary_def) lemma type_I_chain_vert_bound_valid: (*This and the previous one need to be used in all proofs*) assumes "\two_cube \ two_chain. typeI_twoCube two_cube" shows "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms(1) by (force simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) lemma members_of_only_vertical_div_line_integrable': assumes "only_vertical_division one_chain two_chain" "(k::int, \)\one_chain" "(k::int, \)\one_chain" "finite two_chain" shows "line_integral_exists F {i} \" proof - have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by auto have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using type_I_chain_vert_bound_valid valid_typeI_div by linarith have integ_exis_vert: "(\k \. (\(k', \')\two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_horizontal_boundary_def two_chain_vertical_boundary_def boundary_def) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "one_chain = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using assms(1) unfolding only_vertical_division_def by blast have finite_i: "finite {i}" by auto show "line_integral_exists F {i} \" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True show ?thesis using gen_common_subdivision_imp_eq_line_integral(2)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) integ_exis_horiz finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(5) finite_i] integ_exis_vert[of "\"] assms(3) case_prod_conv hv_props(2) hv_props(3) by fastforce next case False have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms field_cont_on_typeI_region_cont_on_edges F_anal_valid valid_typeI_div by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) show "line_integral_exists F {i} \" using common_reparam_exists_imp_eq_line_integral(2)[OF finite_i hv_props(5) finite_two_chain_horizontal_boundary[OF assms(4)] hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii _ hv_props(7) type_I_chain_horiz_bound_valid] integ_exis_vert[of "\"] False assms(3) hv_props(2-4) by fastforce qed qed lemma GreenThm_typeI_two_chain: "two_chain_integral two_chain (\a. - partial_vector_derivative (\x. (F x) \ i) j a) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof (simp add: two_chain_boundary_def one_chain_line_integral_def two_chain_integral_def) let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) have "(\(k,g)\ boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" if "twoCube \ two_chain" for twoCube proof - have analytically_val: " analytically_valid (cubeImage twoCube) (\x. F x \ i) j" using that F_anal_valid by auto show "(\(k, g)\boundary twoCube. k * line_integral F {i} g) = integral (cubeImage twoCube) (\a. - ?F_b' a)" using green_typeI_cube.GreenThm_typeI_twoCube apply (simp add: one_chain_line_integral_def) by (simp add: R2_axioms analytically_val green_typeI_cube_axioms_def green_typeI_cube_def that two_chain_valid_valid_cubes valid_typeI_div) qed then have double_sum_eq_sum: "(\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g)) = (\twoCube\(two_chain). integral (cubeImage twoCube) (\a. - ?F_b' a))" using Finite_Cartesian_Product.sum_cong_aux by auto have pairwise_disjoint_boundaries: "\x\ (boundary ` two_chain). (\y\ (boundary ` two_chain). (x \ y \ (x \ y = {})))" using no_shared_edges_have_similar_orientations by (force simp add: image_def disjoint_iff_not_equal) have finite_boundaries: "\B \ (boundary` two_chain). finite B" using all_two_cubes_have_four_distict_edges using image_iff by fastforce have boundary_inj: "inj_on boundary two_chain" using all_two_cubes_have_four_distict_edges and no_shared_edges_have_similar_orientations by (force simp add: inj_on_def) have "(\x\(\(boundary` two_chain)). case x of (k, g) \ k * line_integral F {i} g) = (sum \ sum) (\x. case x of (k, g) \ (k::int) * line_integral F {i} g) (boundary ` two_chain)" using sum.Union_disjoint[OF finite_boundaries pairwise_disjoint_boundaries] by simp also have "... = (\twoCube\(two_chain).(\(k,g)\ boundary twoCube. k * line_integral F {i} g))" using sum.reindex[OF boundary_inj, of "\x. (\(k,g)\ x. k * line_integral F {i} g)"] by auto finally show "(\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\x. F x \ i) j)) = (\x\(\x\two_chain. boundary x). case x of (k, g) \ real_of_int k * line_integral F {i} g)" using double_sum_eq_sum by auto qed lemma GreenThm_typeI_divisible: assumes gen_division: "gen_division s (cubeImage ` two_chain)" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} (two_chain_boundary two_chain)" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i) j" have "integral s (\x. - ?F_b' x) = two_chain_integral two_chain (\a. - ?F_b' a)" proof (simp add: two_chain_integral_def) have "(\f\two_chain. integral (cubeImage f) (partial_vector_derivative (\p. F p \ i) j)) = integral s (partial_vector_derivative (\p. F p \ i) j)" by (metis analytically_valid_imp_part_deriv_integrable_on F_anal_valid gen_division two_chain_integral_def two_chain_integral_eq_integral_divisable valid_typeI_div) then show "- integral s (partial_vector_derivative (\a. F a \ i) j) = (\C\two_chain. - integral (cubeImage C) (partial_vector_derivative (\a. F a \ i) j))" by (simp add: sum_negf) qed then show ?thesis using GreenThm_typeI_two_chain assms by auto qed lemma GreenThm_typeI_divisible_region_boundary: assumes gen_division: "gen_division s (cubeImage ` two_chain)" and two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain \ \" and boundary_of_region_is_subset_of_partition_boundary: "\ \ two_chain_boundary two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" using vert_edge_def Product_Type.snd_conv by (auto simp add: mult.commute right_diff_distrib') show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def by (metis all_two_cubes_have_four_distict_edges card.infinite finite_UN_I finite_imageD gen_division gen_division_def zero_neq_numeral valid_typeI_div valid_two_chain_def) have boundary_is_vert_hor: "(two_chain_boundary two_chain) = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by(auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \z. x (z, 0)), (- 1, \z. x (z, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \z. twoCube2 (z, 0)), (- 1, \z. twoCube2 (z, 1))} = {}" if "twoCube\two_chain" "twoCube2\two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0::real, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: True card_insert_if split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(-1::int, \y. twoCube (0,y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1, \y. twoCube (y, 0)), (-1, \z. twoCube (z, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) have "\\. \ \ (two_chain_vertical_boundary two_chain) \ \ = \ \ (two_chain_horizontal_boundary two_chain)" proof let ?\ = "\ - (two_chain_horizontal_boundary two_chain)" show "?\ \ two_chain_vertical_boundary two_chain \ \ = ?\ \ two_chain_horizontal_boundary two_chain" using two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundary boundary_is_vert_hor by blast qed then obtain \ where v_props: "\ \ (two_chain_vertical_boundary two_chain)" "\ = \ \ (two_chain_horizontal_boundary two_chain)" by auto have v_horiz_disj: "\ \ (two_chain_horizontal_boundary two_chain) = {}" using horiz_verti_disjoint v_props(1) by auto have v_finite: "finite \" using hor_vert_finite v_props(1) Finite_Set.rev_finite_subset by force have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" by(auto simp add: one_chain_line_integral_def v_props sum.union_disjoint[OF v_finite hor_vert_finite(2) v_horiz_disj]) (*Since \ \ \\<^sub>V(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have "one_chain_line_integral F {i} \ = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) have "\x. ?vert_edge differentiable at x" by (metis mult.commute vert_edge_is_straight_path straight_path_diffrentiable_x) then show "line_integral F {i} (snd oneCube) = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast qed then have "\oneCube \ \. line_integral F {i} (snd oneCube) = 0" using v_props by auto then show "(\x\\. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal by (simp add: line_integral_on_path) then show ?thesis using assms and GreenThm_typeI_divisible by auto qed lemma GreenThm_typeI_divisible_region_boundary_gen: assumes valid_typeI_div: "valid_typeI_division s two_chain" and f_analytically_valid: "\twoC \ two_chain. analytically_valid (cubeImage twoC) (\a. F(a) \ i) j" and only_vertical_division: "only_vertical_division \ two_chain" shows "integral s (\x. - partial_vector_derivative (\a. F(a) \ i) j x) = one_chain_line_integral F {i} \" proof - let ?F_b' = "partial_vector_derivative (\a. F(a) \ i)" have all_two_cubes_have_four_distict_edges: "\twoCube \ two_chain. card (boundary twoCube) = 4" using valid_typeI_div valid_two_chain_def valid_two_cube_def by auto have no_shared_edges_have_similar_orientations: "\twoCube1 \ two_chain. \twoCube2 \ two_chain. twoCube1 \ twoCube2 \ boundary twoCube1 \ boundary twoCube2 = {}" using valid_typeI_div valid_two_chain_def by (auto simp add: pairwise_def) (*Proving that line_integral on the x axis is 0 for the vertical 1-cubes*) have vert_line_integral_zero: "one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) = 0" proof (simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ two_chain_vertical_boundary(two_chain)" for oneCube proof - obtain x y1 y2 k where vert_edge_def: "oneCube = (k, (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div oneCube by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) let ?vert_edge = "(snd oneCube)" have vert_edge_x_const: "\t. (?vert_edge t) \ i = x" by (simp add: i_is_x_axis vert_edge_def) have vert_edge_is_straight_path: "?vert_edge = (\t. (x, y2 + t * (y1 - y2)))" by (auto simp: vert_edge_def algebra_simps) show ?thesis by (simp add: i_is_x_axis line_integral_on_pair_straight_path(1) mult.commute straight_path_diffrentiable_x vert_edge_is_straight_path) qed then show "(\x\two_chain_vertical_boundary two_chain. case x of (k, g) \ k * line_integral F {i} g) = 0" using comm_monoid_add_class.sum.neutral by (simp add: prod.case_eq_if) qed (*then, the x axis line_integral on the boundaries of the twoCube is equal to the line_integral on the horizontal boundaries of the twoCube \ 1*) have boundary_is_finite: "finite (two_chain_boundary two_chain)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain" using assms(1) finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed have boundary_is_vert_hor: "two_chain_boundary two_chain = (two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain)" by (auto simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) then have hor_vert_finite: "finite (two_chain_vertical_boundary two_chain)" "finite (two_chain_horizontal_boundary two_chain)" using boundary_is_finite by auto have horiz_verti_disjoint: "(two_chain_vertical_boundary two_chain) \ (two_chain_horizontal_boundary two_chain) = {}" proof (simp add: two_chain_vertical_boundary_def two_chain_horizontal_boundary_def horizontal_boundary_def vertical_boundary_def) show "(\x\two_chain. {(- 1, \y. x (0, y)), (1::int, \y. x (1::real, y))}) \ (\x\two_chain. {(1, \y. x (y, 0)), (- 1, \y. x (y, 1))}) = {}" proof - have "{(- 1, \y. twoCube (0, y)), (1::int, \y. twoCube (1, y))} \ {(1, \y. twoCube2 (y, 0)), (- 1, \y. twoCube2 (y, 1))} = {}" if "twoCube \ two_chain" "twoCube2 \ two_chain" for twoCube twoCube2 proof (cases "twoCube = twoCube2") case True have "card {(- 1::int, \y. twoCube2 (0, y)), (1::int, \y. twoCube2 (1, y)), (1, \x. twoCube2 (x, 0)), (- 1, \x. twoCube2 (x, 1))} = 4" using all_two_cubes_have_four_distict_edges that(2) by (auto simp add: boundary_def vertical_boundary_def horizontal_boundary_def) then show ?thesis by (auto simp: card_insert_if True split: if_split_asm) next case False then show ?thesis using no_shared_edges_have_similar_orientations by (simp add: that boundary_def vertical_boundary_def horizontal_boundary_def) qed then have "\ ((\twoCube. {(- 1, \y. twoCube (0, y)), (1, \y. twoCube (1, y))}) ` two_chain) \ \ ((\twoCube. {(1::int, \y. twoCube (y, 0)), (- 1, \y. twoCube (y, 1))}) ` two_chain) = {}" using Complete_Lattices.Union_disjoint by force then show ?thesis by force qed qed have "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_vertical_boundary two_chain) + one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using boundary_is_vert_hor horiz_verti_disjoint by (auto simp add: one_chain_line_integral_def hor_vert_finite sum.union_disjoint) then have x_axis_line_integral_is_only_horizontal: "one_chain_line_integral F {i} (two_chain_boundary two_chain) = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using vert_line_integral_zero by auto (*Since \ \ the boundary of the 2-chain and the horizontal boundaries are \ \, then there is some \ \ \\<^sub>V(2-\) such that \ = \ \ \\<^sub>H(2-\)*) obtain \ \ where hv_props: "finite \" "(\(k,\) \ \. (\(k', \') \ two_chain_vertical_boundary two_chain. (\a \ {0 .. 1}. \b \ {0..1}. a \ b \ subpath a b \' = \)))" "\ = \ \ \" "(common_sudiv_exists (two_chain_horizontal_boundary two_chain) \) \ common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" "finite \" "boundary_chain \" "\(k,\)\\. valid_path \" using only_vertical_division by (auto simp add: only_vertical_division_def) have "finite {i}" by auto then have eq_integrals: " one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof(cases "common_sudiv_exists (two_chain_horizontal_boundary two_chain) \") case True then show ?thesis using gen_common_subdivision_imp_eq_line_integral(1)[OF True two_chain_horizontal_boundary_is_boundary_chain hv_props(6) _ hor_vert_finite(2) hv_props(5)] typeI_cube_line_integral_exists_boundary'' by force next case False have integ_exis_horiz: "\(k,\) \ two_chain_horizontal_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary'' assms by (fastforce simp add: valid_two_chain_def) have integ_exis: "\(k,\) \ two_chain_boundary two_chain. line_integral_exists F {i} \" using typeI_cube_line_integral_exists_boundary by blast have valid_paths: "\(k,\) \ two_chain_vertical_boundary two_chain. valid_path \" using typeI_edges_are_valid_paths assms by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def boundary_def) have integ_exis_vert: "(\k \. (\(k', \') \ two_chain_vertical_boundary two_chain. \a\{0..1}. \b\{0..1}. a \ b \ subpath a b \' = \) \ line_integral_exists F {i} \)" using integ_exis valid_paths line_integral_exists_subpath[of "F" "{i}"] by (fastforce simp add: two_chain_boundary_def two_chain_vertical_boundary_def two_chain_horizontal_boundary_def boundary_def) have finite_i: "finite {i}" by auto have i: "{i} \ Basis" using i_is_x_axis real_pair_basis by auto have ii: " \(k2, \2)\two_chain_horizontal_boundary two_chain. \b\{i}. continuous_on (path_image \2) (\x. F x \ b)" using assms(1) field_cont_on_typeI_region_cont_on_edges assms(2) by (fastforce simp add: analytically_valid_def two_chain_horizontal_boundary_def boundary_def path_image_def) have *: "common_reparam_exists \ (two_chain_horizontal_boundary two_chain)" using hv_props(4) False by auto show "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" using common_reparam_exists_imp_eq_line_integral(1)[OF finite_i hv_props(5) hor_vert_finite(2) hv_props(6) two_chain_horizontal_boundary_is_boundary_chain ii * hv_props(7) type_I_chain_horiz_bound_valid] by fastforce qed (*Since \ \ \\<^sub>H(2-\), then the line_integral on the x axis along \ is 0, and from 1 Q.E.D. *) have line_integral_on_path: "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_horizontal_boundary two_chain)" proof (auto simp add: one_chain_line_integral_def) have "line_integral F {i} (snd oneCube) = 0" if oneCube: "oneCube \ \" for oneCube proof - obtain k \ where k_gamma: "(k,\) = oneCube" by (metis coeff_cube_to_path.cases) then obtain k' \' a b where kp_gammap: "(k'::int, \') \ two_chain_vertical_boundary two_chain" "a \ {0 .. 1}" "b \ {0..1}" "subpath a b \' = \" using hv_props oneCube by (smt case_prodE split_conv) obtain x y1 y2 where vert_edge_def: "(k',\') = (k', (\t::real. (x::real, (1 - t) * (y2) + t * y1)))" using valid_typeI_div kp_gammap by (auto simp add: typeI_twoCube_def two_chain_vertical_boundary_def vertical_boundary_def) have vert_edge_x_const: "\t. \ (t) \ i = x" by (metis (no_types, lifting) Pair_inject fstI i_is_x_axis inner_Pair_0(2) kp_gammap(4) real_inner_1_right subpath_def vert_edge_def) have "\ = (\t::real. (x::real, (1 - (b - a)*t - a) * (y2) + ((b-a)*t + a) * y1))" using vert_edge_def Product_Type.snd_conv Product_Type.fst_conv kp_gammap(4) by (simp add: subpath_def diff_diff_eq[symmetric]) also have "... = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" by (simp add: algebra_simps) finally have vert_edge_is_straight_path: "\ = (\t::real. (x::real, (1*y2 - a*y2) + a*y1 + ((b-a)*y1 - (b - a)*y2)*t))" . show "line_integral F {i} (snd oneCube) = 0" proof - have "\x. \ differentiable at x" by (simp add: straight_path_diffrentiable_x vert_edge_is_straight_path) then have "line_integral F {i} \ = 0" using line_integral_on_pair_straight_path(1) vert_edge_x_const by blast then show ?thesis using Product_Type.snd_conv k_gamma by auto qed qed then have "\x\\. (case x of (k, g) \ (k::int) * line_integral F {i} g) = 0" by auto then show "(\x\\. case x of (k, g) \ real_of_int k * line_integral F {i} g) = (\x\two_chain_horizontal_boundary two_chain. case x of (k, g) \ of_int k * line_integral F {i} g)" using hv_props(1) hv_props(3) hv_props(5) sum_zero_set hor_vert_finite(2) eq_integrals apply(auto simp add: one_chain_line_integral_def) by (smt Un_commute sum_zero_set) qed then have "one_chain_line_integral F {i} \ = one_chain_line_integral F {i} (two_chain_boundary two_chain)" using x_axis_line_integral_is_only_horizontal line_integral_on_path by auto then show ?thesis using assms GreenThm_typeI_divisible by auto qed end locale green_typeI_typeII_chain = R2: R2 i j + T1: green_typeI_chain i j F two_chain_typeI + T2: green_typeII_chain i j F two_chain_typeII for i j F two_chain_typeI two_chain_typeII begin lemma GreenThm_typeI_typeII_divisible_region_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ \" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ \" and boundary_of_region_is_subset_of_partition_boundaries: "\ \ two_chain_boundary two_chain_typeI" "\ \ two_chain_boundary two_chain_typeII" shows "integral s (\x. partial_vector_derivative (\a. F a \ j) i x - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i, j} \" proof - let ?F_b' = "partial_vector_derivative (\a. F a \ i) j" let ?F_a' = "partial_vector_derivative (\a. F a \ j) i" have typeI_regions_integral: "integral s (\x. - partial_vector_derivative (\a. F a \ i) j x) = one_chain_line_integral F {i} \" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by blast have typeII_regions_integral: "integral s (partial_vector_derivative (\x. F x \ j) i) = one_chain_line_integral F {j} \" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "\twoCube \ two_chain_typeII. (?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_iff) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed have line_integral_dist: "one_chain_line_integral F {i, j} \ = one_chain_line_integral F {i} \ + one_chain_line_integral F {j} \" proof (simp add: one_chain_line_integral_def) have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if kg: "(k,g) \ \" for k g proof - obtain twoCube_typeI where twoCube_typeI_props: "twoCube_typeI \ two_chain_typeI" "(k, g) \ boundary twoCube_typeI" "typeI_twoCube twoCube_typeI" "continuous_on (cubeImage twoCube_typeI) (\x. F(x) \ i)" using boundary_of_region_is_subset_of_partition_boundaries(1) two_chain_boundary_def T1.valid_typeI_div T1.F_anal_valid kg by (auto simp add: analytically_valid_def) obtain twoCube_typeII where twoCube_typeII_props: "twoCube_typeII \ two_chain_typeII" "(k, g) \ boundary twoCube_typeII" "typeII_twoCube twoCube_typeII" "continuous_on (cubeImage twoCube_typeII) (\x. F(x) \ j)" using boundary_of_region_is_subset_of_partition_boundaries(2) two_chain_boundary_def T2.valid_typeII_div kg T2.F_anal_valid by (auto simp add: analytically_valid_def) have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have int_exists_i: "line_integral_exists F {i} g" using T1.typeI_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have int_exists_j: "line_integral_exists F {j} g" using T2.typeII_cube_line_integral_exists_boundary assms kg by (auto simp add: valid_two_chain_def) have finite: "finite {i, j}" by auto show ?thesis using line_integral_sum_gen[OF finite int_exists_i int_exists_j] R2.i_is_x_axis R2.j_is_y_axis by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then have line_integral_distrib: "(\(k,g)\\. k * line_integral F {i, j} g) = (\(k,g)\\. k * line_integral F {i} g + k * line_integral F {j} g)" by (force intro: sum.cong split_cong) have "(\x. (case x of (k, g) \ (k::int) * line_integral F {i} g) + (case x of (k, g) \ (k::int) * line_integral F {j} g)) = (\x. (case x of (k, g) \ (k * line_integral F {i} g) + (k::int) * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then show "(\(k, g)\\. k * line_integral F {i, j} g) = (\(k, g)\\. (k::int) * line_integral F {i} g) + (\(k, g)\\. (k::int) * line_integral F {j} g)" using comm_monoid_add_class.sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" \] line_integral_distrib by presburger qed show ?thesis using integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region': assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_gen_common_subdiv: "common_sudiv_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have "finite two_chain_typeI" using T1.valid_typeI_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k, \)\one_chain_typeI. line_integral_exists F {i} \" using T1.members_of_only_vertical_div_line_integrable' assms by fastforce have "finite (two_chain_horizontal_boundary two_chain_typeI)" by (meson T1.valid_typeI_div finite_imageD finite_two_chain_horizontal_boundary gen_division_def valid_two_chain_def) then have "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" and "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using gen_common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_gen_common_subdiv only_vertical_division(2) only_horizontal_division(2)] ii by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \)" proof (intro conjI) have "finite two_chain_typeII" using T2.valid_typeII_div finite_image_iff by (auto simp add: gen_division_def valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeII. line_integral_exists F {j} \" using T2.members_of_only_horiz_div_line_integrable' assms T2.two_chain_valid_valid_cubes by blast have typeII_and_typI_one_chains_have_common_subdiv: "common_sudiv_exists one_chain_typeII one_chain_typeI" by (simp add: common_sudiv_exists_comm typeI_and_typII_one_chains_have_gen_common_subdiv) have iv: "finite one_chain_typeI" using only_vertical_division(1) only_vertical_division_def by auto moreover have iv': "finite one_chain_typeII" using only_horizontal_division(1) only_horizontal_division_def by auto ultimately show "one_chain_line_integral F {j} one_chain_typeII = one_chain_line_integral F {j} one_chain_typeI" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using gen_common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv only_horizontal_division(2) only_vertical_division(2) ii] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary_gen T1.valid_typeI_div T1.F_anal_valid only_vertical_division(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary_gen T2.valid_typeII_div T2.F_anal_valid only_horizontal_division(1) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral_exists F {i} g" "line_integral_exists F {j} g" "finite {i, j}" using one_chain_i_integrals one_chain_j_integrals that by fastforce+ moreover have "{i} \ {j} = {}" by (simp add: R2.i_is_x_axis R2.j_is_y_axis) ultimately have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (metis insert_is_Un line_integral_sum_gen(1)) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" by (smt that disjoint_insert(2) finite.emptyI finite.insertI R2.i_is_x_axis inf_bot_right insert_absorb insert_commute insert_is_Un R2.j_is_y_axis line_integral_sum_gen(1) one_chain_i_integrals one_chain_j_integrals prod.case_eq_if singleton_inject snd_conv) then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" using T2.valid_typeII_div unfolding gen_division_def by (metis has_integral_Union) then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" using analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid by blast then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" using T1.valid_typeI_div unfolding gen_division_def by (metis has_integral_Union) then show ?thesis by (simp add: F_a'_integrable Henstock_Kurzweil_Integration.integral_diff has_integral_iff) qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed lemma GreenThm_typeI_typeII_divisible_region: assumes only_vertical_division: "only_vertical_division one_chain_typeI two_chain_typeI" "boundary_chain one_chain_typeI" and only_horizontal_division: "only_horizontal_division one_chain_typeII two_chain_typeII" "boundary_chain one_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" using GreenThm_typeI_typeII_divisible_region' only_vertical_division only_horizontal_division common_subdiv_imp_gen_common_subdiv[OF typeI_and_typII_one_chains_have_common_subdiv] by auto lemma GreenThm_typeI_typeII_divisible_region_finite_holes: assumes valid_cube_boundary: "\(k,\)\boundary C. valid_path \" and only_vertical_division: "only_vertical_division (boundary C) two_chain_typeI" and only_horizontal_division: "only_horizontal_division (boundary C) two_chain_typeII" and s_is_oneCube: "s = cubeImage C" shows "integral (cubeImage C) (\x. partial_vector_derivative (\x. F x \ j) i x - partial_vector_derivative (\x. F x \ i) j x) = one_chain_line_integral F {i, j} (boundary C)" using GreenThm_typeI_typeII_divisible_region[OF only_vertical_division two_cube_boundary_is_boundary only_horizontal_division two_cube_boundary_is_boundary common_boundary_subdiv_exists_refl[OF assms(1)]] s_is_oneCube by auto lemma GreenThm_typeI_typeII_divisible_region_equivallent_boundary: assumes gen_divisions: "gen_division s (cubeImage ` two_chain_typeI)" "gen_division s (cubeImage ` two_chain_typeII)" and typeI_two_cubes_trace_horizontal_boundaries: "two_chain_horizontal_boundary two_chain_typeI \ one_chain_typeI" and typeII_two_cubes_trace_vertical_boundaries: "two_chain_vertical_boundary two_chain_typeII \ one_chain_typeII" and boundary_of_region_is_subset_of_partition_boundaries: "one_chain_typeI \ two_chain_boundary two_chain_typeI" "one_chain_typeII \ two_chain_boundary two_chain_typeII" and typeI_and_typII_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeI one_chain_typeII" shows "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeI" "integral s (\x. partial_vector_derivative (\x. (F x) \ j) i x - partial_vector_derivative (\x. (F x) \ i) j x) = one_chain_line_integral F {i, j} one_chain_typeII" proof - let ?F_b' = "partial_vector_derivative (\x. (F x) \ i) j" let ?F_a' = "partial_vector_derivative (\x. (F x) \ j) i" have one_chain_i_integrals: "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {i} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {i} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(1) by blast have i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_chain_def boundary_of_region_is_subset_of_partition_boundaries(2) by blast have "\k \. (k,\)\one_chain_typeI \ line_integral_exists F {i} \" using T1.typeI_cube_line_integral_exists_boundary assms by (fastforce simp add: valid_two_chain_def) then show ii: "\(k,\)\one_chain_typeI. line_integral_exists F {i} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce moreover have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce ultimately show "one_chain_line_integral F {i} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeII" "\(k, \)\one_chain_typeII. line_integral_exists F {i} \" using ii common_subdivision_imp_eq_line_integral[OF typeI_and_typII_one_chains_have_common_subdiv i i' ii] by auto qed have one_chain_j_integrals: "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII \ (\(k,\)\one_chain_typeI. line_integral_exists F {j} \) \ (\(k,\)\one_chain_typeII. line_integral_exists F {j} \)" proof (intro conjI) have i: "boundary_chain one_chain_typeI" and i': "boundary_chain one_chain_typeII" using two_chain_boundary_is_boundary_chain boundary_of_region_is_subset_of_partition_boundaries unfolding boundary_chain_def by blast+ have "line_integral_exists F {j} \" if "(k,\)\one_chain_typeII" for k \ proof - have F_is_continuous: "\twoC \ two_chain_typeII. continuous_on (cubeImage twoC) (\a. F(a) \ j)" using T2.F_anal_valid by(simp add: analytically_valid_def) show "line_integral_exists F {j} \" using that T2.valid_typeII_div boundary_of_region_is_subset_of_partition_boundaries(2) using green_typeII_cube.line_integral_exists_on_typeII_Cube_boundaries' assms valid_two_chain_def apply (simp add: two_chain_boundary_def) by (metis T2.typeII_cube_line_integral_exists_boundary case_prodD subset_iff that two_chain_boundary_def) qed then show ii: " \(k,\)\one_chain_typeII. line_integral_exists F {j} \" by auto have "finite (two_chain_boundary two_chain_typeI)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeI" using T1.valid_typeI_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeI \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv: "finite one_chain_typeI" using boundary_of_region_is_subset_of_partition_boundaries(1) finite_subset by fastforce have "finite (two_chain_boundary two_chain_typeII)" unfolding two_chain_boundary_def proof (rule finite_UN_I) show "finite two_chain_typeII" using T2.valid_typeII_div finite_imageD gen_division_def valid_two_chain_def by auto show "\a. a \ two_chain_typeII \ finite (boundary a)" by (simp add: boundary_def horizontal_boundary_def vertical_boundary_def) qed then have iv': "finite one_chain_typeII" using boundary_of_region_is_subset_of_partition_boundaries(2) finite_subset by fastforce have typeII_and_typI_one_chains_have_common_subdiv: "common_boundary_sudivision_exists one_chain_typeII one_chain_typeI" using typeI_and_typII_one_chains_have_common_subdiv common_boundary_sudivision_commutative by auto show "one_chain_line_integral F {j} one_chain_typeI = one_chain_line_integral F {j} one_chain_typeII" "\(k, \)\one_chain_typeI. line_integral_exists F {j} \" using common_subdivision_imp_eq_line_integral[OF typeII_and_typI_one_chains_have_common_subdiv i' i ii iv' iv] ii by auto qed have typeI_regions_integral: "integral s (\x. - ?F_b' x) = one_chain_line_integral F {i} one_chain_typeI" using T1.GreenThm_typeI_divisible_region_boundary gen_divisions(1) typeI_two_cubes_trace_horizontal_boundaries boundary_of_region_is_subset_of_partition_boundaries(1) by auto have typeII_regions_integral: "integral s ?F_a' = one_chain_line_integral F {j} one_chain_typeII" using T2.GreenThm_typeII_divisible_region_boundary gen_divisions(2) typeII_two_cubes_trace_vertical_boundaries boundary_of_region_is_subset_of_partition_boundaries(2) by auto have line_integral_dist: "one_chain_line_integral F {i, j} one_chain_typeI = one_chain_line_integral F {i} one_chain_typeI + one_chain_line_integral F {j} one_chain_typeI \ one_chain_line_integral F {i, j} one_chain_typeII = one_chain_line_integral F {i} one_chain_typeII + one_chain_line_integral F {j} one_chain_typeII" proof (simp add: one_chain_line_integral_def) have line_integral_distrib: "(\(k,g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeI. k * line_integral F {i} g + k * line_integral F {j} g) \ (\(k,g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k,g)\one_chain_typeII. k * line_integral F {i} g + k * line_integral F {j} g)" proof - have 0: "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeII" for k g proof - have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto show ?thesis using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto qed then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed have "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" if "(k,g) \ one_chain_typeI" for k g proof - have finite: "finite {i, j}" by auto have line_integral_all: "\i\{i, j}. line_integral_exists F {i} g" using one_chain_i_integrals one_chain_j_integrals that by auto have "line_integral F {i, j} g = line_integral F {i} g + line_integral F {j} g" using line_integral_sum_gen[OF finite] R2.i_is_x_axis R2.j_is_y_axis line_integral_all by auto then show "k * line_integral F {i, j} g = k * line_integral F {i} g + k * line_integral F {j} g" by (simp add: distrib_left) qed then show ?thesis using 0 by (smt sum.cong split_cong) qed show "(\(k::int, g)\one_chain_typeI. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeI. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeI. k * line_integral F {j} g) \ (\(k::int, g)\one_chain_typeII. k * line_integral F {i, j} g) = (\(k, g)\one_chain_typeII. k * line_integral F {i} g) + (\(k::int, g)\one_chain_typeII. k * line_integral F {j} g)" proof - have 0: "(\x. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x. (case x of (k::int, g) \ (k * line_integral F {i} g) + k * line_integral F {j} g))" using comm_monoid_add_class.sum.distrib by auto then have 1: "(\x\one_chain_typeI. (case x of (k::int, g) \ k * line_integral F {i} g) + (case x of (k::int, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeI. (case x of (k::int, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" by presburger have "(\x\one_chain_typeII. (case x of (k, g) \ k * line_integral F {i} g) + (case x of (k, g) \ k * line_integral F {j} g)) = (\x\one_chain_typeII. (case x of (k, g) \( k * line_integral F {i} g + k * line_integral F {j} g)))" using 0 by presburger then show ?thesis using sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeI"] sum.distrib[of "(\(k, g). k * line_integral F {i} g)" "(\(k, g). k * line_integral F {j} g)" "one_chain_typeII"] line_integral_distrib 1 by auto qed qed have integral_dis: "integral s (\x. ?F_a' x - ?F_b' x) = integral s (\x. ?F_a' x) + integral s (\x. - ?F_b' x)" proof - have "(?F_a' has_integral (\img\cubeImage ` two_chain_typeII. integral img ?F_a')) s" proof - have "(?F_a' has_integral integral (cubeImage twoCube) ?F_a') (cubeImage twoCube)" if "twoCube \ two_chain_typeII" for twoCube by (simp add: analytically_valid_imp_part_deriv_integrable_on T2.F_anal_valid has_integral_integrable_integral that) then have "\u. u \ (cubeImage ` two_chain_typeII) \ (?F_a' has_integral integral u ?F_a') u" by auto then show ?thesis using gen_divisions(2) unfolding gen_division_def by (metis has_integral_Union) qed then have F_a'_integrable: "(?F_a' integrable_on s)" by auto have "(?F_b' has_integral (\img\cubeImage ` two_chain_typeI. integral img ?F_b')) s" proof - have "\twoCube \ two_chain_typeI. (?F_b' has_integral integral (cubeImage twoCube) ?F_b') (cubeImage twoCube)" by (simp add: analytically_valid_imp_part_deriv_integrable_on T1.F_anal_valid has_integral_integrable_integral) then have "\u. u \ (cubeImage ` two_chain_typeI) \ (?F_b' has_integral integral u ?F_b') u" by auto then show ?thesis using gen_divisions(1) unfolding gen_division_def by (metis has_integral_Union) qed then show ?thesis using F_a'_integrable Henstock_Kurzweil_Integration.integral_diff by auto qed show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeI" using one_chain_j_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto show "integral s (\x. ?F_a' x - ?F_b' x) = one_chain_line_integral F {i, j} one_chain_typeII" using one_chain_i_integrals integral_dis line_integral_dist typeI_regions_integral typeII_regions_integral by auto qed end end diff --git a/thys/MFMC_Countable/MFMC_Bounded.thy b/thys/MFMC_Countable/MFMC_Bounded.thy --- a/thys/MFMC_Countable/MFMC_Bounded.thy +++ b/thys/MFMC_Countable/MFMC_Bounded.thy @@ -1,653 +1,653 @@ (* Author: Andreas Lochbihler, Digital Asset *) section \The max-flow min-cut theorem in bounded networks\ subsection \Linkages in unhindered bipartite webs\ theory MFMC_Bounded imports Matrix_For_Marginals MFMC_Reduction begin context countable_bipartite_web begin lemma countable_A [simp]: "countable (A \)" using A_vertex countable_V by(blast intro: countable_subset) lemma unhindered_criterion [rule_format]: assumes "\ hindered \" shows "\X \ A \. finite X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\\<^bold>E `` X. weight \ y)" using assms proof(rule contrapos_np) assume "\ ?thesis" then obtain X where "X \ {X. X \ A \ \ finite X \ (\\<^sup>+ y\\<^bold>E `` X. weight \ y) < (\\<^sup>+ x\X. weight \ x)}" (is "_ \ Collect ?P") by(auto simp add: not_le) from wf_eq_minimal[THEN iffD1, OF wf_finite_psubset, rule_format, OF this, simplified] obtain X where X_A: "X \ A \" and fin_X [simp]: "finite X" and less: "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) < (\\<^sup>+ x\X. weight \ x)" and minimal: "\X'. X' \ X \ (\\<^sup>+ x\X'. weight \ x) \ (\\<^sup>+ y\\<^bold>E `` X'. weight \ y)" by(clarsimp simp add: not_less)(meson finite_subset order_trans psubset_imp_subset) have nonempty: "X \ {}" using less by auto then obtain xx where xx: "xx \ X" by auto define f where "f x = (if x = xx then (\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x) else if x \ X then weight \ x else 0)" for x define g where "g y = (if y \ \<^bold>E `` X then weight \ y else 0)" for y define E' where "E' \ \<^bold>E \ X \ UNIV" have Xxx: "X - {xx} \ X" using xx by blast have E [simp]: "E' `` X' = \<^bold>E `` X'" if "X' \ X" for X' using that by(auto simp add: E'_def) have in_E': "(x, y) \ E' \ x \ X \ (x, y) \ \<^bold>E" for x y by(auto simp add: E'_def) have "(\\<^sup>+ x\X. f x) = (\\<^sup>+ x\X - {xx}. f x) + (\\<^sup>+ x\{xx}. f x)" using xx by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+ x\X - {xx}. weight \ x) + ((\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x))" by(rule arg_cong2[where f="(+)"])(auto simp add: f_def xx nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y\\<^bold>E `` X. g y)" using minimal[OF Xxx] xx by(subst add_diff_eq_iff_ennreal[THEN iffD2])(fastforce simp add: g_def[abs_def] nn_integral_count_space_indicator intro!: nn_integral_cong intro: nn_integral_mono elim: order_trans split: split_indicator)+ finally have sum_eq: "(\\<^sup>+ x\X. f x) = (\\<^sup>+ y\\<^bold>E `` X. g y)" . have "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = (\\<^sup>+ y\\<^bold>E `` X. g y)" by(auto simp add: nn_integral_count_space_indicator g_def intro!: nn_integral_cong) then have fin: "\ \ \" using less by auto have fin2: "(\\<^sup>+ x\X'. weight \ x) \ \" if "X' \ X" for X' proof - have "(\\<^sup>+ x\\<^bold>E `` X'. weight \ x) \ (\\<^sup>+ x\\<^bold>E `` X. weight \ x)" using that by(auto 4 3 simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator split_indicator_asm) then show ?thesis using minimal[OF that] less by(auto simp add: top_unique) qed have "f xx = (\\<^sup>+ y\\<^bold>E `` X. weight \ y) - (\\<^sup>+ x\X - {xx}. weight \ x)" by (simp add: f_def) also have "\ < (\\<^sup>+ x\X. weight \ x) - (\\<^sup>+ x\X - {xx}. weight \ x)" using less fin2[OF Xxx] minimal[OF Xxx] by(subst minus_less_iff_ennreal)(fastforce simp add: less_top[symmetric] nn_integral_count_space_indicator diff_add_self_ennreal intro: nn_integral_mono elim: order_trans split: split_indicator)+ also have "\ = (\\<^sup>+ x\{xx}. weight \ x)" using fin2[OF Xxx] xx apply(simp add: nn_integral_count_space_indicator del: nn_integral_indicator_singleton) apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space split: split_indicator simp del: nn_integral_indicator_singleton intro!: nn_integral_cong) done also have "\ = weight \ xx" by(simp add: nn_integral_count_space_indicator) finally have fxx: "f xx < weight \ xx" . have le: "(\\<^sup>+ x\X'. f x) \ (\\<^sup>+ y\\<^bold>E `` X'. g y)" if "X' \ X" for X' proof(cases "X' = X") case True then show ?thesis using sum_eq by simp next case False hence X': "X' \ X" using that by blast have "(\\<^sup>+ x\X'. f x) = (\\<^sup>+ x\X' - {xx}. f x) + (\\<^sup>+ x\{xx}. f x * indicator X' xx)" by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ \ (\\<^sup>+ x\X' - {xx}. f x) + (\\<^sup>+ x\{xx}. weight \ x * indicator X' xx)" using fxx by(intro add_mono)(auto split: split_indicator simp add: nn_integral_count_space_indicator) also have "\ = (\\<^sup>+ x\X'. weight \ x)" using xx that by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] f_def simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ \ (\\<^sup>+ y\\<^bold>E `` X'. weight \ y)" by(rule minimal[OF X']) also have "\ = (\\<^sup>+ y\\<^bold>E `` X'. g y)" using that by(auto 4 3 intro!: nn_integral_cong simp add: g_def Image_iff) finally show ?thesis . qed have "countable X" using X_A A_vertex countable_V by(blast intro: countable_subset) moreover have "\<^bold>E `` X \ \<^bold>V" by(auto simp add: vertex_def) with countable_V have "countable (\<^bold>E `` X)" by(blast intro: countable_subset) moreover have "E' \ X \ \<^bold>E `` X" by(auto simp add: E'_def) ultimately obtain h' where h'_dom: "\x y. 0 < h' x y \ (x, y) \ E'" and h'_fin: "\x y. h' x y \ \" and h'_f: "\x. x \ X \ (\\<^sup>+ y\E' `` X. h' x y) = f x" and h'_g: "\y. y \ E' `` X \ (\\<^sup>+ x\X. h' x y) = g y" using bounded_matrix_for_marginals_ennreal[where f=f and g=g and A=X and B="E' `` X" and R=E' and thesis=thesis] sum_eq fin le by(auto) have h'_outside: "(x, y) \ E' \ h' x y = 0" for x y using h'_dom[of x y] not_gr_zero by(fastforce) define h where "h = (\(x, y). if x \ X \ edge \ x y then h' x y else 0)" have h_OUT: "d_OUT h x = (if x \ X then f x else 0)" for x by(auto 4 3 simp add: h_def d_OUT_def h'_f[symmetric] E'_def nn_integral_count_space_indicator intro!: nn_integral_cong intro: h'_outside split: split_indicator) have h_IN: "d_IN h y = (if y \ \<^bold>E `` X then weight \ y else 0)" for y using h'_g[of y, symmetric] by(auto 4 3 simp add: h_def d_IN_def g_def nn_integral_count_space_indicator nn_integral_0_iff_AE in_E' intro!: nn_integral_cong intro: h'_outside split: split_indicator split_indicator_asm) have h: "current \ h" proof show "d_OUT h x \ weight \ x" for x using fxx by(auto simp add: h_OUT f_def) show "d_IN h y \ weight \ y" for y by(simp add: h_IN) show "h e = 0" if "e \ \<^bold>E" for e using that by(cases e)(auto simp add: h_def) qed have "separating \ (TER h)" proof fix x y p assume x: "x \ A \" and y: "y \ B \" and p: "path \ x p y" then obtain [simp]: "p = [y]" and xy: "(x, y) \ \<^bold>E" using disjoint by -(erule rtrancl_path.cases; auto dest: bipartite_E)+ show "(\z\set p. z \ TER h) \ x \ TER h" proof(rule disjCI) assume "x \ TER h" hence "x \ X" using x by(auto simp add: SAT.simps SINK.simps h_OUT split: if_split_asm) hence "y \ TER h" using xy currentD_OUT[OF h y] by(auto simp add: SAT.simps h_IN SINK.simps) thus "\z\set p. z \ TER h" by simp qed qed then have w: "wave \ h" using h .. have "xx \ A \" using xx X_A by blast moreover have "xx \ \ (TER h)" proof assume "xx \ \ (TER h)" then obtain p y where y: "y \ B \" and p: "path \ xx p y" and bypass: "\z. \ xx \ y; z \ set p \ \ z = xx \ z \ TER h" by(rule \_E) auto from p obtain [simp]: "p = [y]" and xy: "edge \ xx y" and neq: "xx \ y" using disjoint X_A xx y by -(erule rtrancl_path.cases; auto dest: bipartite_E)+ from neq bypass[of y] have "y \ TER h" by simp moreover from xy xx currentD_OUT[OF h y] have "y \ TER h" by(auto simp add: SAT.simps h_IN SINK.simps) ultimately show False by contradiction qed moreover have "d_OUT h xx < weight \ xx" using fxx xx by(simp add: h_OUT) ultimately have "hindrance \ h" .. then show "hindered \" using h w .. qed end lemma nn_integral_count_space_top_approx: fixes f :: "nat => ennreal" and b :: ennreal assumes "nn_integral (count_space UNIV) f = top" and "b < top" obtains n where "b < sum f {.. of_nat x \ 1 \ x" by (metis of_nat_le_iff of_nat_1) locale bounded_countable_bipartite_web = countable_bipartite_web \ for \ :: "('v, 'more) web_scheme" (structure) + assumes bounded_B: "x \ A \ \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y) < \" begin theorem unhindered_linkable_bounded: assumes "\ hindered \" shows "linkable \" proof(cases "A \ = {}") case True hence "linkage \ (\_. 0)" by(auto simp add: linkage.simps) moreover have "web_flow \ (\_. 0)" by(auto simp add: web_flow.simps) ultimately show ?thesis by blast next case nonempty: False define A_n :: "nat \ 'v set" where "A_n n = from_nat_into (A \) ` {..n}" for n have fin_A_n [simp]: "finite (A_n n)" for n by(simp add: A_n_def) have A_n_A: "A_n n \ A \" for n by(auto simp add: A_n_def from_nat_into[OF nonempty]) have countable2: "countable (\<^bold>E `` A_n n)" for n using countable_V by(rule countable_subset[rotated])(auto simp add: vertex_def) have "\Y2. \n. \X \ A_n n. Y2 n X \ \<^bold>E `` X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y2 n X. weight \ y) \ (\\<^sup>+ y\Y2 n X. weight \ y) \ \" proof(rule choice strip ex_simps(6)[THEN iffD2])+ fix n X assume X: "X \ A_n n" then have [simp]: "finite X" by(rule finite_subset) simp have X_count: "countable (\<^bold>E `` X)" using countable2 by(rule countable_subset[rotated])(rule Image_mono[OF order_refl X]) show "\Y. Y \ \<^bold>E `` X \ (\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y. weight \ y) \ (\\<^sup>+ y\Y. weight \ y) \ \" (is "Ex ?P") proof(cases "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = \") case True define Y' where "Y' = to_nat_on (\<^bold>E `` X) ` (\<^bold>E `` X)" have inf: "infinite (\<^bold>E `` X)" using True by(intro notI)(auto simp add: nn_integral_count_space_finite) then have Y': "Y' = UNIV" using X_count by(auto simp add: Y'_def intro!: image_to_nat_on) have "(\\<^sup>+ y\\<^bold>E `` X. weight \ y) = (\\<^sup>+ y\from_nat_into (\<^bold>E `` X) ` Y'. weight \ y * indicator (\<^bold>E `` X) y)" using X_count by(auto simp add: nn_integral_count_space_indicator Y'_def image_image intro!: nn_integral_cong from_nat_into_to_nat_on[symmetric] rev_image_eqI split: split_indicator) also have "\ = (\\<^sup>+ y. weight \ (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" using X_count inf by(subst nn_integral_count_space_reindex)(auto simp add: inj_on_def Y') finally have "\ = \" using True by simp from nn_integral_count_space_top_approx[OF this, of "sum (weight \) X"] obtain yy where yy: "sum (weight \) X < (\ y (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" by(auto simp add: less_top[symmetric]) define Y where "Y = from_nat_into (\<^bold>E `` X) ` {.. \<^bold>E `` X" have [simp]: "finite Y" by(simp add: Y_def) have "(\\<^sup>+ x\X. weight \ x) = sum (weight \) X" by(simp add: nn_integral_count_space_finite) also have "\ \ (\ y (from_nat_into (\<^bold>E `` X) y) * indicator (\<^bold>E `` X) (from_nat_into (\<^bold>E `` X) y))" using yy by simp also have "\ = (\ y \ from_nat_into (\<^bold>E `` X) ` {.. y * indicator (\<^bold>E `` X) y)" using X_count inf by(subst sum.reindex)(auto simp add: inj_on_def) also have "\ = (\ y \ Y. weight \ y)" by(auto intro!: sum.cong simp add: Y_def) also have "\ = (\\<^sup>+ y\Y. weight \ y)" by(simp add: nn_integral_count_space_finite) also have "Y \ \<^bold>E `` X" by(simp add: Y_def) moreover have "(\\<^sup>+ y\Y. weight \ y) \ \" by(simp add: nn_integral_count_space_finite) ultimately show ?thesis by blast next case False with unhindered_criterion[OF assms, of X] X A_n_A[of n] have "?P (\<^bold>E `` X)" by auto then show ?thesis .. qed qed then obtain Y2 where Y2_A: "Y2 n X \ \<^bold>E `` X" and le: "(\\<^sup>+ x\X. weight \ x) \ (\\<^sup>+ y\Y2 n X. weight \ y)" and finY2: "(\\<^sup>+ y\Y2 n X. weight \ y) \ \" if "X \ A_n n" for n X by iprover define Y where "Y n = (\ X \ Pow (A_n n). Y2 n X)" for n define s where "s n = (\\<^sup>+ y\Y n. weight \ y)" for n have Y_vertex: "Y n \ \<^bold>V" for n by(auto 4 3 simp add: Y_def vertex_def dest!: Y2_A[of _ n]) have Y_B: "Y n \ B \" for n unfolding Y_def by(auto dest!: Y2_A[of _ n] dest: bipartite_E) have s_top [simp]: "s n \ \" for n proof - have "\x \ Y2 n X; X \ A_n n\ \ Suc 0 \ card {X. X \ A_n n \ x \ Y2 n X}" for x X by(subst card_le_Suc_iff)(auto intro!: exI[where x=X] exI[where x="{X. X \ A_n n \ x \ Y2 n X} - {X}"]) then have "(\\<^sup>+ y\Y n. weight \ y) \ (\\<^sup>+ y\Y n. \ X\Pow (A_n n). weight \ y * indicator (Y2 n X) y)" by(intro nn_integral_mono)(auto simp add: Y_def One_le_of_nat_ennreal intro!: mult_right_mono[of "1 :: ennreal", simplified]) also have "\ = (\ X\Pow (A_n n). \\<^sup>+ y\Y n. weight \ y * indicator (Y2 n X) y)" by(subst nn_integral_sum) auto also have "\ = (\ X\Pow (A_n n). \\<^sup>+ y\Y2 n X. weight \ y)" by(auto intro!: sum.cong nn_integral_cong simp add: nn_integral_count_space_indicator Y_def split: split_indicator) also have "\ < \" by(simp add: less_top[symmetric] finY2) finally show ?thesis by(simp add: less_top s_def) qed define f :: "nat \ 'v option \ real" where "f n xo = (case xo of Some x \ if x \ A_n n then enn2real (weight \ x) else 0 | None \ enn2real (s n - sum (weight \) (A_n n)))" for n xo define g :: "nat \ 'v \ real" where "g n y = enn2real (weight \ y * indicator (Y n) y)" for n y define R :: "nat \ ('v option \ 'v) set" where "R n = map_prod Some id ` (\<^bold>E \ A_n n \ Y n) \ {None} \ Y n" for n define A_n' where "A_n' n = Some ` A_n n \ {None}" for n have f_simps: "f n (Some x) = (if x \ A_n n then enn2real (weight \ x) else 0)" "f n None = enn2real (s n - sum (weight \) (A_n n))" for n x by(simp_all add: f_def) have g_s: "(\\<^sup>+ y\Y n. g n y) = s n" for n by(auto simp add: s_def g_def ennreal_enn2real_if intro!: nn_integral_cong) have "(\\<^sup>+ x\A_n' n. f n x) = (\\<^sup>+ x\Some`A_n n. weight \ (the x)) + (\\<^sup>+ x\{None}. f n x)" for n by(auto simp add: nn_integral_count_space_indicator nn_integral_add[symmetric] f_simps A_n'_def ennreal_enn2real_if simp del: nn_integral_indicator_singleton intro!: nn_integral_cong split: split_indicator) also have "\ n = sum (weight \) (A_n n) + (s n - sum (weight \) (A_n n))" for n by(subst nn_integral_count_space_reindex)(auto simp add: nn_integral_count_space_finite f_simps ennreal_enn2real_if) also have "\ n = s n" for n using le[OF order_refl, of n] by(simp add: s_def nn_integral_count_space_finite)(auto elim!: order_trans simp add: nn_integral_count_space_indicator Y_def intro!: nn_integral_mono split: split_indicator) finally have sum_eq: "(\\<^sup>+ x\A_n' n. f n x) = (\\<^sup>+ y\Y n. g n y)" for n using g_s by simp have "\h'. \n. (\x y. (x, y) \ R n \ h' n x y = 0) \ (\x y. h' n x y \ \) \ (\x\A_n' n. (\\<^sup>+ y\Y n. h' n x y) = f n x) \ (\y\Y n. (\\<^sup>+ x\A_n' n. h' n x y) = g n y)" (is "Ex (\h'. \n. ?Q n (h' n))") proof(rule choice allI)+ fix n note sum_eq moreover have "(\\<^sup>+ y\Y n. g n y) \ \" using g_s by simp moreover have le_fg: "(\\<^sup>+ x\X. f n x) \ (\\<^sup>+ y\R n `` X. g n y)" if "X \ A_n' n" for X proof(cases "None \ X") case True have "(\\<^sup>+ x\X. f n x) \ (\\<^sup>+ x\A_n' n. f n x)" using that by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = (\\<^sup>+ y\Y n. g n y)" by(simp add: sum_eq) also have "R n `` X = Y n" using True by(auto simp add: R_def) ultimately show ?thesis by simp next case False then have *: "Some ` (the ` X) = X" by(auto simp add: image_image)(metis (no_types, lifting) image_iff notin_range_Some option.sel option.collapse)+ from False that have X: "the ` X \ A_n n" by(auto simp add: A_n'_def) from * have "(\\<^sup>+ x\X. f n x) = (\\<^sup>+ x\Some ` (the ` X). f n x)" by simp also have "\ = (\\<^sup>+ x\the ` X. f n (Some x))" by(rule nn_integral_count_space_reindex) simp also have "\ = (\\<^sup>+ x\the ` X. weight \ x)" using that False by(auto 4 3intro!: nn_integral_cong simp add: f_simps A_n'_def ennreal_enn2real_if) also have "\ \ (\\<^sup>+ y\Y2 n (the ` X). weight \ y)" using False that by(intro le)(auto simp add: A_n'_def) also have "\ \ (\\<^sup>+ y\R n `` X. weight \ y)" using False Y2_A[of "the ` X" n] X by(auto simp add: A_n'_def nn_integral_count_space_indicator R_def Image_iff Y_def intro: rev_image_eqI intro!: nn_integral_mono split: split_indicator) (drule (1) subsetD; clarify; drule (1) bspec; auto 4 3 intro: rev_image_eqI) also have "\ = (\\<^sup>+ y\R n `` X. g n y)" by(auto intro!: nn_integral_cong simp add: R_def g_def ennreal_enn2real_if) finally show ?thesis . qed moreover have "countable (A_n' n)" by(simp add: A_n'_def countable_finite) moreover have "countable (Y2 n X)" if "X \ A_n n" for X using Y2_A[OF that] by(rule countable_subset)(rule countable_subset[OF _ countable_V]; auto simp add: vertex_def) then have "countable (Y n)" unfolding Y_def by(intro countable_UN)(simp_all add: countable_finite) moreover have "R n \ A_n' n \ Y n" by(auto simp add: R_def A_n'_def) ultimately obtain h' where "\x y. 0 < h' x y \ (x, y) \ R n" "\x y. h' x y \ \" "\x. x \ A_n' n \ (\\<^sup>+ y\Y n. h' x y) = (f n x)" "\y. y \ Y n \ (\\<^sup>+ x\A_n' n. h' x y) = g n y" by(rule bounded_matrix_for_marginals_ennreal) blast+ hence "?Q n h'" by(auto)(use not_gr_zero in blast) thus "Ex (?Q n)" by blast qed then obtain h' where dom_h': "\x y. (x, y) \ R n \ h' n x y = 0" and fin_h' [simp]: "\x y. h' n x y \ \" and h'_f: "\x. x \ A_n' n \ (\\<^sup>+ y\Y n. h' n x y) = f n x" and h'_g: "\y. y \ Y n \ (\\<^sup>+ x\A_n' n. h' n x y) = g n y" for n by blast define h :: "nat \ 'v \ 'v \ real" where "h n = (\(x, y). if x \ A_n n \ y \ Y n then enn2real (h' n (Some x) y) else 0)" for n have h_nonneg: "0 \ h n xy" for n xy by(simp add: h_def split_def) have h_notB: "h n (x, y) = 0" if "y \ B \" for n x y using Y_B[of n] that by(auto simp add: h_def) have h_le_weight2: "h n (x, y) \ weight \ y" for n x y proof(cases "x \ A_n n \ y \ Y n") case True have "h' n (Some x) y \ (\\<^sup>+ x\A_n' n. h' n x y)" by(rule nn_integral_ge_point)(auto simp add: A_n'_def True) also have "\ \ weight \ y" using h'_g[of y n] True by(simp add: g_def ennreal_enn2real_if) finally show ?thesis using True by(simp add: h_def ennreal_enn2real_if) qed(auto simp add: h_def) have h_OUT: "d_OUT (h n) x = (if x \ A_n n then weight \ x else 0)" for n x using h'_f[of "Some x" n, symmetric] by(auto simp add: h_def d_OUT_def A_n'_def f_simps ennreal_enn2real_if nn_integral_count_space_indicator intro!: nn_integral_cong) have h_IN: "d_IN (h n) y = (if y \ Y n then enn2real (weight \ y - h' n None y) else 0)" for n y proof(cases "y \ Y n") case True have "d_IN (h n) y = (\\<^sup>+ x\Some ` A_n n. h' n x y)" by(subst nn_integral_count_space_reindex) (auto simp add: d_IN_def h_def nn_integral_count_space_indicator ennreal_enn2real_if R_def intro!: nn_integral_cong dom_h' split: split_indicator) also have "\ = (\\<^sup>+ x\A_n' n. h' n x y) - (\\<^sup>+ x\{None}. h' n x y)" apply(simp add: nn_integral_count_space_indicator del: nn_integral_indicator_singleton) apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space A_n'_def nn_integral_count_space_indicator split: split_indicator intro!: nn_integral_cong) done also have "\ = g n y - h' n None y" using h'_g[OF True] by(simp add: nn_integral_count_space_indicator) finally show ?thesis using True by(simp add: g_def ennreal_enn2real_if) qed(auto simp add: d_IN_def ennreal_enn2real_if nn_integral_0_iff_AE AE_count_space h_def g_def) let ?Q = "\<^bold>V \ \<^bold>V" have "bounded (range (\n. h n xy))" if "xy \ ?Q" for xy unfolding bounded_def dist_real_def proof(rule exI strip|erule imageE|hypsubst)+ fix n obtain x y where [simp]: "xy = (x, y)" by(cases xy) have "h n (x, y) \ d_OUT (h n) x" unfolding d_OUT_def by(rule nn_integral_ge_point) simp also have "\ \ weight \ x" by(simp add: h_OUT) finally show "\0 - h n xy\ \ enn2real (weight \ (fst xy))" by(simp add: h_nonneg)(metis enn2real_ennreal ennreal_cases ennreal_le_iff weight_finite) qed moreover have "countable ?Q" using countable_V by(simp) ultimately obtain k where k: "strict_mono k" and conv: "\xy. xy \ ?Q \ convergent (\n. h (k n) xy)" by(rule convergent_bounded_family) blast+ have h_outside: "h n xy = 0" if "xy \ ?Q" for xy n using that A_n_A[of n] A_vertex Y_vertex by(auto simp add: h_def split: prod.split) have h_outside_AB: "h n (x, y) = 0" if "x \ A \ \ y \ B \" for n x y using that A_n_A[of n] Y_B[of n] by(auto simp add: h_def) have h_outside_E: "h n (x, y) = 0" if "(x, y) \ \<^bold>E" for n x y using that unfolding h_def by(clarsimp)(subst dom_h', auto simp add: R_def) define H where "H xy = lim (\n. h (k n) xy)" for xy have H: "(\n. h (k n) xy) \ H xy" for xy using conv[of xy] unfolding H_def by(cases "xy \ ?Q")(auto simp add: convergent_LIMSEQ_iff h_outside) have H_outside: "H (x, y) = 0" if "x \ A \ \ y \ B \" for x y using that by(simp add: H_def convergent_LIMSEQ_iff h_outside_AB) have H': "(\n. ennreal (h (k n) xy)) \ H xy" for xy using H by(rule tendsto_ennrealI) have H_def': "H xy = lim (\n. ennreal (h (k n) xy))" for xy by (metis H' limI) have H_OUT: "d_OUT H x = weight \ x" if x: "x \ A \" for x proof - let ?w = "\y. if (x, y) \ \<^bold>E then weight \ y else 0" have sum_w: "(\\<^sup>+ y. if edge \ x y then weight \ y else 0) = (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y)" - by(simp add: nn_integral_count_space_indicator indicator_def if_distrib cong: if_cong) + by(simp add: nn_integral_count_space_indicator indicator_def of_bool_def if_distrib cong: if_cong) have "(\n. d_OUT (h (k n)) x) \ d_OUT H x" unfolding d_OUT_def by(rule nn_integral_dominated_convergence[where w="?w"])(use bounded_B x in \simp_all add: AE_count_space H h_outside_E h_le_weight2 sum_w\) moreover define n_x where "n_x = to_nat_on (A \) x" have x': "x \ A_n (k n)" if "n \ n_x" for n using that seq_suble[OF k, of n] x unfolding A_n_def by(intro rev_image_eqI[where x=n_x])(simp_all add: A_n_def n_x_def) have "(\n. d_OUT (h (k n)) x) \ weight \ x" by(intro tendsto_eventually eventually_sequentiallyI[where c="n_x"])(simp add: h_OUT x') ultimately show ?thesis by(rule LIMSEQ_unique) qed then have "linkage \ H" .. moreover have "web_flow \ H" unfolding web_flow_iff proof show "d_OUT H x \ weight \ x" for x by(cases "x \ A \")(simp_all add: H_OUT[unfolded d_OUT_def] H_outside d_OUT_def) show "d_IN H y \ weight \ y" for y proof - have "d_IN H y = (\\<^sup>+ x. liminf (\n. ennreal (h (k n) (x, y))))" unfolding d_IN_def H_def' by(rule nn_integral_cong convergent_liminf_cl[symmetric] convergentI H')+ also have "\ \ liminf (\n. d_IN (h (k n)) y)" unfolding d_IN_def by(rule nn_integral_liminf) simp also have "\ \ liminf (\n. weight \ y)" unfolding h_IN by(rule Liminf_mono)(auto simp add: ennreal_enn2real_if) also have "\ = weight \ y" by(simp add: Liminf_const) finally show "?thesis" . qed show "ennreal (H e) = 0" if "e \ \<^bold>E" for e proof(rule LIMSEQ_unique[OF H']) obtain x y where [simp]: "e = (x, y)" by(cases e) have "ennreal (h (k n) (x, y)) = 0" for n using dom_h'[of "Some x" y "k n"] that by(auto simp add: h_def R_def enn2real_eq_0_iff elim: meta_mp) then show "(\n. ennreal (h (k n) e)) \ 0" using that by(intro tendsto_eventually eventually_sequentiallyI) simp qed qed ultimately show ?thesis by blast qed end subsection \Glueing the reductions together\ locale bounded_countable_web = countable_web \ for \ :: "('v, 'more) web_scheme" (structure) + assumes bounded_out: "x \ \<^bold>V - B \ \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y) < \" begin lemma bounded_countable_bipartite_web_of: "bounded_countable_bipartite_web (bipartite_web_of \)" (is "bounded_countable_bipartite_web ?\") proof - interpret bi: countable_bipartite_web ?\ by(rule countable_bipartite_web_of) show ?thesis proof fix x assume "x \ A ?\" then obtain x' where x: "x = Inl x'" and x': "vertex \ x'" "x' \ B \" by auto have "\<^bold>E\<^bsub>?\\<^esub> `` {x} \ Inr ` ({x'} \ (\<^bold>E `` {x'}))" using x by(auto simp add: bipartite_web_of_def vertex_def split: sum.split_asm) hence "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) \ (\\<^sup>+ y \ \. weight ?\ y)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = (\\<^sup>+ y\{x'} \ (\<^bold>E `` {x'}). weight (bipartite_web_of \) (Inr y))" by(rule nn_integral_count_space_reindex)(auto) also have "\ \ weight \ x' + (\\<^sup>+ y \ \<^bold>E `` {x'}. weight \ y)" apply(subst (1 2) nn_integral_count_space_indicator, simp, simp) apply(cases "\ edge \ x' x'") apply(subst nn_integral_disjoint_pair) apply(auto intro!: nn_integral_mono add_increasing split: split_indicator) done also have "\ < \" using bounded_out[of x'] x' using weight_finite[of x'] by(simp del: weight_finite add: less_top) finally show "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) < \" . qed qed theorem loose_linkable_bounded: assumes "loose \" shows "linkable \" proof - interpret bi: bounded_countable_bipartite_web "bipartite_web_of \" by(rule bounded_countable_bipartite_web_of) have "\ hindered (bipartite_web_of \)" using assms by(rule unhindered_bipartite_web_of) then have "linkable (bipartite_web_of \)" by(rule bi.unhindered_linkable_bounded) then show ?thesis by(rule linkable_bipartite_web_ofD) simp qed lemma bounded_countable_web_quotient_web: "bounded_countable_web (quotient_web \ f)" (is "bounded_countable_web ?\") proof - interpret r: countable_web ?\ by(rule countable_web_quotient_web) show ?thesis proof fix x assume "x \ \<^bold>V\<^bsub>quotient_web \ f\<^esub> - B (quotient_web \ f)" then have x: "x \ \<^bold>V - B \" by(auto dest: vertex_quotient_webD) have "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) \ (\\<^sup>+ y \ \<^bold>E `` {x}. weight \ y)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ < \" using x by(rule bounded_out) finally show "(\\<^sup>+ y \ \<^bold>E\<^bsub>?\\<^esub> `` {x}. weight ?\ y) < \" . qed qed lemma ex_orthogonal_current: "\f S. web_flow \ f \ separating \ S \ orthogonal_current \ f S" proof - from ex_maximal_wave[OF countable] obtain f where f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" by blast from ex_trimming[OF f w countable weight_finite] obtain h where h: "trimming \ f h" .. let ?\ = "quotient_web \ f" interpret \: bounded_countable_web "?\" by(rule bounded_countable_web_quotient_web) have "loose ?\" using f w maximal by(rule loose_quotient_web[OF weight_finite]) then have "linkable ?\" by(rule \.loose_linkable_bounded) then obtain g where wg: "web_flow ?\ g" and link: "linkage ?\ g" by blast let ?k = "plus_current h g" have "web_flow \ ?k" "orthogonal_current \ ?k (\ (TER f))" by(rule linkage_quotient_webD[OF f w wg link h])+ moreover have "separating \ (\ (TER f))" using waveD_separating[OF w] by(rule separating_essential) ultimately show ?thesis by blast qed end locale bounded_countable_network = countable_network \ for \ :: "('v, 'more) network_scheme" (structure) + assumes out: "\ x \ \<^bold>V; x \ source \ \ \ d_OUT (capacity \) x < \" context antiparallel_edges begin lemma \''_bounded_countable_network: "bounded_countable_network \''" if "\x. \ x \ \<^bold>V; x \ source \ \ \ d_OUT (capacity \) x < \" proof - interpret ae: countable_network \'' by(rule \''_countable_network) show ?thesis proof fix x assume x: "x \ \<^bold>V\<^bsub>\''\<^esub>" and not_source: "x \ source \''" from x consider (Vertex) x' where "x = Vertex x'" "x' \ \<^bold>V" | (Edge) y z where "x = Edge y z" "edge \ y z" unfolding "\<^bold>V_\''" by auto then show "d_OUT (capacity \'') x < \" proof cases case Vertex then show ?thesis using x not_source that[of x'] by auto next case Edge then show ?thesis using capacity_finite[of "(y, z)"] by(simp del: capacity_finite add: less_top) qed qed qed end context bounded_countable_network begin lemma bounded_countable_web_web_of_network: assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" shows "bounded_countable_web (web_of_network \)" (is "bounded_countable_web ?\") proof - interpret web: countable_web ?\ by(rule countable_web_web_of_network) fact+ show ?thesis proof fix e assume "e \ \<^bold>V\<^bsub>?\\<^esub> - B ?\" then obtain x y where e: "e = (x, y)" and xy: "edge \ x y" by(cases e) simp from xy have y: "y \ source \" using source_in[of x] by auto have "\<^bold>E\<^bsub>?\\<^esub> `` {e} \ \<^bold>E \ {y} \ UNIV" using e by auto hence "(\\<^sup>+ e' \ \<^bold>E\<^bsub>?\\<^esub> `` {e}. weight ?\ e') \ (\\<^sup>+ e \ \<^bold>E \ {y} \ UNIV. capacity \ e)" using e by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ \ (\\<^sup>+ e \ Pair y ` UNIV. capacity \ e)" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = d_OUT (capacity \) y" unfolding d_OUT_def by(rule nn_integral_count_space_reindex) simp also have "\ < \" using out[of y] xy y by(auto simp add: vertex_def) finally show "(\\<^sup>+ e' \ \<^bold>E\<^bsub>?\\<^esub> `` {e}. weight ?\ e') < \" . qed qed context begin qualified lemma max_flow_min_cut'_bounded: assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" by(rule max_flow_min_cut')(rule bounded_countable_web.ex_orthogonal_current[OF bounded_countable_web_web_of_network], fact+) qualified lemma max_flow_min_cut''_bounded: assumes sink_out: "\y. \ edge \ (sink \) y" and source_in: "\x. \ edge \ x (source \)" and no_loop: "\x. \ edge \ x x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': bounded_countable_network \'' by(rule \''_bounded_countable_network)(rule out) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut'_bounded)(auto simp add: sink_out source_in no_loop capacity_pos elim: edg.cases) then obtain f S where f: "flow \'' f" and cut: "cut \'' S" and ortho: "orthogonal \'' f S" by blast have "flow \ (collect f)" using f by(rule flow_collect) moreover have "cut \ (cut' S)" using cut by(rule cut_cut') moreover have "orthogonal \ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut') ultimately show ?thesis by blast qed qualified lemma max_flow_min_cut'''_bounded: assumes sink_out: "\y. \ edge \ (sink \) y" and source_in: "\x. \ edge \ x (source \)" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': bounded_countable_network \'' by(rule \''_bounded_countable_network)(rule out) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut''_bounded)(auto simp add: sink_out source_in capacity_pos elim: edg.cases) then obtain f S where f: "flow \'' f" and cut: "cut \'' S" and ortho: "orthogonal \'' f S" by blast have "flow \ (collect f)" using f by(rule flow_collect) moreover have "cut \ (cut' S)" using cut by(rule cut_cut') moreover have "orthogonal \ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut') ultimately show ?thesis by blast qed lemma \'''_bounded_countable_network: "bounded_countable_network \'''" proof - interpret \''': countable_network \''' by(rule \'''_countable_network) show ?thesis proof fix x assume x: "x \ \<^bold>V\<^bsub>\'''\<^esub>" and not_source: "x \ source \'''" from x have x': "x \ \<^bold>V" by(auto simp add: vertex_def) have "d_OUT (capacity \''') x \ d_OUT (capacity \) x" by(rule d_OUT_mono) simp also have "\ < \" using x' not_source by(simp add: out) finally show "d_OUT (capacity \''') x < \" . qed qed theorem max_flow_min_cut_bounded: "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret \': bounded_countable_network \''' by(rule \'''_bounded_countable_network) have "\f S. flow \''' f \ cut \''' S \ orthogonal \''' f S" by(rule \'.max_flow_min_cut'''_bounded) auto then obtain f S where f: "flow \''' f" and cut: "cut \''' S" and ortho: "orthogonal \''' f S" by blast from flow_\'''[OF this] show ?thesis by blast qed end end end diff --git a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy --- a/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy +++ b/thys/Prime_Number_Theorem/Prime_Number_Theorem_Library.thy @@ -1,2041 +1,2041 @@ section \Auxiliary material\ theory Prime_Number_Theorem_Library imports Zeta_Function.Zeta_Function "HOL-Real_Asymp.Real_Asymp" begin (* TODO: Move *) lemma asymp_equivD_strong: assumes "f \[F] g" "eventually (\x. f x \ 0 \ g x \ 0) F" shows "((\x. f x / g x) \ 1) F" proof - from assms(1) have "((\x. if f x = 0 \ g x = 0 then 1 else f x / g x) \ 1) F" by (rule asymp_equivD) also have "?this \ ?thesis" by (intro filterlim_cong eventually_mono[OF assms(2)]) auto finally show ?thesis . qed lemma frontier_real_Ici [simp]: fixes a :: real shows "frontier {a..} = {a}" unfolding frontier_def by (auto simp: interior_real_atLeast) lemma sum_upto_ln_conv_sum_upto_mangoldt: "sum_upto (\n. ln (real n)) x = sum_upto (\n. mangoldt n * nat \x / real n\) x" proof - have "sum_upto (\n. ln (real n)) x = sum_upto (\n. \d | d dvd n. mangoldt d) x" by (intro sum_upto_cong) (simp_all add: mangoldt_sum) also have "\ = sum_upto (\k. sum_upto (\d. mangoldt k) (x / real k)) x" by (rule sum_upto_sum_divisors) also have "\ = sum_upto (\n. mangoldt n * nat \x / real n\) x" unfolding sum_upto_altdef by (simp add: mult_ac) finally show ?thesis . qed lemma ln_fact_conv_sum_upto_mangoldt: "ln (fact n) = sum_upto (\k. mangoldt k * (n div k)) n" proof - have [simp]: "{0<..Suc n} = insert (Suc n) {0<..n}" for n by auto have "ln (fact n) = sum_upto (\n. ln (real n)) n" by (induction n) (auto simp: sum_upto_altdef nat_add_distrib ln_mult) also have "\ = sum_upto (\k. mangoldt k * (n div k)) n" unfolding sum_upto_ln_conv_sum_upto_mangoldt by (intro sum_upto_cong) (auto simp: floor_divide_of_nat_eq) finally show ?thesis . qed lemma powr_sum: "x \ 0 \ finite A \ x powr sum f A = (\y\A. x powr f y)" by (simp add: powr_def exp_sum sum_distrib_right) lemma fds_abs_converges_comparison_test: fixes s :: "'a :: dirichlet_series" assumes "eventually (\n. norm (fds_nth f n) \ fds_nth g n) at_top" and "fds_converges g (s \ 1)" shows "fds_abs_converges f s" unfolding fds_abs_converges_def proof (rule summable_comparison_test_ev) from assms(2) show "summable (\n. fds_nth g n / n powr (s \ 1))" by (auto simp: fds_converges_def) from assms(1) eventually_gt_at_top[of 0] show "eventually (\n. norm (norm (fds_nth f n / nat_power n s)) \ fds_nth g n / real n powr (s \ 1)) at_top" by eventually_elim (auto simp: norm_divide norm_nat_power intro!: divide_right_mono) qed lemma fds_converges_scaleR [intro]: assumes "fds_converges f s" shows "fds_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. c *\<^sub>R (fds_nth f n / nat_power n s))" by (intro summable_scaleR_right) (auto simp: fds_converges_def) also have "(\n. c *\<^sub>R (fds_nth f n / nat_power n s)) = (\n. (c *\<^sub>R fds_nth f n / nat_power n s))" by (simp add: scaleR_conv_of_real) finally show ?thesis by (simp add: fds_converges_def) qed lemma fds_abs_converges_scaleR [intro]: assumes "fds_abs_converges f s" shows "fds_abs_converges (c *\<^sub>R f) s" proof - from assms have "summable (\n. abs c * norm (fds_nth f n / nat_power n s))" by (intro summable_mult) (auto simp: fds_abs_converges_def) also have "(\n. abs c * norm (fds_nth f n / nat_power n s)) = (\n. norm ((c *\<^sub>R fds_nth f n) / nat_power n s))" by (simp add: norm_divide) finally show ?thesis by (simp add: fds_abs_converges_def) qed lemma conv_abscissa_scaleR: "conv_abscissa (scaleR c f) \ conv_abscissa f" by (rule conv_abscissa_mono) auto lemma abs_conv_abscissa_scaleR: "abs_conv_abscissa (scaleR c f) \ abs_conv_abscissa f" by (rule abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_left [intro]: "fds_converges f s \ fds_converges (fds_const c * f) s" by (auto simp: fds_converges_def dest: summable_mult[of _ c]) lemma fds_abs_converges_mult_const_left [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_const c * f) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult[of _ "norm c"]) lemma conv_abscissa_mult_const_left: "conv_abscissa (fds_const c * f) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_left: "abs_conv_abscissa (fds_const c * f) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma fds_converges_mult_const_right [intro]: "fds_converges f s \ fds_converges (f * fds_const c) s" by (auto simp: fds_converges_def dest: summable_mult2[of _ c]) lemma fds_abs_converges_mult_const_right [intro]: "fds_abs_converges f s \ fds_abs_converges (f * fds_const c) s" by (auto simp: fds_abs_converges_def norm_mult norm_divide dest: summable_mult2[of _ "norm c"]) lemma conv_abscissa_mult_const_right: "conv_abscissa (f * fds_const c) \ conv_abscissa f" by (intro conv_abscissa_mono) auto lemma abs_conv_abscissa_mult_const_right: "abs_conv_abscissa (f * fds_const c) \ abs_conv_abscissa f" by (intro abs_conv_abscissa_mono) auto lemma bounded_coeffs_imp_fds_abs_converges: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (fds_nth f)" "s \ 1 > 1" shows "fds_abs_converges f s" proof - from assms obtain C where C: "\n. norm (fds_nth f n) \ C" by (auto simp: Bseq_def) show ?thesis proof (rule fds_abs_converges_comparison_test) from \s \ 1 > 1\ show "fds_converges (C *\<^sub>R fds_zeta) (s \ 1)" by (intro fds_abs_converges_imp_converges) auto from C show "eventually (\n. norm (fds_nth f n) \ fds_nth (C *\<^sub>R fds_zeta) n) at_top" by (intro always_eventually) (auto simp: fds_nth_zeta) qed qed lemma bounded_coeffs_imp_fds_abs_converges': fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n * nat_power n s0)" "s \ 1 > 1 - s0 \ 1" shows "fds_abs_converges f s" proof - have "fds_nth (fds_shift s0 f) = (\n. fds_nth f n * nat_power n s0)" by (auto simp: fun_eq_iff) with assms have "Bseq (fds_nth (fds_shift s0 f))" by simp with assms(2) have "fds_abs_converges (fds_shift s0 f) (s + s0)" by (intro bounded_coeffs_imp_fds_abs_converges) (auto simp: algebra_simps) thus ?thesis by simp qed lemma bounded_coeffs_imp_abs_conv_abscissa_le: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" and c :: ereal assumes "Bseq (\n. fds_nth f n * nat_power n s)" "1 - s \ 1 \ c" shows "abs_conv_abscissa f \ c" proof (rule abs_conv_abscissa_leI_weak) fix x assume "c < ereal x" have "ereal (1 - s \ 1) \ c" by fact also have "\ < ereal x" by fact finally have "1 - s \ 1 < ereal x" by simp thus "fds_abs_converges f (of_real x)" by (intro bounded_coeffs_imp_fds_abs_converges'[OF assms(1)]) auto qed lemma bounded_coeffs_imp_abs_conv_abscissa_le_1: fixes s :: "'a :: dirichlet_series" and f :: "'a fds" assumes "Bseq (\n. fds_nth f n)" shows "abs_conv_abscissa f \ 1" proof - have [simp]: "fds_nth f n * nat_power n 0 = fds_nth f n" for n by (cases "n = 0") auto show ?thesis by (rule bounded_coeffs_imp_abs_conv_abscissa_le[where s = 0]) (insert assms, auto simp:) qed (* TODO: replace library version *) (* EXAMPLE: This might make a good example to illustrate real_asymp *) lemma fixes a b c :: real assumes ab: "a + b > 0" and c: "c < -1" shows set_integrable_powr_at_top: "(\x. (b + x) powr c) absolutely_integrable_on {a<..}" and set_lebesgue_integral_powr_at_top: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" and powr_has_integral_at_top: "((\x. (b + x) powr c) has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" proof - let ?f = "\x. (b + x) powr c" and ?F = "\x. (b + x) powr (c + 1) / (c + 1)" have limits: "((?F \ real_of_ereal) \ ?F a) (at_right (ereal a))" "((?F \ real_of_ereal) \ 0) (at_left \)" using c ab unfolding ereal_tendsto_simps1 by (real_asymp simp: field_simps)+ have 1: "set_integrable lborel (einterval a \) ?f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 2: "?f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "LBINT x=ereal a..\. (b + x) powr c = 0 - ?F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: derivative_eq_intros) thus 3: "(\x\{a<..}. ((b + x) powr c) \lborel) = -((b + a) powr (c + 1) / (c + 1))" by (simp add: interval_integral_to_infinity_eq) show "(?f has_integral -((b + a) powr (c + 1) / (c + 1))) {a<..}" using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff) qed lemma fds_converges_altdef2: "fds_converges f s \ convergent (\N. eval_fds (fds_truncate N f) s)" unfolding fds_converges_def summable_iff_convergent' eval_fds_truncate by (auto simp: not_le intro!: convergent_cong always_eventually sum.mono_neutral_right) lemma tendsto_eval_fds_truncate: assumes "fds_converges f s" shows "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s" proof - have "(\N. eval_fds (fds_truncate N f) s) \ eval_fds f s \ (\N. \i\N. fds_nth f i / nat_power i s) \ eval_fds f s" unfolding eval_fds_truncate by (intro filterlim_cong always_eventually allI sum.mono_neutral_left) (auto simp: not_le) also have \ using assms by (simp add: fds_converges_iff sums_def' atLeast0AtMost) finally show ?thesis . qed lemma linepath_translate_left: "linepath (c + a) (c + a) = (\x. c + a) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma linepath_translate_right: "linepath (a + c) (b + c) = (\x. x + c) \ linepath a b" by (auto simp: fun_eq_iff linepath_def algebra_simps) lemma integrable_on_affinity: assumes "m \ 0" "f integrable_on (cbox a b)" shows "(\x. f (m *\<^sub>R x + c)) integrable_on ((\x. (1 / m) *\<^sub>R x - ((1 / m) *\<^sub>R c)) ` cbox a b)" proof - from assms obtain I where "(f has_integral I) (cbox a b)" by (auto simp: integrable_on_def) from has_integral_affinity[OF this assms(1), of c] show ?thesis by (auto simp: integrable_on_def) qed lemma has_integral_cmul_iff: assumes "c \ 0" shows "((\x. c *\<^sub>R f x) has_integral (c *\<^sub>R I)) A \ (f has_integral I) A" using assms has_integral_cmul[of f I A c] has_integral_cmul[of "\x. c *\<^sub>R f x" "c *\<^sub>R I" A "inverse c"] by (auto simp: field_simps) lemma has_integral_affinity': fixes a :: "'a::euclidean_space" assumes "(f has_integral i) (cbox a b)" and "m > 0" shows "((\x. f(m *\<^sub>R x + c)) has_integral (i /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m))" proof (cases "cbox a b = {}") case True hence "(cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) = {}" using \m > 0\ unfolding box_eq_empty by (auto simp: algebra_simps) with True and assms show ?thesis by simp next case False have "((\x. f (m *\<^sub>R x + c)) has_integral (1 / \m\ ^ DIM('a)) *\<^sub>R i) ((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b)" using assms by (intro has_integral_affinity) auto also have "((\x. (1 / m) *\<^sub>R x + - ((1 / m) *\<^sub>R c)) ` cbox a b) = ((\x. - ((1 / m) *\<^sub>R c) + x) ` (\x. (1 / m) *\<^sub>R x) ` cbox a b)" by (simp add: image_image algebra_simps) also have "(\x. (1 / m) *\<^sub>R x) ` cbox a b = cbox ((1 / m) *\<^sub>R a) ((1 / m) *\<^sub>R b)" using \m > 0\ False by (subst image_smult_cbox) simp_all also have "(\x. - ((1 / m) *\<^sub>R c) + x) ` \ = cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)" by (subst cbox_translation [symmetric]) (simp add: field_simps vector_add_divide_simps) finally show ?thesis using \m > 0\ by (simp add: field_simps) qed lemma has_integral_affinity_iff: fixes f :: "'a :: euclidean_space \ 'b :: real_normed_vector" assumes "m > 0" shows "((\x. f (m *\<^sub>R x + c)) has_integral (I /\<^sub>R m ^ DIM('a))) (cbox ((a - c) /\<^sub>R m) ((b - c) /\<^sub>R m)) \ (f has_integral I) (cbox a b)" (is "?lhs = ?rhs") proof assume ?lhs from has_integral_affinity'[OF this, of "1 / m" "-c /\<^sub>R m"] and \m > 0\ show ?rhs by (simp add: vector_add_divide_simps) (simp add: field_simps) next assume ?rhs from has_integral_affinity'[OF this, of m c] and \m > 0\ show ?lhs by simp qed lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma has_contour_integral_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x + Im a * \)) has_integral I) {Re a..Re b}" proof - have deriv: "vector_derivative ((\x. x - Im a * \) \ linepath a b) (at y) = b - a" for y using linepath_translate_right[of a "-Im a * \" b, symmetric] by simp have "(f has_contour_integral I) (linepath a b) \ ((\x. f (x + Im a * \)) has_contour_integral I) (linepath (a - Im a * \) (b - Im a * \))" using linepath_translate_right[of a "-Im a * \" b] deriv by (simp add: has_contour_integral) also have "\ \ ((\x. f (x + Im a * \)) has_integral I) {Re a..Re b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (auto simp: complex_is_Real_iff) finally show ?thesis . qed lemma contour_integrable_linepath_same_Im_iff: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x + Im a * \)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_same_Im_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_same_Im: fixes a b :: complex and f :: "complex \ complex" assumes "Im a = Im b" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (x + Im a * \))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_same_Im_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_same_Im_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemmas [simp del] = div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1 lemma continuous_on_compact_bound: assumes "compact A" "continuous_on A f" obtains B where "B \ 0" "\x. x \ A \ norm (f x) \ B" proof - from assms(2,1) have "compact (f ` A)" by (rule compact_continuous_image) then obtain B where "\x\A. norm (f x) \ B" by (auto dest!: compact_imp_bounded simp: bounded_iff) hence "max B 0 \ 0" and "\x\A. norm (f x) \ max B 0" by auto thus ?thesis using that by blast qed interpretation cis: periodic_fun_simple cis "2 * pi" by standard (simp_all add: complex_eq_iff) lemma open_contains_cbox: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "cbox a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from assms obtain R where R: "R > 0" "ball x R \ A" by (auto simp: open_contains_ball) define r :: real where "r = R / (2 * sqrt DIM('a))" from \R > 0\ have [simp]: "r > 0" by (auto simp: r_def) define d :: 'a where "d = r *\<^sub>R Topology_Euclidean_Space.One" have "cbox (x - d) (x + d) \ A" proof safe fix y assume y: "y \ cbox (x - d) (x + d)" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" by (subst euclidean_dist_l2) (auto simp: L2_set_def) also from y have "sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2) \ sqrt (\i\(Basis::'a set). r\<^sup>2)" by (intro real_sqrt_le_mono sum_mono power_mono) (auto simp: dist_norm d_def cbox_def algebra_simps) also have "\ = sqrt (DIM('a) * r\<^sup>2)" by simp also have "DIM('a) * r\<^sup>2 = (R / 2) ^ 2" by (simp add: r_def power_divide) also have "sqrt \ = R / 2" using \R > 0\ by simp also from \R > 0\ have "\ < R" by simp finally have "y \ ball x R" by simp with R show "y \ A" by blast qed thus ?thesis using that[of "x - d" "x + d"] by (auto simp: algebra_simps d_def box_def) qed lemma open_contains_box: fixes x :: "'a :: euclidean_space" assumes "open A" "x \ A" obtains a b where "box a b \ A" "x \ box a b" "\i\Basis. a \ i < b \ i" proof - from open_contains_cbox[OF assms] guess a b . with that[of a b] box_subset_cbox[of a b] show ?thesis by auto qed lemma analytic_onE_box: assumes "f analytic_on A" "s \ A" obtains a b where "Re a < Re b" "Im a < Im b" "s \ box a b" "f analytic_on box a b" proof - from assms obtain r where r: "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) with open_contains_box[of "ball s r" s] obtain a b where "box a b \ ball s r" "s \ box a b" "\i\Basis. a \ i < b \ i" by auto moreover from r have "f analytic_on ball s r" by (simp add: analytic_on_open) ultimately show ?thesis using that[of a b] analytic_on_subset[of _ "ball s r" "box a b"] by (auto simp: Basis_complex_def) qed lemma inner_image_box: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i < b \ i" shows "(\x. x \ i) ` box a b = {a \ i<.. i}" proof safe fix x assume x: "x \ {a \ i<.. i}" let ?y = "(\j\Basis. (if i = j then x else (a + b) \ j / 2) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` box a b" by (intro imageI) (auto simp: box_def algebra_simps) also have "?y \ i = (\j\Basis. (if i = j then x else (a + b) \ j / 2) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` box a b" . qed (insert assms, auto simp: box_def) lemma Re_image_box: assumes "Re a < Re b" "Im a < Im b" shows "Re ` box a b = {Re a<..::complex" a b] assms by (auto simp: Basis_complex_def) lemma inner_image_cbox: assumes "(i :: 'a :: euclidean_space) \ Basis" assumes "\i\Basis. a \ i \ b \ i" shows "(\x. x \ i) ` cbox a b = {a \ i..b \ i}" proof safe fix x assume x: "x \ {a \ i..b \ i}" let ?y = "(\j\Basis. (if i = j then x else a \ j) *\<^sub>R j)" from x assms have "?y \ i \ (\x. x \ i) ` cbox a b" by (intro imageI) (auto simp: cbox_def) also have "?y \ i = (\j\Basis. (if i = j then x else a \ j) * (j \ i))" by (simp add: inner_sum_left) also have "\ = (\j\Basis. if i = j then x else 0)" by (intro sum.cong) (auto simp: inner_not_same_Basis assms) also have "\ = x" using assms by simp finally show "x \ (\x. x \ i) ` cbox a b" . qed (insert assms, auto simp: cbox_def) lemma Re_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Re ` cbox a b = {Re a..Re b}" using inner_image_cbox[of "1::complex" a b] assms by (auto simp: Basis_complex_def) lemma Im_image_cbox: assumes "Re a \ Re b" "Im a \ Im b" shows "Im ` cbox a b = {Im a..Im b}" using inner_image_cbox[of "\::complex" a b] assms by (auto simp: Basis_complex_def) lemma analytic_onE_cball: assumes "f analytic_on A" "s \ A" "ub > (0::real)" obtains R where "R > 0" "R < ub" "f analytic_on cball s R" proof - from assms obtain r where "r > 0" "f holomorphic_on ball s r" by (auto simp: analytic_on_def) hence "f analytic_on ball s r" by (simp add: analytic_on_open) hence "f analytic_on cball s (min (ub / 2) (r / 2))" by (rule analytic_on_subset, subst cball_subset_ball_iff) (use \r > 0\ in auto) moreover have "min (ub / 2) (r / 2) > 0" and "min (ub / 2) (r / 2) < ub" using \r > 0\ and \ub > 0\ by (auto simp: min_def) ultimately show ?thesis using that[of "min (ub / 2) (r / 2)"] by blast qed corollary analytic_pre_zeta' [analytic_intros]: assumes "f analytic_on A" "a > 0" shows "(\x. pre_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_pre_zeta[of a UNIV]] assms(2) by (auto simp: o_def) corollary analytic_hurwitz_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" "a > 0" shows "(\x. hurwitz_zeta a (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_hurwitz_zeta[of a "-{1}"]] assms(2,3) by (auto simp: o_def) corollary analytic_zeta' [analytic_intros]: assumes "f analytic_on A" "(\x. x \ A \ f x \ 1)" shows "(\x. zeta (f x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_zeta[of "-{1}"]] assms(2) by (auto simp: o_def) lemma logderiv_zeta_analytic: "(\s. deriv zeta s / zeta s) analytic_on {s. Re s \ 1} - {1}" using zeta_Re_ge_1_nonzero by (auto intro!: analytic_intros) lemma cis_pi_half [simp]: "cis (pi / 2) = \" by (simp add: complex_eq_iff) lemma mult_real_sqrt: "x \ 0 \ x * sqrt y = sqrt (x ^ 2 * y)" by (simp add: real_sqrt_mult) lemma arcsin_pos: "x \ {0<..1} \ arcsin x > 0" using arcsin_less_arcsin[of 0 x] by simp lemmas analytic_imp_holomorphic' = holomorphic_on_subset[OF analytic_imp_holomorphic] lemma residue_simple': assumes "open s" "0 \ s" "f holomorphic_on s" shows "residue (\w. f w / w) 0 = f 0" using residue_simple[of s 0 f] assms by simp lemma fds_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_converges f s \ fds_converges g s'" unfolding fds_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma fds_abs_converges_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" "s = s'" shows "fds_abs_converges f s \ fds_abs_converges g s'" unfolding fds_abs_converges_def by (intro summable_cong eventually_mono[OF assms(1)]) (simp_all add: assms) lemma conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "conv_abscissa f = conv_abscissa g" proof - have "fds_converges f = fds_converges g" by (intro ext fds_converges_cong assms refl) thus ?thesis by (simp add: conv_abscissa_def) qed lemma abs_conv_abscissa_cong: assumes "eventually (\n. fds_nth f n = fds_nth g n) at_top" shows "abs_conv_abscissa f = abs_conv_abscissa g" proof - have "fds_abs_converges f = fds_abs_converges g" by (intro ext fds_abs_converges_cong assms refl) thus ?thesis by (simp add: abs_conv_abscissa_def) qed definition fds_remainder where "fds_remainder m = fds_subseries (\n. n > m)" lemma fds_nth_remainder: "fds_nth (fds_remainder m f) = (\n. if n > m then fds_nth f n else 0)" by (simp add: fds_remainder_def fds_subseries_def fds_nth_fds') lemma fds_converges_remainder_iff [simp]: "fds_converges (fds_remainder m f) s \ fds_converges f s" by (intro fds_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_abs_converges_remainder_iff [simp]: "fds_abs_converges (fds_remainder m f) s \ fds_abs_converges f s" by (intro fds_abs_converges_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma fds_converges_remainder [intro]: "fds_converges f s \ fds_converges (fds_remainder m f) s" and fds_abs_converges_remainder [intro]: "fds_abs_converges f s \ fds_abs_converges (fds_remainder m f) s" by simp_all lemma conv_abscissa_remainder [simp]: "conv_abscissa (fds_remainder m f) = conv_abscissa f" by (intro conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma abs_conv_abscissa_remainder [simp]: "abs_conv_abscissa (fds_remainder m f) = abs_conv_abscissa f" by (intro abs_conv_abscissa_cong eventually_mono[OF eventually_gt_at_top[of m]]) (auto simp: fds_nth_remainder) lemma eval_fds_remainder: "eval_fds (fds_remainder m f) s = (\n. fds_nth f (n + Suc m) / nat_power (n + Suc m) s)" (is "_ = suminf (\n. ?f (n + Suc m))") proof (cases "fds_converges f s") case False hence "\fds_converges (fds_remainder m f) s" by simp hence "(\x. (\n. fds_nth (fds_remainder m f) n / nat_power n s) sums x) = (\_. False)" by (auto simp: fds_converges_def summable_def) hence "eval_fds (fds_remainder m f) s = (THE _. False)" by (simp add: eval_fds_def suminf_def) moreover from False have "\summable (\n. ?f (n + Suc m))" unfolding fds_converges_def by (subst summable_iff_shift) auto hence "(\x. (\n. ?f (n + Suc m)) sums x) = (\_. False)" by (auto simp: summable_def) hence "suminf (\n. ?f (n + Suc m)) = (THE _. False)" by (simp add: suminf_def) ultimately show ?thesis by simp next case True hence *: "fds_converges (fds_remainder m f) s" by simp have "eval_fds (fds_remainder m f) s = (\n. fds_nth (fds_remainder m f) n / nat_power n s)" unfolding eval_fds_def .. also have "\ = (\n. fds_nth (fds_remainder m f) (n + Suc m) / nat_power (n + Suc m) s)" using * unfolding fds_converges_def by (subst suminf_minus_initial_segment) (auto simp: fds_nth_remainder) also have "(\n. fds_nth (fds_remainder m f) (n + Suc m)) = (\n. fds_nth f (n + Suc m))" by (intro ext) (auto simp: fds_nth_remainder) finally show ?thesis . qed lemma fds_truncate_plus_remainder: "fds_truncate m f + fds_remainder m f = f" by (intro fds_eqI) (auto simp: fds_truncate_def fds_remainder_def fds_subseries_def) lemma holomorphic_fds_eval' [holomorphic_intros]: assumes "g holomorphic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) holomorphic_on A" using holomorphic_on_compose_gen[OF assms(1) holomorphic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma analytic_fds_eval' [analytic_intros]: assumes "g analytic_on A" "\x. x \ A \ Re (g x) > conv_abscissa f" shows "(\x. eval_fds f (g x)) analytic_on A" using analytic_on_compose_gen[OF assms(1) analytic_fds_eval[OF order.refl, of f]] assms(2) by (auto simp: o_def) lemma homotopic_loopsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" "h ` ({0..1} \ {0..1}) \ s" "\x. x \ {0..1} \ h (0, x) = p x" "\x. x \ {0..1} \ h (1, x) = q x" "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathstart (h \ Pair x)" shows "homotopic_loops s p q" using assms unfolding homotopic_loops by (intro exI[of _ h]) auto lemma continuous_on_linepath [continuous_intros]: assumes "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. linepath (a x) (b x) (f x))" using assms by (auto simp: linepath_def intro!: continuous_intros assms) lemma continuous_on_part_circlepath [continuous_intros]: assumes "continuous_on A c" "continuous_on A r" "continuous_on A a" "continuous_on A b" "continuous_on A f" shows "continuous_on A (\x. part_circlepath (c x) (r x) (a x) (b x) (f x))" using assms by (auto simp: part_circlepath_def intro!: continuous_intros assms) lemma homotopic_loops_part_circlepath: assumes "sphere c r \ A" and "r \ 0" and "b1 = a1 + 2 * of_int k * pi" and "b2 = a2 + 2 * of_int k * pi" shows "homotopic_loops A (part_circlepath c r a1 b1) (part_circlepath c r a2 b2)" proof - define h where "h = (\(x,y). part_circlepath c r (linepath a1 a2 x) (linepath b1 b2 x) y)" show ?thesis proof (rule homotopic_loopsI) show "continuous_on ({0..1} \ {0..1}) h" by (auto simp: h_def case_prod_unfold intro!: continuous_intros) next from assms have "h ` ({0..1} \ {0..1}) \ sphere c r" by (auto simp: h_def part_circlepath_def dist_norm norm_mult) also have "\ \ A" by fact finally show "h ` ({0..1} \ {0..1}) \ A" . next fix x :: real assume x: "x \ {0..1}" show "h (0, x) = part_circlepath c r a1 b1 x" and "h (1, x) = part_circlepath c r a2 b2 x" by (simp_all add: h_def linepath_def) have "cis (pi * (real_of_int k * 2)) = 1" using cis.plus_of_int[of 0 k] by (simp add: algebra_simps) thus "pathfinish (h \ Pair x) = pathstart (h \ Pair x)" by (simp add: h_def o_def exp_eq_polar linepath_def algebra_simps cis_mult [symmetric] cis_divide [symmetric] assms) qed qed lemma homotopic_pathsI: fixes h :: "real \ real \ _" assumes "continuous_on ({0..1} \ {0..1}) h" assumes "h ` ({0..1} \ {0..1}) \ s" assumes "\x. x \ {0..1} \ h (0, x) = p x" assumes "\x. x \ {0..1} \ h (1, x) = q x" assumes "\x. x \ {0..1} \ pathstart (h \ Pair x) = pathstart p" assumes "\x. x \ {0..1} \ pathfinish (h \ Pair x) = pathfinish p" shows "homotopic_paths s p q" using assms unfolding homotopic_paths by (intro exI[of _ h]) auto lemma part_circlepath_conv_subpath: "part_circlepath c r a b = subpath (a / (2*pi)) (b / (2*pi)) (circlepath c r)" by (simp add: part_circlepath_def circlepath_def subpath_def linepath_def algebra_simps exp_eq_polar) lemma homotopic_paths_part_circlepath: assumes "a \ b" "b \ c" assumes "path_image (part_circlepath C r a c) \ A" "r \ 0" shows "homotopic_paths A (part_circlepath C r a c) (part_circlepath C r a b +++ part_circlepath C r b c)" (is "homotopic_paths _ ?g (?h1 +++ ?h2)") proof (cases "a = c") case False with assms have "a < c" by simp define slope where "slope = (b - a) / (c - a)" from assms and \a < c\ have slope: "slope \ {0..1}" by (auto simp: field_simps slope_def) define f :: "real \ real" where "f = linepath 0 slope +++ linepath slope 1" show ?thesis proof (rule homotopic_paths_reparametrize) fix t :: real assume t: "t \ {0..1}" show "(?h1 +++ ?h2) t = ?g (f t)" proof (cases "t \ 1 / 2") case True hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from True \a < c\ have "(1 - f t) * a + f t * c = (1 - 2 * t) * a + 2 * t * b" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from True have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. next case False hence "?g (f t) = C + r * cis ((1 - f t) * a + f t * c)" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) also from False \a < c\ have "(1 - f t) * a + f t * c = (2 - 2 * t) * b + (2 * t - 1) * c" unfolding f_def slope_def linepath_def joinpaths_def by (simp add: divide_simps del: div_mult_self3 div_mult_self4 div_mult_self2 div_mult_self1) (simp add: algebra_simps)? also from False have "C + r * cis \ = (?h1 +++ ?h2) t" by (simp add: joinpaths_def part_circlepath_def exp_eq_polar linepath_def) finally show ?thesis .. qed next from slope have "path_image f \ {0..1}" by (auto simp: f_def path_image_join closed_segment_eq_real_ivl) thus "f ` {0..1} \ {0..1}" by (simp add: path_image_def) next have "path f" unfolding f_def by auto thus "continuous_on {0..1} f" by (simp add: path_def) qed (insert assms, auto simp: f_def joinpaths_def linepath_def) next case [simp]: True with assms have [simp]: "b = c" by auto have "part_circlepath C r c c +++ part_circlepath C r c c = part_circlepath C r c c" by (simp add: fun_eq_iff joinpaths_def part_circlepath_def) thus ?thesis using assms by simp qed lemma has_contour_integral_mirror_iff: assumes "valid_path g" shows "(f has_contour_integral I) (-g) \ ((\x. -f (- x)) has_contour_integral I) g" proof - from assms have "g piecewise_differentiable_on {0..1}" by (auto simp: valid_path_def piecewise_C1_imp_differentiable) then obtain S where S: "finite S" "\x. x \ {0..1} - S \ g differentiable at x within {0..1}" unfolding piecewise_differentiable_on_def by blast have S': "g differentiable at x" if "x \ {0..1} - ({0, 1} \ S)" for x proof - from that have "x \ interior {0..1}" by auto with S(2)[of x] that show ?thesis by (auto simp: at_within_interior[of _ "{0..1}"]) qed have "(f has_contour_integral I) (-g) \ ((\x. f (- g x) * vector_derivative (-g) (at x)) has_integral I) {0..1}" by (simp add: has_contour_integral) also have "\ \ ((\x. -f (- g x) * vector_derivative g (at x)) has_integral I) {0..1}" by (intro has_integral_spike_finite_eq[of "S \ {0, 1}"]) (insert \finite S\ S', auto simp: o_def fun_Compl_def) also have "\ \ ((\x. -f (-x)) has_contour_integral I) g" by (simp add: has_contour_integral) finally show ?thesis . qed lemma contour_integral_on_mirror_iff: assumes "valid_path g" shows "f contour_integrable_on (-g) \ (\x. -f (- x)) contour_integrable_on g" by (auto simp: contour_integrable_on_def has_contour_integral_mirror_iff assms) lemma contour_integral_mirror: assumes "valid_path g" shows "contour_integral (-g) f = contour_integral g (\x. -f (- x))" proof (cases "f contour_integrable_on (-g)") case True then obtain I where I: "(f has_contour_integral I) (-g)" by (auto simp: contour_integrable_on_def) also note has_contour_integral_mirror_iff[OF assms] finally have "((\x. - f (- x)) has_contour_integral I) g" . with I show ?thesis using contour_integral_unique by blast next case False hence "\(\x. -f (-x)) contour_integrable_on g" by (auto simp: contour_integral_on_mirror_iff assms) from False and this show ?thesis by (simp add: not_integrable_contour_integral) qed lemma contour_integrable_neg_iff: "(\x. -f x) contour_integrable_on g \ f contour_integrable_on g" using contour_integrable_neg[of f g] contour_integrable_neg[of "\x. -f x" g] by auto lemma contour_integral_neg: shows "contour_integral g (\x. -f x) = -contour_integral g f" proof (cases "f contour_integrable_on g") case True thus ?thesis by (simp add: contour_integral_neg) next case False hence "\(\x. -f x) contour_integrable_on g" by (simp add: contour_integrable_neg_iff) with False show ?thesis by (simp add: not_integrable_contour_integral) qed lemma minus_cis: "-cis x = cis (x + pi)" by (simp add: complex_eq_iff) lemma path_image_part_circlepath_subset: assumes "a \ a'" "a' \ b'" "b' \ b" shows "path_image (part_circlepath c r a' b') \ path_image (part_circlepath c r a b)" using assms by (subst (1 2) path_image_part_circlepath) auto lemma part_circlepath_mirror: assumes "a' = a + pi + 2 * pi * of_int k" "b' = b + pi + 2 * pi * of_int k" "c' = -c" shows "-part_circlepath c r a b = part_circlepath c' r a' b'" proof fix x :: real have "part_circlepath c' r a' b' x = c' + r * cis (linepath a b x + pi + k * (2 * pi))" by (simp add: part_circlepath_def exp_eq_polar assms linepath_translate_right mult_ac) also have "cis (linepath a b x + pi + k * (2 * pi)) = cis (linepath a b x + pi)" by (rule cis.plus_of_int) also have "\ = -cis (linepath a b x)" by (simp add: minus_cis) also have "c' + r * \ = -part_circlepath c r a b x" by (simp add: part_circlepath_def assms exp_eq_polar) finally show "(- part_circlepath c r a b) x = part_circlepath c' r a' b' x" by simp qed lemma path_mirror [intro]: "path (g :: _ \ 'b::topological_group_add) \ path (-g)" by (auto simp: path_def intro!: continuous_intros) lemma path_mirror_iff [simp]: "path (-g :: _ \ 'b::topological_group_add) \ path g" using path_mirror[of g] path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma valid_path_mirror [intro]: "valid_path g \ valid_path (-g)" by (auto simp: valid_path_def fun_Compl_def piecewise_C1_differentiable_neg) lemma valid_path_mirror_iff [simp]: "valid_path (-g) \ valid_path g" using valid_path_mirror[of g] valid_path_mirror[of "-g"] by (auto simp: fun_Compl_def) lemma pathstart_mirror [simp]: "pathstart (-g) = -pathstart g" and pathfinish_mirror [simp]: "pathfinish (-g) = -pathfinish g" by (simp_all add: pathstart_def pathfinish_def) lemma path_image_mirror: "path_image (-g) = uminus ` path_image g" by (auto simp: path_image_def) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma contour_integral_spike_finite_simple_path: assumes "finite A" "simple_path g" "g = g'" "\x. x \ path_image g - A \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral proof (rule integral_spike) have "finite (g -` A \ {0<..<1})" using \simple_path g\ \finite A\ by (intro finite_vimage_IntI simple_path_inj_on) auto hence "finite ({0, 1} \ g -` A \ {0<..<1})" by auto thus "negligible ({0, 1} \ g -` A \ {0<..<1})" by (rule negligible_finite) next fix x assume "x \ {0..1} - ({0, 1} \ g -` A \ {0<..<1})" hence "g x \ path_image g - A" by (auto simp: path_image_def) from assms(4)[OF this] and assms(3) show "f' (g' x) * vector_derivative g' (at x) = f (g x) * vector_derivative g (at x)" by simp qed proposition contour_integral_bound_part_circlepath_strong: assumes fi: "f contour_integrable_on part_circlepath z r s t" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod (contour_integral (part_circlepath z r s t) f) \ B * r * (t - s)" proof - from fi have "(f has_contour_integral contour_integral (part_circlepath z r s t) f) (part_circlepath z r s t)" by (rule has_contour_integral_integral) from has_contour_integral_bound_part_circlepath_strong[OF this assms(2-)] show ?thesis by auto qed lemma cos_le_zero: assumes "x \ {pi/2..3*pi/2}" shows "cos x \ 0" proof - have "cos x = -cos (x - pi)" by (simp add: cos_diff) moreover from assms have "cos (x - pi) \ 0" by (intro cos_ge_zero) auto ultimately show ?thesis by simp qed lemma cos_le_zero': "x \ {-3*pi/2..-pi/2} \ cos x \ 0" using cos_le_zero[of "-x"] by simp lemma cis_minus_pi_half [simp]: "cis (- (pi / 2)) = -\" by (simp add: complex_eq_iff) lemma winding_number_join_pos_combined': "\valid_path \1 \ z \ path_image \1 \ 0 < Re (winding_number \1 z); valid_path \2 \ z \ path_image \2 \ 0 < Re (winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) lemma Union_atLeastAtMost_real_of_nat: assumes "a < b" shows "(\n\{a.. {real a..real b}" thus "x \ (\n\{a.. real a" "x < real b" by simp_all hence "x \ real (nat \x\)" "x \ real (Suc (nat \x\))" by linarith+ moreover from x have "nat \x\ \ a" "nat \x\ < b" by linarith+ ultimately have "\n\{a.. {real n..real (n + 1)}" by (intro bexI[of _ "nat \x\"]) simp_all thus ?thesis by blast qed qed auto lemma nat_sum_has_integral_floor: fixes f :: "nat \ 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m..i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Inf X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Inf X\" if "x \ X - {Sup X}" for x using that X by (auto simp: D_def nat_eq_iff floor_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X \ ((\x. f (nat \Inf X\)) has_integral f (nat \Inf X\)) X" using X by (intro has_integral_spike_eq[of "{Sup X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Inf X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Inf X\)) X" . qed fact+ also have "(\X\D. f (nat \Inf X\)) = (\k\{m.. 'a :: banach" assumes mn: "m < n" shows "((\x. f (nat \x\)) has_integral sum f {m<..n}) {real m..real n}" proof - define D where "D = (\i. {real i..real (Suc i)}) ` {m..x. f (nat \x\)) has_integral (\X\D. f (nat \Sup X\))) {real m..real n}" proof (rule has_integral_combine_division) fix X assume X: "X \ D" have "nat \x\ = nat \Sup X\" if "x \ X - {Inf X}" for x using that X by (auto simp: D_def nat_eq_iff ceiling_eq_iff) hence "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X \ ((\x. f (nat \Sup X\)) has_integral f (nat \Sup X\)) X" using X by (intro has_integral_spike_eq[of "{Inf X}"]) auto also from X have "\" using has_integral_const_real[of "f (nat \Sup X\)" "Inf X" "Sup X"] by (auto simp: D_def) finally show "((\x. f (nat \x\)) has_integral f (nat \Sup X\)) X" . qed fact+ also have "(\X\D. f (nat \Sup X\)) = (\k\{m.. = (\k\{m<..n}. f k)" by (intro sum.reindex_bij_witness[of _ "\x. x - 1" Suc]) auto finally show ?thesis . qed lemma zeta_partial_sum_le: fixes x :: real and m :: nat assumes x: "x \ {0<..1}" shows "(\k=1..m. real k powr (x - 1)) \ real m powr x / x" proof - consider "m = 0" | "m = 1" | "m > 1" by force thus ?thesis proof cases assume m: "m > 1" hence "{1..m} = insert 1 {1<..m}" by auto also have "(\k\\. real k powr (x - 1)) = 1 + (\k\{1<..m}. real k powr (x - 1))" by simp also have "(\k\{1<..m}. real k powr (x - 1)) \ real m powr x / x - 1 / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{1<..m}. n powr (x - 1))) {real 1..m}" using m by (intro nat_sum_has_integral_ceiling) auto next have "((\t. t powr (x - 1)) has_integral (real m powr x / x - real 1 powr x / x)) {real 1..real m}" by (intro fundamental_theorem_of_calculus) (insert x m, auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: derivative_eq_intros) thus "((\t. t powr (x - 1)) has_integral (real m powr x / x - 1 / x)) {real 1..real m}" by simp qed (insert x, auto intro!: powr_mono2') also have "1 + (real m powr x / x - 1 / x) \ real m powr x / x" using x by (simp add: field_simps) finally show ?thesis by simp qed (use assms in auto) qed lemma zeta_partial_sum_le': fixes x :: real and m :: nat assumes x: "x > 0" and m: "m > 0" shows "(\n=1..m. real n powr (x - 1)) \ m powr x * (1 / x + 1 / m)" proof (cases "x > 1") case False with assms have "(\n=1..m. real n powr (x - 1)) \ m powr x / x" by (intro zeta_partial_sum_le) auto also have "\ \ m powr x * (1 / x + 1 / m)" using assms by (simp add: field_simps) finally show ?thesis . next case True have "(\n\{1..m}. n powr (x - 1)) = (\n\insert m {0.. = m powr (x - 1) + (\n\{0..n\{0.. real m powr x / x" proof (rule has_integral_le) show "((\t. (nat \t\) powr (x - 1)) has_integral (\n\{0..t. t powr (x - 1)) has_integral (real m powr x / x)) {real 0..real m}" using has_integral_powr_from_0[of "x - 1"] x by auto next fix t assume "t \ {real 0..real m}" with \x > 1\ show "real (nat \t\) powr (x - 1) \ t powr (x - 1)" by (cases "t = 0") (auto intro: powr_mono2) qed also have "m powr (x - 1) + m powr x / x = m powr x * (1 / x + 1 / m)" using m x by (simp add: powr_diff field_simps) finally show ?thesis by simp qed lemma natfun_bigo_1E: assumes "(f :: nat \ _) \ O(\_. 1)" obtains C where "C \ lb" "\n. norm (f n) \ C" proof - from assms obtain C N where "\n\N. norm (f n) \ C" by (auto elim!: landau_o.bigE simp: eventually_at_top_linorder) hence *: "norm (f n) \ Max ({C, lb} \ (norm ` f ` {.. N") (subst Max_ge_iff; force simp: image_iff)+ moreover have "Max ({C, lb} \ (norm ` f ` {.. lb" by (intro Max.coboundedI) auto ultimately show ?thesis using that by blast qed lemma natfun_bigo_iff_Bseq: "f \ O(\_. 1) \ Bseq f" proof assume "Bseq f" then obtain C where "C > 0" "\n. norm (f n) \ C" by (auto simp: Bseq_def) thus "f \ O(\_. 1)" by (intro bigoI[of _ C]) auto next assume "f \ O(\_. 1)" from natfun_bigo_1E[OF this, where lb = 1] obtain C where "C \ 1" "\n. norm (f n) \ C" by auto thus "Bseq f" by (auto simp: Bseq_def intro!: exI[of _ C]) qed lemma enn_decreasing_sum_le_set_nn_integral: fixes f :: "real \ ennreal" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" shows "(\n. f (real (Suc n))) \ set_nn_integral lborel {0..} f" proof - have "(\n. (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def) also have "\ \ (\\<^sup>+x\{0..}. (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def intro!: decreasing) finally show ?thesis . qed (* TODO replace version in library *) lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" - using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) + using nonneg by (simp add: indicator_def of_bool_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto also have "(\x. abs (indicator \ x * f x)) = (\x. indicator \ x * f x)" using nonneg by (auto simp: indicator_def fun_eq_iff) finally show ?thesis . qed lemma decreasing_sum_le_integral: fixes f :: "real \ real" assumes nonneg: "\x. x \ 0 \ f x \ 0" assumes decreasing: "\x y. 0 \ x \ x \ y \ f y \ f x" assumes integral: "(f has_integral I) {0..}" shows "summable (\i. f (real (Suc i)))" and "suminf (\i. f (real (Suc i))) \ I" proof - have [simp]: "I \ 0" by (intro has_integral_nonneg[OF integral] nonneg) auto have "(\n. ennreal (f (Suc n))) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel)" by (subst nn_integral_cmult_indicator) auto also have "nat \x\ = Suc n" if "x \ {real n<..real (Suc n)}" for x n using that by (auto simp: nat_eq_iff ceiling_eq_iff) hence "(\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (Suc n)) \lborel) = (\n. \\<^sup>+x\{real n<..real (Suc n)}. ennreal (f (real (nat \x\))) \lborel)" by (intro suminf_cong nn_integral_cong) (auto simp: indicator_def) also have "\ = (\\<^sup>+x\(\i. {real i<..real (Suc i)}). ennreal (f (nat \x::real\)) \lborel)" by (subst nn_integral_disjoint_family) (auto simp: disjoint_family_on_def intro!: measurable_completion) also have "\ \ (\\<^sup>+x\{0..}. ennreal (f x) \lborel)" by (intro nn_integral_mono) (auto simp: indicator_def nonneg intro!: decreasing) also have "\ = (\\<^sup>+ x. ennreal (indicat_real {0..} x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using nn_integral_has_integral_lebesgue[OF nonneg integral] by (auto simp: nonneg) finally have *: "(\n. ennreal (f (Suc n))) \ ennreal I" . from * show summable: "summable (\i. f (real (Suc i)))" by (intro summable_suminf_not_top) (auto simp: top_unique intro: nonneg) note * also from summable have "(\n. ennreal (f (Suc n))) = ennreal (\n. f (Suc n))" by (subst suminf_ennreal2) (auto simp: o_def nonneg) finally show "(\n. f (real (Suc n))) \ I" by (subst (asm) ennreal_le_iff) auto qed lemma decreasing_sum_le_integral': fixes f :: "real \ real" assumes "\x. x \ 0 \ f x \ 0" assumes "\x y. 0 \ x \ x \ y \ f y \ f x" assumes "(f has_integral I) {0..}" shows "summable (\i. f (real i))" and "suminf (\i. f (real i)) \ f 0 + I" proof - have "summable ((\i. f (real (Suc i))))" using decreasing_sum_le_integral[OF assms] by (simp add: o_def) thus *: "summable (\i. f (real i))" by (subst (asm) summable_Suc_iff) have "(\n. f (real (Suc n))) \ I" by (intro decreasing_sum_le_integral assms) thus "suminf (\i. f (real i)) \ f 0 + I" using * by (subst (asm) suminf_split_head) auto qed lemma norm_suminf_le: assumes "\n. norm (f n :: 'a :: banach) \ g n" "summable g" shows "norm (suminf f) \ suminf g" proof - have *: "summable (\n. norm (f n))" using assms by (intro summable_norm summable_comparison_test[OF _ assms(2)] exI[of _ 0]) auto hence "norm (suminf f) \ (\n. norm (f n))" by (intro summable_norm) auto also have "\ \ suminf g" by (intro suminf_le * assms allI) finally show ?thesis . qed lemma of_nat_powr_neq_1_complex [simp]: assumes "n > 1" "Re s \ 0" shows "of_nat n powr s \ (1::complex)" proof - have "norm (of_nat n powr s) = real n powr Re s" by (simp add: norm_powr_real_powr) also have "\ \ 1" using assms by (auto simp: powr_def) finally show ?thesis by auto qed lemma abs_summable_on_uminus_iff: "(\x. -f x) abs_summable_on A \ f abs_summable_on A" using abs_summable_on_uminus[of f A] abs_summable_on_uminus[of "\x. -f x" A] by auto lemma abs_summable_on_cmult_right_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. c * f x) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_right[of c f A] abs_summable_on_cmult_right[of "inverse c" "\x. c * f x" A] by (auto simp: field_simps) lemma abs_summable_on_cmult_left_iff: fixes f :: "'a \ 'b :: {banach, real_normed_field, second_countable_topology}" assumes "c \ 0" shows "(\x. f x * c) abs_summable_on A \ f abs_summable_on A" using assms abs_summable_on_cmult_left[of c f A] abs_summable_on_cmult_left[of "inverse c" "\x. f x * c" A] by (auto simp: field_simps) lemma fds_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n)" proof - have "fds_deriv f / f = - fds (\n. fds_nth f n * mangoldt n) * f / f" using completely_multiplicative_fds_deriv[of "fds_nth f"] assms by simp also have "\ = - fds (\n. fds_nth f n * mangoldt n)" using assms by (simp add: divide_fds_def fds_right_inverse) finally show ?thesis . qed lemma fds_nth_logderiv_completely_multiplicative: fixes f :: "'a :: {real_normed_field} fds" assumes "completely_multiplicative_function (fds_nth f)" "fds_nth f 1 \ 0" shows "fds_nth (fds_deriv f / f) n = -fds_nth f n * mangoldt n" using assms by (subst fds_logderiv_completely_multiplicative) (simp_all add: fds_nth_fds') lemma eval_fds_logderiv_completely_multiplicative: fixes s :: "'a :: dirichlet_series" and l :: 'a and f :: "'a fds" defines "h \ fds_deriv f / f" assumes "completely_multiplicative_function (fds_nth f)" and [simp]: "fds_nth f 1 \ 0" assumes "s \ 1 > abs_conv_abscissa f" shows "(\p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)) abs_summable_on {p. prime p}" (is ?th1) and "eval_fds h s = -(\\<^sub>ap | prime p. of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" (is ?th2) proof - let ?P = "{p::nat. prime p}" interpret f: completely_multiplicative_function "fds_nth f" by fact have "fds_abs_converges h s" using abs_conv_abscissa_completely_multiplicative_log_deriv[OF assms(2)] assms by (intro fds_abs_converges) auto hence *: "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV" by (auto simp: h_def fds_abs_converges_altdef') note * also have "(\n. fds_nth h n / nat_power n s) abs_summable_on UNIV \ (\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro abs_summable_on_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) finally have sum1: "(\x. -fds_nth f x * mangoldt x / nat_power x s) abs_summable_on Collect primepow" by (rule abs_summable_on_subset) auto also have "?this \ (\(p,k). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s) abs_summable_on (?P \ UNIV)" using bij_betw_primepows unfolding case_prod_unfold by (intro abs_summable_on_reindex_bij_betw [symmetric]) also have "\ \ (\(p,k). -((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) abs_summable_on (?P \ UNIV)" unfolding case_prod_unfold by (intro abs_summable_on_cong, subst mangoldt_primepow) (auto simp: f.mult f.power nat_power_mult_distrib nat_power_power_left power_divide dest: prime_gt_1_nat) finally have sum2: \ . have sum4: "summable (\n. (norm (fds_nth f p / nat_power p s)) ^ Suc n)" if p: "prime p" for p proof - have "summable (\n. \ln (real p)\ * (norm (fds_nth f p / nat_power p s)) ^ Suc n)" using p abs_summable_on_Sigma_project2[OF sum2, of p] unfolding abs_summable_on_nat_iff' by (simp add: norm_power norm_mult norm_divide mult_ac del: power_Suc) thus ?thesis by (rule summable_mult_D) (insert p, auto dest: prime_gt_1_nat) qed have sums: "(\n. (fds_nth f p / nat_power p s) ^ Suc n) sums (1 / (1 - fds_nth f p / nat_power p s) - 1)" if p: "prime p" for p :: nat proof - from sum4[OF p] have "norm (fds_nth f p / nat_power p s) < 1" unfolding summable_Suc_iff by (simp add: summable_geometric_iff) from geometric_sums[OF this] show ?thesis by (subst sums_Suc_iff) auto qed have eq: "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))" if p: "prime p" for p proof - have "(\\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k * of_real (ln (real p)))) = (\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) * of_real (-ln (real p))" using sum4[of p] p by (subst infsetsum_cmult_left [symmetric]) (auto simp: abs_summable_on_nat_iff' norm_power simp del: power_Suc) also have "(\\<^sub>ak. (fds_nth f p / nat_power p s) ^ Suc k) = (1 / (1 - fds_nth f p / nat_power p s) - 1)" using sum4[OF p] sums[OF p] by (subst infsetsum_nat') (auto simp: sums_iff abs_summable_on_nat_iff' norm_power simp del: power_Suc) finally show ?thesis by (simp add: mult_ac) qed have sum3: "(\x. \\<^sub>ay. - ((fds_nth f x / nat_power x s) ^ Suc y * of_real (ln (real x)))) abs_summable_on {p. prime p}" using sum2 by (rule abs_summable_on_Sigma_project1') auto also have "?this \ (\p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1))) abs_summable_on {p. prime p}" by (intro abs_summable_on_cong eq) auto also have "\ \ ?th1" by (subst abs_summable_on_uminus_iff) auto finally show ?th1 . have "eval_fds h s = (\\<^sub>an. fds_nth h n / nat_power n s)" using * unfolding eval_fds_def by (subst infsetsum_nat') auto also have "\ = (\\<^sub>an \ {n. primepow n}. -fds_nth f n * mangoldt n / nat_power n s)" unfolding h_def using fds_nth_logderiv_completely_multiplicative[OF assms(2)] by (intro infsetsum_cong_neutral) (auto simp: fds_nth_fds mangoldt_def) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -fds_nth f (p ^ Suc k) * mangoldt (p ^ Suc k) / nat_power (p ^ Suc k) s)" using bij_betw_primepows unfolding case_prod_unfold by (intro infsetsum_reindex_bij_betw [symmetric]) also have "\ = (\\<^sub>a(p,k)\(?P \ UNIV). -((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" by (intro infsetsum_cong) (auto simp: f.mult f.power mangoldt_def aprimedivisor_prime_power ln_realpow prime_gt_0_nat nat_power_power_left divide_simps simp del: power_Suc) also have "\ = (\\<^sub>ap | prime p. \\<^sub>ak. - ((fds_nth f p / nat_power p s) ^ Suc k) * of_real (ln (real p)))" using sum2 by (subst infsetsum_Times) (auto simp: case_prod_unfold) also have "\ = (\\<^sub>ap | prime p. -(of_real (ln (real p)) * (1 / (1 - fds_nth f p / nat_power p s) - 1)))" using eq by (intro infsetsum_cong) auto finally show ?th2 by (subst (asm) infsetsum_uminus) qed lemma eval_fds_logderiv_zeta: assumes "Re s > 1" shows "(\p. of_real (ln (real p)) / (p powr s - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta s / zeta s = -(\\<^sub>ap | prime p. of_real (ln (real p)) / (p powr s - 1))" (is ?th2) proof - have *: "completely_multiplicative_function (fds_nth fds_zeta :: _ \ complex)" by standard auto note abscissa = le_less_trans[OF abs_conv_abscissa_completely_multiplicative_log_deriv[OF *]] have "(\p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1)) abs_summable_on {p. prime p}" using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?this \ (\p. ln (real p) / (p powr s - 1)) abs_summable_on {p. prime p}" using assms by (intro abs_summable_on_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th1 . from assms have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto have "deriv zeta s = deriv (eval_fds fds_zeta) s" by (intro deriv_cong_ev[OF eventually_mono[OF ev]]) (auto simp: eval_fds_zeta) also have "deriv (eval_fds fds_zeta) s / zeta s = eval_fds (fds_deriv fds_zeta / fds_zeta) s" using assms zeta_Re_gt_1_nonzero[of s] by (subst eval_fds_log_deriv) (auto simp: eval_fds_zeta eval_fds_deriv intro!: abscissa) also have "eval_fds (fds_deriv fds_zeta / fds_zeta) s = -(\\<^sub>ap | prime p. ln (real p) * (1 / (1 - fds_nth fds_zeta p / p powr s) - 1))" (is "_ = -?S") using eval_fds_logderiv_completely_multiplicative[OF *, of s] assms by auto also have "?S = (\\<^sub>ap | prime p. ln (real p) / (p powr s - 1))" using assms by (intro infsetsum_cong) (auto simp: fds_nth_zeta divide_simps dest: prime_gt_1_nat) finally show ?th2 . qed lemma sums_logderiv_zeta: assumes "Re s > 1" shows "(\p. if prime p then of_real (ln (real p)) / (of_nat p powr s - 1) else 0) sums -(deriv zeta s / zeta s)" (is "?f sums _") proof - note * = eval_fds_logderiv_zeta[OF assms] from sums_infsetsum_nat[OF *(1)] and *(2) show ?thesis by simp qed lemma abs_conv_abscissa_diff_le: "abs_conv_abscissa (f - g :: 'a :: dirichlet_series fds) \ max (abs_conv_abscissa f) (abs_conv_abscissa g)" using abs_conv_abscissa_add_le[of f "-g"] by auto lemma abs_conv_abscissa_diff_leI: "abs_conv_abscissa (f :: 'a :: dirichlet_series fds) \ d \ abs_conv_abscissa g \ d \ abs_conv_abscissa (f - g) \ d" using abs_conv_abscissa_diff_le[of f g] by (auto simp: le_max_iff_disj) lemma range_add_nat: "range (\n. n + c) = {(c::nat)..}" proof safe fix x assume "x \ c" hence "x = x - c + c" by simp thus "x \ range (\n. n + c)" by blast qed auto lemma abs_summable_hurwitz_zeta: assumes "Re s > 1" "a + real b > 0" shows "(\n. 1 / (of_nat n + a) powr s) abs_summable_on {b..}" proof - from assms have "summable (\n. cmod (1 / (of_nat (n + b) + a) powr s))" using summable_hurwitz_zeta_real[of "Re s" "a + b"] by (auto simp: norm_divide powr_minus field_simps norm_powr_real_powr) hence "(\n. 1 / (of_nat (n + b) + a) powr s) abs_summable_on UNIV" by (auto simp: abs_summable_on_nat_iff' add_ac) also have "?this \ (\n. 1 / (of_nat n + a) powr s) abs_summable_on range (\n. n + b)" by (rule abs_summable_on_reindex_iff) auto also have "range (\n. n + b) = {b..}" by (rule range_add_nat) finally show ?thesis . qed lemma hurwitz_zeta_nat_conv_infsetsum: assumes "a > 0" and "Re s > 1" shows "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" proof - have "hurwitz_zeta (real a) s = (\n. of_nat (n + a) powr -s)" using assms by (subst hurwitz_zeta_conv_suminf) auto also have "\ = (\\<^sub>an. of_nat (n + a) powr -s)" using abs_summable_hurwitz_zeta[of s a 0] assms by (intro infsetsum_nat' [symmetric]) (auto simp: powr_minus field_simps) finally show "hurwitz_zeta (real a) s = (\\<^sub>an. of_nat (n + a) powr -s)" . also have "\ = (\\<^sub>an\range (\n. n + a). of_nat n powr -s)" by (rule infsetsum_reindex [symmetric]) auto also have "range (\n. n + a) = {a..}" by (rule range_add_nat) finally show "hurwitz_zeta (real a) s = (\\<^sub>an\{a..}. of_nat n powr -s)" . qed lemma continuous_on_pre_zeta [continuous_intros]: assumes "continuous_on A f" "a > 0" shows "continuous_on A (\x. pre_zeta a (f x))" proof - from assms have "continuous_on UNIV (pre_zeta a)" by (intro holomorphic_on_imp_continuous_on[OF holomorphic_pre_zeta]) auto from continuous_on_compose2[OF this assms(1)] show ?thesis by simp qed lemma continuous_pre_zeta [continuous_intros]: assumes "continuous (at x within A) f" "a > 0" shows "continuous (at x within A) (\x. pre_zeta a (f x))" proof - have "continuous (at z) (pre_zeta a)" for z by (rule continuous_on_interior[of UNIV]) (insert assms, auto intro!: continuous_intros) from continuous_within_compose3[OF this assms(1)] show ?thesis . qed lemma pre_zeta_bound: assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" proof - let ?f = "\x. - (s * (x + a) powr (-1-s))" let ?g' = "\x. norm s * (x + a) powr (-1-Re s)" let ?g = "\x. -norm s / Re s * (x + a) powr (-Re s)" define R where "R = EM_remainder 1 ?f 0" have [simp]: "-Re s - 1 = -1 - Re s" by (simp add: algebra_simps) have "\frac x - 1 / 2\ \ 1 / 2" for x :: real unfolding frac_def by linarith hence "\pbernpoly (Suc 0) x\ \ 1 / 2" for x by (simp add: pbernpoly_def bernpoly_def) moreover have "((\b. cmod s * (b + a) powr - Re s / Re s) \ 0) at_top" using \Re s > 0\ \a > 0\ by real_asymp ultimately have *: "\x. x \ real 0 \ norm (EM_remainder 1 ?f (int x)) \ (1 / 2) / fact 1 * (-?g (real x))" using \a > 0\ \Re s > 0\ by (intro norm_EM_remainder_le_strong_nat'[where g' = ?g' and Y = "{}"]) (auto intro!: continuous_intros derivative_eq_intros simp: field_simps norm_mult norm_powr_real_powr add_eq_0_iff) have R: "norm R \ norm s / (2 * Re s) * a powr -Re s" unfolding R_def using spec[OF *, of 0] by simp from assms have "pre_zeta a s = a powr -s / 2 + R" by (simp add: pre_zeta_def pre_zeta_aux_def R_def) also have "norm \ \ a powr -Re s / 2 + norm s / (2 * Re s) * a powr -Re s" using a by (intro order.trans[OF norm_triangle_ineq] add_mono R) (auto simp: norm_powr_real_powr) also have "\ = (1 + norm s / Re s) / 2 * a powr -Re s" by (simp add: field_simps) finally show ?thesis . qed lemma pre_zeta_bound': assumes "0 < Re s" and a: "a > 0" shows "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" proof - from assms have "norm (pre_zeta a s) \ (1 + norm s / Re s) / 2 * a powr -Re s" by (intro pre_zeta_bound) auto also have "\ = (Re s + norm s) / 2 / (Re s * a powr Re s)" using assms by (auto simp: field_simps powr_minus) also have "Re s + norm s \ norm s + norm s" by (intro add_right_mono complex_Re_le_cmod) also have "(norm s + norm s) / 2 = norm s" by simp finally show "norm (pre_zeta a s) \ norm s / (Re s * a powr Re s)" using assms by (simp add: divide_right_mono) qed lemma summable_comparison_test_bigo: fixes f :: "nat \ real" assumes "summable (\n. norm (g n))" "f \ O(g)" shows "summable f" proof - from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" by (auto elim: landau_o.bigE) thus ?thesis by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) qed lemma deriv_zeta_eq: assumes s: "s \ 1" shows "deriv zeta s = deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2" proof - from s have ev: "eventually (\z. z \ 1) (nhds s)" by (intro t1_space_nhds) have [derivative_intros]: "(pre_zeta 1 has_field_derivative deriv (pre_zeta 1) s) (at s)" by (intro holomorphic_derivI[of _ UNIV] holomorphic_intros) auto have "((\s. pre_zeta 1 s + 1 / (s - 1)) has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" using s by (auto intro!: derivative_eq_intros simp: power2_eq_square) also have "?this \ (zeta has_field_derivative (deriv (pre_zeta 1) s - 1 / (s - 1)\<^sup>2)) (at s)" by (intro has_field_derivative_cong_ev eventually_mono[OF ev]) (auto simp: zeta_def hurwitz_zeta_def) finally show ?thesis by (rule DERIV_imp_deriv) qed lemma zeta_remove_zero: assumes "Re s \ 1" shows "(s - 1) * pre_zeta 1 s + 1 \ 0" proof (cases "s = 1") case False hence "(s - 1) * pre_zeta 1 s + 1 = (s - 1) * zeta s" by (simp add: zeta_def hurwitz_zeta_def divide_simps) also from False assms have "\ \ 0" using zeta_Re_ge_1_nonzero[of s] by auto finally show ?thesis . qed auto lemma eval_fds_deriv_zeta: assumes "Re s > 1" shows "eval_fds (fds_deriv fds_zeta) s = deriv zeta s" proof - have ev: "eventually (\z. z \ {z. Re z > 1}) (nhds s)" using assms by (intro eventually_nhds_in_open open_halfspace_Re_gt) auto from assms have "eval_fds (fds_deriv fds_zeta) s = deriv (eval_fds fds_zeta) s" by (subst eval_fds_deriv) auto also have "\ = deriv zeta s" by (intro deriv_cong_ev eventually_mono[OF ev]) (auto simp: eval_fds_zeta) finally show ?thesis . qed lemma length_sorted_list_of_set [simp]: "finite A \ length (sorted_list_of_set A) = card A" by (metis length_remdups_card_conv length_sort set_sorted_list_of_set sorted_list_of_set_sort_remdups) lemma le_nat_iff': "x \ nat y \ x = 0 \ y \ 0 \ int x \ y" by auto lemma sum_upto_plus1: assumes "x \ 0" shows "sum_upto f (x + 1) = sum_upto f x + f (Suc (nat \x\))" proof - have "sum_upto f (x + 1) = sum f {0<..Suc (nat \x\)}" using assms by (simp add: sum_upto_altdef nat_add_distrib) also have "{0<..Suc (nat \x\)} = insert (Suc (nat \x\)) {0<..nat \x\}" by auto also have "sum f \ = sum_upto f x + f (Suc (nat \x\))" by (subst sum.insert) (auto simp: sum_upto_altdef add_ac) finally show ?thesis . qed lemma sum_upto_minus1: assumes "x \ 1" shows "sum_upto f (x - 1) = (sum_upto f x - f (nat \x\) :: 'a :: ab_group_add)" using sum_upto_plus1[of "x - 1" f] assms by (simp add: algebra_simps nat_diff_distrib) lemma integral_smallo: fixes f g g' :: "real \ real" assumes "f \ o(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x)" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ o(g)" proof (rule landau_o.smallI) fix c :: real assume c: "c > 0" note [continuous_intros] = continuous_on_subset[OF cont] define c' where "c' = c / 2" from c have c': "c' > 0" by (simp add: c'_def) from landau_o.smallD[OF assms(1) this] obtain b where b: "\x. x \ b \ norm (f x) \ c' * norm (g' x)" unfolding eventually_at_top_linorder by blast define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c' * g x) at_top at_top" using c' by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c' * g x \ D - c' * g b') at_top" by (auto simp: filterlim_at_top) thus "eventually (\x. norm (integral {a..x} f) \ c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c' * norm (g' x))" using b' elim assms c' integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c' * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: has_field_derivative_at_within[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c' * (g x - g b') \ c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed qed lemma integral_bigo: fixes f g g' :: "real \ real" assumes "f \ O(g')" and "filterlim g at_top at_top" assumes "\a' x. a \ a' \ a' \ x \ f integrable_on {a'..x}" assumes deriv: "\x. x \ a \ (g has_field_derivative g' x) (at x within {a..})" assumes cont: "continuous_on {a..} g'" assumes nonneg: "\x. x \ a \ g' x \ 0" shows "(\x. integral {a..x} f) \ O(g)" proof - note [continuous_intros] = continuous_on_subset[OF cont] from landau_o.bigE[OF assms(1)] obtain c b where c: "c > 0" and b: "\x. x \ b \ norm (f x) \ c * norm (g' x)" unfolding eventually_at_top_linorder by metis define c' where "c' = c / 2" define b' where "b' = max a b" define D where "D = norm (integral {a..b'} f)" have "filterlim (\x. c * g x) at_top at_top" using c by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] assms) hence "eventually (\x. c * g x \ D - c * g b') at_top" by (auto simp: filterlim_at_top) hence "eventually (\x. norm (integral {a..x} f) \ 2 * c * norm (g x)) at_top" using eventually_ge_at_top[of b'] proof eventually_elim case (elim x) have b': "a \ b'" "b \ b'" by (auto simp: b'_def) from elim b' have integrable: "(\x. \g' x\) integrable_on {b'..x}" by (intro integrable_continuous_real continuous_intros) auto have "integral {a..x} f = integral {a..b'} f + integral {b'..x} f" using elim b' by (intro Henstock_Kurzweil_Integration.integral_combine [symmetric] assms) auto also have "norm \ \ D + norm (integral {b'..x} f)" unfolding D_def by (rule norm_triangle_ineq) also have "norm (integral {b'..x} f) \ integral {b'..x} (\x. c * norm (g' x))" using b' elim assms c integrable by (intro integral_norm_bound_integral b assms) auto also have "\ = c * integral {b'..x} (\x. \g' x\)" by simp also have "integral {b'..x} (\x. \g' x\) = integral {b'..x} g'" using assms b' by (intro integral_cong) auto also have "(g' has_integral (g x - g b')) {b'..x}" using b' elim by (intro fundamental_theorem_of_calculus) (auto simp flip: has_field_derivative_iff_has_vector_derivative intro!: DERIV_subset[OF deriv]) hence "integral {b'..x} g' = g x - g b'" by (simp add: has_integral_iff) also have "D + c * (g x - g b') \ 2 * c * g x" using elim by (simp add: field_simps c'_def) also have "\ \ 2 * c * norm (g x)" using c by (intro mult_left_mono) auto finally show ?case by simp qed thus ?thesis by (rule bigoI) qed lemma primepows_le_subset: assumes x: "x > 0" and l: "l > 0" shows "{(p, i). prime p \ l \ i \ real (p ^ i) \ x} \ {..nat \root l x\} \ {..nat \log 2 x\}" proof safe fix p i :: nat assume pi: "prime p" "i \ l" "real (p ^ i) \ x" have "real p ^ l \ real p ^ i" using pi x l by (intro power_increasing) (auto dest: prime_gt_0_nat) also have "\ \ x" using pi by simp finally have "root l (real p ^ l) \ root l x" using x pi l by (subst real_root_le_iff) auto also have "root l (real p ^ l) = real p" using pi l by (subst real_root_pos2) auto finally show "p \ nat \root l x\" using pi l x by (simp add: le_nat_iff' le_floor_iff) from pi have "2 ^ i \ real p ^ i" using l by (intro power_mono) (auto dest: prime_gt_1_nat) also have "\ \ x" using pi by simp finally show "i \ nat \log 2 x\" using pi x by (auto simp: le_nat_iff' le_floor_iff le_log_iff powr_realpow) qed lemma mangoldt_non_primepow: "\primepow n \ mangoldt n = 0" by (auto simp: mangoldt_def) lemma le_imp_bigo_real: assumes "c \ 0" "eventually (\x. f x \ c * (g x :: real)) F" "eventually (\x. 0 \ f x) F" shows "f \ O[F](g)" proof - have "eventually (\x. norm (f x) \ c * norm (g x)) F" using assms(2,3) proof eventually_elim case (elim x) have "norm (f x) \ c * g x" using elim by simp also have "\ \ c * norm (g x)" by (intro mult_left_mono assms) auto finally show ?case . qed thus ?thesis by (intro bigoI[of _ c]) auto qed (* TODO: unneeded. But why does real_asymp not work? *) lemma ln_minus_ln_floor_bigo: "(\x. ln x - ln (real (nat \x\))) \ O(\_. 1)" proof (intro le_imp_bigo_real[of 1] eventually_mono[OF eventually_ge_at_top[of 1]]) fix x :: real assume x: "x \ 1" from x have *: "x - real (nat \x\) \ 1" by linarith from x have "ln x - ln (real (nat \x\)) \ (x - real (nat \x\)) / real (nat \x\)" by (intro ln_diff_le) auto also have "\ \ 1 / 1" using x * by (intro frac_le) auto finally show "ln x - ln (real (nat \x\)) \ 1 * 1" by simp qed auto lemma cos_geD: assumes "cos x \ cos a" "0 \ a" "a \ pi" "-pi \ x" "x \ pi" shows "x \ {-a..a}" proof (cases "x \ 0") case True with assms show ?thesis by (subst (asm) cos_mono_le_eq) auto next case False with assms show ?thesis using cos_mono_le_eq[of a "-x"] by auto qed (* TODO: Could be generalised *) lemma path_image_part_circlepath_same_Re: assumes "0 \ b" "b \ pi" "a = -b" "r \ 0" shows "path_image (part_circlepath c r a b) = sphere c r \ {s. Re s \ Re c + r * cos a}" proof safe fix z assume "z \ path_image (part_circlepath c r a b)" with assms obtain t where t: "t \ {a..b}" "z = c + of_real r * cis t" by (auto simp: path_image_part_circlepath exp_eq_polar) from t and assms show "z \ sphere c r" by (auto simp: dist_norm norm_mult) from t and assms show "Re z \ Re c + r * cos a" using cos_monotone_0_pi_le[of t b] cos_monotone_minus_pi_0'[of a t] by (cases "t \ 0") (auto intro!: mult_left_mono) next fix z assume z: "z \ sphere c r" "Re z \ Re c + r * cos a" show "z \ path_image (part_circlepath c r a b)" proof (cases "r = 0") case False with assms have r: "r > 0" by simp with z have z_eq: "z = c + r * cis (Arg (z - c))" using Arg_eq[of "z - c"] by (auto simp: dist_norm exp_eq_polar norm_minus_commute) moreover from z(2) r assms have "cos b \ cos (Arg (z - c))" by (subst (asm) z_eq) auto with assms have "Arg (z - c) \ {-b..b}" using Arg_le_pi[of "z - c"] mpi_less_Arg[of "z - c"] by (intro cos_geD) auto ultimately show "z \ path_image (part_circlepath c r a b)" using assms by (subst path_image_part_circlepath) (auto simp: exp_eq_polar) qed (insert assms z, auto simp: path_image_part_circlepath) qed lemma part_circlepath_rotate_left: "part_circlepath c r (x + a) (x + b) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma part_circlepath_rotate_right: "part_circlepath c r (a + x) (b + x) = (\z. c + cis x * (z - c)) \ part_circlepath c r a b" by (simp add: part_circlepath_def exp_eq_polar fun_eq_iff linepath_translate_left linepath_translate_right cis_mult add_ac) lemma path_image_semicircle_Re_ge: assumes "r \ 0" shows "path_image (part_circlepath c r (-pi/2) (pi/2)) = sphere c r \ {s. Re s \ Re c}" by (subst path_image_part_circlepath_same_Re) (simp_all add: assms) lemma sphere_rotate: "(\z. c + cis x * (z - c)) ` sphere c r = sphere c r" proof safe fix z assume z: "z \ sphere c r" hence "z = c + cis x * (c + cis (-x) * (z - c) - c)" "c + cis (-x) * (z - c) \ sphere c r" by (auto simp: dist_norm norm_mult norm_minus_commute cis_conv_exp exp_minus field_simps norm_divide) with z show "z \ (\z. c + cis x * (z - c)) ` sphere c r" by blast qed (auto simp: dist_norm norm_minus_commute norm_mult) lemma path_image_semicircle_Re_le: assumes "r \ 0" shows "path_image (part_circlepath c r (pi/2) (3/2*pi)) = sphere c r \ {s. Re s \ Re c}" proof - let ?f = "(\z. c + cis pi * (z - c))" have *: "part_circlepath c r (pi/2) (3/2*pi) = part_circlepath c r (pi + (-pi/2)) (pi + pi/2)" by simp have "path_image (part_circlepath c r (pi/2) (3/2*pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Re c \ Re s}" by (auto simp: image_iff intro!: exI[of _ "2 * c - x" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_ge: assumes "r \ 0" shows "path_image (part_circlepath c r 0 pi) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (pi/2) * (z - c))" have *: "part_circlepath c r 0 pi = part_circlepath c r (pi / 2 + (-pi/2)) (pi / 2 + pi/2)" by simp have "path_image (part_circlepath c r 0 pi) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c - \ * (x - c)" for x]) finally show ?thesis . qed lemma path_image_semicircle_Im_le: assumes "r \ 0" shows "path_image (part_circlepath c r pi (2 * pi)) = sphere c r \ {s. Im s \ Im c}" proof - let ?f = "(\z. c + cis (3*pi/2) * (z - c))" have *: "part_circlepath c r pi (2*pi) = part_circlepath c r (3*pi/2 + (-pi/2)) (3*pi/2 + pi/2)" by simp have "path_image (part_circlepath c r pi (2 * pi)) = ?f ` sphere c r \ ?f ` {s. Re c \ Re s}" unfolding * part_circlepath_rotate_left path_image_compose path_image_semicircle_Re_ge[OF assms] by auto also have "?f ` sphere c r = sphere c r" by (rule sphere_rotate) also have "cis (3 * pi / 2) = -\" using cis_mult[of pi "pi / 2"] by simp hence "?f ` {s. Re c \ Re s} = {s. Im c \ Im s}" by (auto simp: image_iff intro!: exI[of _ "c + \ * (x - c)" for x]) finally show ?thesis . qed lemma powr_numeral [simp]: "x \ 0 \ (x::real) powr numeral y = x ^ numeral y" using powr_numeral[of x y] by (cases "x = 0") auto lemma eval_fds_logderiv_zeta_real: assumes "x > (1 :: real)" shows "(\p. ln (real p) / (p powr x - 1)) abs_summable_on {p. prime p}" (is ?th1) and "deriv zeta (of_real x) / zeta (of_real x) = -of_real (\\<^sub>ap | prime p. ln (real p) / (p powr x - 1))" (is ?th2) proof - have "(\p. Re (of_real (ln (real p)) / (of_nat p powr of_real x - 1))) abs_summable_on {p. prime p}" using assms by (intro abs_summable_Re eval_fds_logderiv_zeta) auto also have "?this \ ?th1" by (intro abs_summable_on_cong) (auto simp: powr_Reals_eq) finally show ?th1 . show ?th2 using assms by (subst eval_fds_logderiv_zeta) (auto simp: infsetsum_of_real [symmetric] powr_Reals_eq) qed lemma fixes a b c d :: real assumes ab: "d * a + b \ 1" and c: "c < -1" and d: "d > 0" defines "C \ - ((ln (d * a + b) - 1 / (c + 1)) * (d * a + b) powr (c + 1) / (d * (c + 1)))" shows set_integrable_ln_powr_at_top: "(\x. (ln (d * x + b) * ((d * x + b) powr c))) absolutely_integrable_on {a<..}" (is ?th1) and set_lebesgue_integral_ln_powr_at_top: "(\x\{a<..}. (ln (d * x + b) * ((d * x + b) powr c)) \lborel) = C" (is ?th2) and ln_powr_has_integral_at_top: "((\x. ln (d * x + b) * (d * x + b) powr c) has_integral C) {a<..}" (is ?th3) proof - define f where "f = (\x. ln (d * x + b) * (d * x + b) powr c)" define F where "F = (\x. (ln (d * x + b) - 1 / (c + 1)) * (d * x + b) powr (c + 1) / (d * (c + 1)))" have *: "(F has_field_derivative f x) (at x)" "isCont f x" "f x \ 0" if "x > a" for x proof - have "1 \ d * a + b" by fact also have "\ < d * x + b" using that assms by (intro add_strict_right_mono mult_strict_left_mono) finally have gt_1: "d * x + b > 1" . show "(F has_field_derivative f x) (at x)" "isCont f x" using ab c d gt_1 by (auto simp: F_def f_def divide_simps intro!: derivative_eq_intros continuous_intros) (auto simp: algebra_simps powr_add)? show "f x \ 0" using gt_1 by (auto simp: f_def) qed have limits: "((F \ real_of_ereal) \ F a) (at_right (ereal a))" "((F \ real_of_ereal) \ 0) (at_left \)" using c ab d unfolding ereal_tendsto_simps1 F_def by (real_asymp; simp add: field_simps)+ have 1: "set_integrable lborel (einterval a \) f" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: * AE_I2) thus 2: "f absolutely_integrable_on {a<..}" by (auto simp: set_integrable_def integrable_completion) have "(LBINT x=ereal a..\. f x) = 0 - F a" using ab c limits by (intro interval_integral_FTC_nonneg) (auto intro!: *) thus 3: ?th2 by (simp add: interval_integral_to_infinity_eq F_def f_def C_def) show ?th3 using set_borel_integral_eq_integral[OF 1] 3 by (simp add: has_integral_iff f_def C_def) qed lemma ln_fact_conv_sum_upto: "ln (fact n) = sum_upto ln n" by (induction n) (auto simp: sum_upto_plus1 add.commute[of 1] ln_mult) lemma sum_upto_ln_conv_ln_fact: "sum_upto ln x = ln (fact (nat \x\))" by (simp add: ln_fact_conv_sum_upto sum_upto_altdef) lemma real_of_nat_div: "real (a div b) = real_of_int \real a / real b\" by (subst floor_divide_of_nat_eq) auto lemma integral_subset_negligible: fixes f :: "'a :: euclidean_space \ 'b :: banach" assumes "S \ T" "negligible (T - S)" shows "integral S f = integral T f" proof - have "integral T f = integral T (\x. if x \ S then f x else 0)" by (rule integral_spike[of "T - S"]) (use assms in auto) also have "\ = integral (S \ T) f" by (subst integral_restrict_Int) auto also have "S \ T = S" using assms by auto finally show ?thesis .. qed lemma integrable_on_cong [cong]: assumes "\x. x \ A \ f x = g x" "A = B" shows "f integrable_on A \ g integrable_on B" using has_integral_cong[of A f g, OF assms(1)] assms(2) by (auto simp: integrable_on_def) lemma measurable_sum_upto [measurable]: fixes f :: "'a \ nat \ real" assumes [measurable]: "\y. (\t. f t y) \ M \\<^sub>M borel" assumes [measurable]: "x \ M \\<^sub>M borel" shows "(\t. sum_upto (f t) (x t)) \ M \\<^sub>M borel" proof - have meas: "(\t. set_lebesgue_integral lborel {y. y \ 0 \ y - real (nat \x t\) \ 0} (\y. f t (nat \y\))) \ M \\<^sub>M borel" (is "?f \ _") unfolding set_lebesgue_integral_def by measurable also have "?f = (\t. sum_upto (f t) (x t))" proof fix t :: 'a show "?f t = sum_upto (f t) (x t)" proof (cases "x t < 1") case True hence "{y. y \ 0 \ y - real (nat \x t\) \ 0} = {0}" by auto thus ?thesis using True by (simp add: set_integral_at_point sum_upto_altdef) next case False define n where "n = nat \x t\" from False have "n > 0" by (auto simp: n_def) have *: "((\x. f t (nat \x\)) has_integral sum (f t) {0<..n}) {real 0..real n}" using \n > 0\ by (intro nat_sum_has_integral_ceiling) auto have **: "(\x. f t (nat \x\)) absolutely_integrable_on {real 0..real n}" proof (rule absolutely_integrable_absolutely_integrable_ubound) show "(\_. MAX n\{0..n}. \f t n\) absolutely_integrable_on {real 0..real n}" using \n > 0\ by (subst absolutely_integrable_on_iff_nonneg) (auto simp: Max_ge_iff intro!: exI[of _ "f t 0"]) show "(\x. f t (nat \x\)) integrable_on {real 0..real n}" using * by (simp add: has_integral_iff) next fix y :: real assume y: "y \ {real 0..real n}" have "f t (nat \y\) \ \f t (nat \y\)\" by simp also have "\ \ (MAX n\{0..n}. \f t n\)" using y by (intro Max.coboundedI) auto finally show "f t (nat \y\) \ (MAX n\{0..n}. \f t n\)" . qed have "sum (f t) {0<..n} = (\x\{real 0..real n}. f t (nat \x\) \lebesgue)" using has_integral_set_lebesgue[OF **] * by (simp add: has_integral_iff) also have "\ = (\x\{real 0..real n}. f t (nat \x\) \lborel)" unfolding set_lebesgue_integral_def by (subst integral_completion) auto also have "{real 0..real n} = {y. 0 \ y \ y - real (nat \x t\) \ 0}" by (auto simp: n_def) also have "sum (f t) {0<..n} = sum_upto (f t) (x t)" by (simp add: sum_upto_altdef n_def) finally show ?thesis .. qed qed finally show ?thesis . qed end diff --git a/thys/Treaps/Random_List_Permutation.thy b/thys/Treaps/Random_List_Permutation.thy --- a/thys/Treaps/Random_List_Permutation.thy +++ b/thys/Treaps/Random_List_Permutation.thy @@ -1,473 +1,474 @@ (* File: Random_List.thy Authors: Manuel Eberl Some facts about randomly permuted lists and how to obtain them by drawing i.i.d. priorities for every element. *) section \Randomly-permuted lists\ theory Random_List_Permutation imports Probability_Misc Comparison_Sort_Lower_Bound.Linorder_Relations begin subsection \General facts about linear orderings\ text \ We define the set of all linear orderings on a given set and show some properties about it. \ definition linorders_on :: "'a set \ ('a \ 'a) set set" where "linorders_on A = {R. linorder_on A R}" lemma linorders_on_empty [simp]: "linorders_on {} = {{}}" by (auto simp: linorders_on_def linorder_on_def refl_on_def) lemma linorders_finite_nonempty: assumes "finite A" shows "linorders_on A \ {}" proof - from finite_distinct_list[OF assms] obtain xs where "set xs = A" "distinct xs" by blast hence "linorder_on A (linorder_of_list xs)" by auto thus ?thesis by (auto simp: linorders_on_def) qed text \ There is an obvious bijection between permutations of a set (i.\,e.\ lists with all elements from that set without repetition) and linear orderings on it. \ lemma bij_betw_linorders_on: assumes "finite A" shows "bij_betw linorder_of_list (permutations_of_set A) (linorders_on A)" using bij_betw_linorder_of_list[of A] assms unfolding linorders_on_def by simp lemma sorted_wrt_list_of_set_linorder_of_list [simp]: assumes "distinct xs" shows "sorted_wrt_list_of_set (linorder_of_list xs) (set xs) = xs" by (rule sorted_wrt_list_of_set_eqI[of "set xs"]) (insert assms, auto) lemma linorder_of_list_sorted_wrt_list_of_set [simp]: assumes "linorder_on A R" "finite A" shows "linorder_of_list (sorted_wrt_list_of_set R A) = R" proof - from assms(1) have subset: "R \ A \ A" by (auto simp: linorder_on_def refl_on_def) from assms and subset show ?thesis by (auto simp: linorder_of_list_def linorder_sorted_wrt_list_of_set sorted_wrt_linorder_index_le_iff) qed lemma bij_betw_linorders_on': assumes "finite A" shows "bij_betw (\R. sorted_wrt_list_of_set R A) (linorders_on A) (permutations_of_set A)" by (rule bij_betw_byWitness[where f' = linorder_of_list]) (insert assms, auto simp: linorders_on_def permutations_of_set_def linorder_sorted_wrt_list_of_set) lemma finite_linorders_on [intro]: assumes "finite A" shows "finite (linorders_on A)" proof - have "finite (permutations_of_set A)" by simp also have "?this \ finite (linorders_on A)" using assms by (rule bij_betw_finite [OF bij_betw_linorders_on]) finally show ?thesis . qed text \ Next, we look at the ordering defined by a list that is permuted with some permutation function. For this, we first define the composition of a relation with a function. \ definition map_relation :: "'a set \ ('a \ 'b) \ ('b \ 'b) set \ ('a \ 'a) set" where "map_relation A f R = {(x,y)\A\A. (f x, f y) \ R}" lemma index_distinct_eqI: assumes "distinct xs" "i < length xs" "xs ! i = x" shows "index xs x = i" using assms by (induction xs arbitrary: i) (auto simp: nth_Cons split: nat.splits) lemma index_permute_list: assumes "\ permutes {.. set xs" shows "index (permute_list \ xs) x = inv \ (index xs x)" proof - have *: "inv \ permutes {.. permutes {.. xs) = map_relation (set xs) ((!) xs \ inv \ \ index xs) (linorder_of_list xs)" proof - note * = permutes_inv[OF assms(1)] have less: "inv \ i < length xs" if "i < length xs" for i using permutes_in_image[OF *] and that by simp from assms and * show ?thesis by (auto simp: linorder_of_list_def map_relation_def index_nth_id index_permute_list less) qed lemma inj_on_conv_Ex1: "inj_on f A \ (\y\f`A. \!x\A. y = f x)" by (auto simp: inj_on_def) lemma bij_betw_conv_Ex1: "bij_betw f A B \ (\y\B. \!x\A. f x = y) \ B = f ` A" unfolding bij_betw_def inj_on_conv_Ex1 by (auto simp: eq_commute) lemma permutesI: assumes "bij_betw f A A" "\x. x \ A \ f x = x" shows "f permutes A" unfolding permutes_def proof (intro conjI allI impI) fix y from assms have [simp]: "f x \ A \ x \ A" for x by (auto simp: bij_betw_def) show "\!x. f x = y" proof (cases "y \ A") case True also from assms have "A = f ` A" by (auto simp: bij_betw_def) finally obtain x where "x \ A" "y = f x" by auto with assms and \y \ A\ show ?thesis by (intro ex1I[of _ x]) (auto simp: bij_betw_def dest: inj_onD) qed (insert assms, auto) qed (insert assms, auto) text \ We now show the important lemma that any two linear orderings on a finite set can be mapped onto each other by a permutation. \ lemma linorder_permutation_exists: assumes "finite A" "linorder_on A R" "linorder_on A R'" obtains \ where "\ permutes A" "R' = map_relation A \ R" proof - define xs where "xs = sorted_wrt_list_of_set R A" define ys where "ys = sorted_wrt_list_of_set R' A" have xs_ys: "distinct xs" "distinct ys" "set xs = A" "set ys = A" using assms by (simp_all add: linorder_sorted_wrt_list_of_set xs_def ys_def) from xs_ys have "mset ys = mset xs" by (simp add: set_eq_iff_mset_eq_distinct [symmetric]) from mset_eq_permutation[OF this] guess \ . note \ = this define \' where "\' = (\x. if x \ A then x else xs ! inv \ (index xs x))" have \': "\' permutes A" proof (rule permutesI) have "bij_betw ((!) xs \ inv \) {.. bij_betw_nth)+ (simp_all add: xs_ys) hence "bij_betw ((!) xs \ inv \ \ index xs) A A" by (rule bij_betw_trans [rotated] bij_betw_index)+ (insert bij_betw_index[of xs A "length xs"], simp_all add: xs_ys atLeast0LessThan) also have "bij_betw ((!) xs \ inv \ \ index xs) A A \ bij_betw \' A A" by (rule bij_betw_cong) (auto simp: \'_def) finally show \ . qed (simp_all add: \'_def) from assms have "R' = linorder_of_list ys" by (simp add: ys_def) also from \ have "ys = permute_list \ xs" by simp also have "linorder_of_list (permute_list \ xs) = map_relation A ((!) xs \ inv \ \ index xs) (linorder_of_list xs)" using \ by (subst linorder_of_list_permute) (simp_all add: xs_ys) also from assms have "linorder_of_list xs = R" by (simp add: xs_def) finally have "R' = map_relation A ((!) xs \ inv \ \ index xs) R" . also have "\ = map_relation A \' R" by (auto simp: map_relation_def \'_def) finally show ?thesis using \' and that[of \'] by simp qed text \ We now define the linear ordering defined by some priority function, i.\,e.\ a function that injectively associates priorities to every element such that elements with lower priority are smaller in the resulting ordering. \ definition linorder_from_keys :: "'a set \ ('a \ 'b :: linorder) \ ('a \ 'a) set" where "linorder_from_keys A f = {(x,y)\A\A. f x \ f y}" lemma linorder_from_keys_permute: assumes "g permutes A" shows "linorder_from_keys A (f \ g) = map_relation A g (linorder_from_keys A f)" using permutes_in_image[OF assms] by (auto simp: map_relation_def linorder_from_keys_def) lemma linorder_on_linorder_from_keys [intro]: assumes "inj_on f A" shows "linorder_on A (linorder_from_keys A f)" using assms by (auto simp: linorder_on_def refl_on_def antisym_def linorder_from_keys_def trans_def total_on_def dest: inj_onD) lemma linorder_from_keys_empty [simp]: "linorder_from_keys {} = (\_. {})" by (simp add: linorder_from_keys_def fun_eq_iff) text \ We now show another important fact, namely that when we draw $n$ values i.\,i.\,d.\ uniformly from a non-trivial real interval, we almost surely get distinct values. \ lemma emeasure_PiM_diagonal: fixes a b :: real assumes "x \ A" "y \ A" "x \ y" assumes "a < b" "finite A" defines "M \ uniform_measure lborel {a..b}" shows "emeasure (PiM A (\_. M)) {h\A \\<^sub>E UNIV. h x = h y} = 0" proof - from assms have M: "prob_space M" unfolding M_def by (intro prob_space_uniform_measure) auto then interpret product_prob_space "\_. M" A unfolding product_prob_space_def product_prob_space_axioms_def product_sigma_finite_def by (auto simp: prob_space_imp_sigma_finite) from M interpret pair_sigma_finite M M by unfold_locales have [measurable]: "{h\extensional {x, y}. h x = h y} \ sets (Pi\<^sub>M {x, y} (\i. lborel))" proof - have "{h\extensional {x,y}. h x = h y} = {h \ space (Pi\<^sub>M {x, y} (\i. lborel)). h x = h y}" by (auto simp: extensional_def space_PiM) also have "... \ sets (Pi\<^sub>M {x, y} (\i. lborel))" by measurable finally show ?thesis . qed have [simp]: "sets (PiM A (\_. M)) = sets (PiM A (\_. lborel))" for A :: "'a set" by (intro sets_PiM_cong refl) (simp_all add: M_def) have sets_M_M: "sets (M \\<^sub>M M) = sets (borel \\<^sub>M borel)" by (intro sets_pair_measure_cong) (auto simp: M_def) have [measurable]: "(\(b, a). if b = a then 1 else 0) \ borel_measurable (M \\<^sub>M M)" unfolding measurable_split_conv by (subst measurable_cong_sets[OF sets_M_M refl]) (auto intro!: measurable_If measurable_const measurable_equality_set) have "{h\A \\<^sub>E UNIV. h x = h y} = (\h. restrict h {x, y}) -` {h\extensional {x, y}. h x = h y} \ space (PiM A (\_. M :: real measure))" by (auto simp: space_PiM PiE_def extensional_def M_def) also have "emeasure (PiM A (\_. M)) \ = emeasure (distr (PiM A (\_. M)) (PiM {x,y} (\_. lborel :: real measure)) (\h. restrict h {x,y})) {h\extensional {x, y}. h x = h y}" proof (rule emeasure_distr [symmetric]) have "(\h. restrict h {x, y}) \ Pi\<^sub>M A (\_. lborel) \\<^sub>M Pi\<^sub>M {x, y} (\_. lborel)" using assms by (intro measurable_restrict_subset) auto also have "\ = Pi\<^sub>M A (\_. M) \\<^sub>M Pi\<^sub>M {x, y} (\_. lborel)" by (intro sets_PiM_cong measurable_cong_sets refl) (simp_all add: M_def) finally show "(\h. restrict h {x, y}) \ \" . next show "{h\extensional {x, y}. h x = h y} \ sets (Pi\<^sub>M {x, y} (\_. lborel))" by simp qed also have "distr (PiM A (\_. M)) (PiM {x,y} (\_. lborel :: real measure)) (\h. restrict h {x,y}) = distr (PiM A (\_. M)) (PiM {x,y} (\_. M)) (\h. restrict h {x,y})" by (intro distr_cong refl sets_PiM_cong) (simp_all add: M_def) also from assms have "\ = Pi\<^sub>M {x, y} (\i. M)" by (intro distr_restrict [symmetric]) auto also have "emeasure \ {h\extensional {x, y}. h x = h y} = nn_integral \ (\h. indicator {h\extensional {x, y}. h x = h y} h)" by (intro nn_integral_indicator [symmetric]) simp_all also have "\ = nn_integral (Pi\<^sub>M {x, y} (\i. M)) (\h. if h x = h y then 1 else 0)" by (intro nn_integral_cong) (auto simp add: indicator_def space_PiM PiE_def) also from \x \ y\ have "\ = (\\<^sup>+ z. (if fst z = snd z then 1 else 0) \(M \\<^sub>M M))" by (intro product_nn_integral_pair) auto also have "\ = (\\<^sup>+ x. (\\<^sup>+y. (if x = y then 1 else 0) \M) \M)" by (subst M.nn_integral_fst [symmetric]) simp_all - also have "\ = (\\<^sup>+ x. (\\<^sup>+y. indicator {x} y \M) \M)" by (simp add: indicator_def eq_commute) + also have "\ = (\\<^sup>+ x. (\\<^sup>+y. indicator {x} y \M) \M)" + by (simp add: indicator_def of_bool_def eq_commute) also have "\ = (\\<^sup>+ x. emeasure M {x} \M)" by (subst nn_integral_indicator) (simp_all add: M_def) also have "\ = (\\<^sup>+ x. 0 \M)" unfolding M_def by (intro nn_integral_cong_AE refl AE_uniform_measureI) auto also have "\ = 0" by simp finally show ?thesis . qed lemma measurable_linorder_from_keys_restrict: assumes fin: "finite A" shows "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (A \ A))" (is "_ : ?M \\<^sub>M _") apply (subst measurable_count_space_eq2) apply (auto simp add: fin linorder_from_keys_def) proof - note fin[simp] fix R assume "R \ A \ A" then have "linorder_from_keys A -` {R} \ space ?M = {f \ space ?M. \x\A. \y\A. (x, y) \ R \ f x \ f y}" by (auto simp add: linorder_from_keys_def set_eq_iff) also have "\ \ sets ?M" by measurable finally show "linorder_from_keys A -` {R} \ space ?M \ sets ?M" . qed lemma measurable_count_space_extend: assumes "f \ measurable M (count_space A)" "A \ B" shows "f \ measurable M (count_space B)" proof - note assms(1) also have "count_space A = restrict_space (count_space B) A" using assms(2) by (subst restrict_count_space) (simp_all add: Int_absorb2) finally show ?thesis by (simp add: measurable_restrict_space2_iff) qed lemma measurable_linorder_from_keys_restrict': assumes fin: "finite A" "A \ B" shows "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (B \ B))" apply (rule measurable_count_space_extend) apply (rule measurable_linorder_from_keys_restrict[OF assms(1)]) using assms by auto context fixes a b :: real and A :: "'a set" and M and B assumes fin: "finite A" and ab: "a < b" and B: "A \ B" defines "M \ distr (PiM A (\_. uniform_measure lborel {a..b})) (count_space (Pow (B \ B))) (linorder_from_keys A)" begin lemma measurable_linorder_from_keys [measurable]: "linorder_from_keys A \ Pi\<^sub>M A (\_. borel :: real measure) \\<^sub>M count_space (Pow (B \ B))" by (rule measurable_linorder_from_keys_restrict') (auto simp: fin B) text \ The ordering defined by randomly-chosen priorities is almost surely linear: \ theorem almost_everywhere_linorder: "AE R in M. linorder_on A R" proof - define N where "N = PiM A (\_. uniform_measure lborel {a..b})" have [simp]: "sets (Pi\<^sub>M A (\_. uniform_measure lborel {a..b})) = sets (PiM A (\_. lborel))" by (intro sets_PiM_cong) simp_all let ?M_A = "(Pi\<^sub>M A (\_. lborel :: real measure))" have meas: "{h \ A \\<^sub>E UNIV. h i = h j} \ sets ?M_A" if [simp]: "i \ A" "j \ A" for i j proof - have "{h \ A \\<^sub>E UNIV. h i = h j} = {h \ space ?M_A. h i = h j}" by (auto simp: space_PiM) also have "... \ sets ?M_A" by measurable finally show ?thesis . qed define X :: "('a \ real) set" where "X = (\x\A. \y\A-{x}. {h\A \\<^sub>E UNIV. h x = h y})" have "AE f in N. inj_on f A" proof (rule AE_I) show "{f \ space N. \ inj_on f A} \ X" by (auto simp: inj_on_def X_def space_PiM N_def) next show "X \ sets N" unfolding X_def N_def using meas by (auto intro!: countable_finite fin sets.countable_UN') next have "emeasure N X \ (\i\A. emeasure N (\y\A - {i}. {h \ A \\<^sub>E UNIV. h i = h y}))" unfolding X_def N_def using fin meas by (intro emeasure_subadditive_finite) (auto simp: disjoint_family_on_def intro!: sets.countable_UN' countable_finite) also have "\ \ (\i\A. \j\A-{i}. emeasure N {h \ A \\<^sub>E UNIV. h i = h j})" unfolding N_def using fin meas by (intro emeasure_subadditive_finite sum_mono) (auto simp: disjoint_family_on_def intro!: sets.countable_UN' countable_finite) also have "\ = (\i\A. \j\A-{i}. 0)" unfolding N_def using fin ab by (intro sum.cong refl emeasure_PiM_diagonal) auto also have "\ = 0" by simp finally show "emeasure N X = 0" by simp qed hence "AE f in N. linorder_on A (linorder_from_keys A f)" by eventually_elim auto thus ?thesis unfolding M_def N_def by (subst AE_distr_iff) auto qed text \ Furthermore, this is equivalent to choosing one of the $|A|!$ linear orderings uniformly at random. \ theorem random_linorder_by_prios: "M = uniform_measure (count_space (Pow (B \ B))) (linorders_on A)" proof - from linorders_finite_nonempty[OF fin] obtain R where R: "linorder_on A R" by (auto simp: linorders_on_def) have *: "emeasure M {R} \ emeasure M {R'}" if "linorder_on A R" "linorder_on A R'" for R R' proof - define N where "N = PiM A (\_. uniform_measure lborel {a..b})" from linorder_permutation_exists[OF fin that] obtain \ where \: "\ permutes A" "R' = map_relation A \ R" by blast have "(\f. f \ \) \ Pi\<^sub>M A (\_. lborel :: real measure) \\<^sub>M Pi\<^sub>M A (\_. lborel :: real measure)" by (auto intro!: measurable_PiM_single' measurable_PiM_component_rev simp: comp_def permutes_in_image[OF \(1)] space_PiM PiE_def extensional_def) also have "\ = N \\<^sub>M Pi\<^sub>M A (\_. lborel)" unfolding N_def by (intro measurable_cong_sets refl sets_PiM_cong) simp_all finally have [measurable]: "(\f. f \ \) \ \" . have [simp]: "measurable N X = measurable (PiM A (\_. lborel)) X" for X :: "('a \ 'a) set measure" unfolding N_def by (intro measurable_cong_sets refl sets_PiM_cong) simp_all have [simp]: "measurable M X = measurable (count_space (Pow (B \ B))) X" for X :: "('a \ 'a) set measure" unfolding M_def by simp have M_eq: "M = distr N (count_space (Pow (B \ B))) (linorder_from_keys A)" by (simp only: M_def N_def) also have "N = distr N (PiM A (\_. lborel)) (\f. f \ \)" unfolding N_def by (rule PiM_uniform_measure_permute [symmetric]) fact+ also have "distr \ (count_space (Pow (B \ B))) (linorder_from_keys A) = distr N (count_space (Pow (B \ B))) (linorder_from_keys A \ (\f. f \ \)) " by (intro distr_distr) simp_all also have "\ = distr N (count_space (Pow (B \ B))) (map_relation A \ \ linorder_from_keys A)" by (intro distr_cong refl) (auto simp: linorder_from_keys_permute[OF \(1)]) also have "\ = distr M (count_space (Pow (B \ B))) (map_relation A \)" unfolding M_eq using B by (intro distr_distr [symmetric]) (auto simp: map_relation_def) finally have M_eq': "distr M (count_space (Pow (B \ B))) (map_relation A \) = M" .. from that have subset': "R \ B \ B" "R' \ B \ B" using B by (auto simp: linorder_on_def refl_on_def) hence "emeasure M {R} \ emeasure M (map_relation A \ -` {R'} \ space M)" using subset' by (intro emeasure_mono) (auto simp: M_def \ ) also have "\ = emeasure (distr M (count_space (Pow (B \ B))) (map_relation A \)) {R'} " by (rule emeasure_distr [symmetric]) (insert subset' B, auto simp: map_relation_def) also note M_eq' finally show ?thesis . qed have same_prob: "emeasure M {R'} = emeasure M {R}" if "linorder_on A R'" for R' using *[of R R'] and *[of R' R] and R and that by simp from ab have "prob_space M" unfolding M_def by (intro prob_space.prob_space_distr prob_space_PiM prob_space_uniform_measure) simp_all hence "1 = emeasure M (Pow (B \ B))" using prob_space.emeasure_space_1[OF \prob_space M\] by (simp add: M_def) also have "(Pow (B \ B)) = linorders_on A \ ((Pow (B \ B))-linorders_on A)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "emeasure M \ = emeasure M (linorders_on A) + emeasure M (Pow (B \ B)-linorders_on A)" using B by (subst plus_emeasure) (auto simp: M_def linorders_on_def linorder_on_def refl_on_def) also have "emeasure M (Pow (B \ B)-linorders_on A) = 0" using almost_everywhere_linorder by (subst (asm) AE_iff_measurable) (auto simp: linorders_on_def M_def) also from fin have "emeasure M (linorders_on A) = (\R'\linorders_on A. emeasure M {R'})" using B by (intro emeasure_eq_sum_singleton) (auto simp: M_def linorders_on_def linorder_on_def refl_on_def) also have "\ = (\R'\linorders_on A. emeasure M {R})" by (rule sum.cong) (simp_all add: linorders_on_def same_prob) also from fin have "\ = fact (card A) * emeasure M {R}" by (simp add: linorders_on_def card_finite_linorders) finally have [simp]: "emeasure M {R} = inverse (fact (card A))" by (simp add: inverse_ennreal_unique) show ?thesis proof (rule measure_eqI_countable_AE') show "sets M = Pow (Pow (B \ B))" by (simp add: M_def) next from \finite A\ show "countable (linorders_on A)" by (blast intro: countable_finite) next show "AE R in uniform_measure (count_space (Pow (B \ B))) (linorders_on A). R \ linorders_on A" by (rule AE_uniform_measureI) (insert B, auto simp: linorders_on_def linorder_on_def refl_on_def) next fix R' assume R': "R' \ linorders_on A" have subset: "linorders_on A \ Pow (B \ B)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) have "emeasure (uniform_measure (count_space (Pow (B \ B))) (linorders_on A)) {R'} = emeasure (count_space (Pow (B \ B))) (linorders_on A \ {R'}) / emeasure (count_space (Pow (B \ B))) (linorders_on A)" using R' B by (subst emeasure_uniform_measure) (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "\ = 1 / emeasure (count_space (Pow (B \ B))) (linorders_on A)" using R' B by (subst emeasure_count_space) (auto simp: linorders_on_def linorder_on_def refl_on_def) also have "\ = 1 / fact (card A)" using fin finite_linorders_on[of A] by (subst emeasure_count_space [OF subset]) (auto simp: divide_ennreal [symmetric] linorders_on_def card_finite_linorders) also have "\ = emeasure M {R}" by (simp add: field_simps divide_ennreal_def) also have "\ = emeasure M {R'}" using R' by (intro same_prob [symmetric]) (auto simp: linorders_on_def) finally show "emeasure M {R'} = emeasure (uniform_measure (count_space (Pow (B \ B))) (linorders_on A)) {R'}" .. next show "linorders_on A \ Pow (B \ B)" using B by (auto simp: linorders_on_def linorder_on_def refl_on_def) qed (auto simp: M_def linorders_on_def almost_everywhere_linorder) qed end end diff --git a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy --- a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy +++ b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy @@ -1,2089 +1,2089 @@ (* File: Zeta_3_Irrational.thy Author: Manuel Eberl, TU München *) section \The Irrationality of $\zeta(3)$\ theory Zeta_3_Irrational imports "E_Transcendental.E_Transcendental" "Prime_Number_Theorem.Prime_Number_Theorem" "Prime_Distribution_Elementary.PNT_Consequences" begin text \ Ap\'{e}ry's original proof of the irrationality of $\zeta(3)$ contained several tricky identities of sums involving binomial coefficients that are difficult to prove. I instead follow Beukers's proof, which consists of several fairly straightforward manipulations of integrals with absolutely no caveats. Both Ap\'{e}ry and Beukers make use of an asymptotic upper bound on $\text{lcm}\{1\ldots n\}$ -- namely $\text{lcm}\{1\ldots n\} \in o(c^n)$ for any $c > e$, which is a consequence of the Prime Number Theorem (which, fortunately, is available in the \emph{Archive of Formal Proofs}~\cite{afp_primes1,afp_primes2}). I follow the concise presentation of Beukers's proof in Filaseta's lecture notes~\cite{filaseta}, which turned out to be very amenable to formalisation. There is another earlier formalisation of the irrationality of $\zeta(3)$ by Mahboubi\ \emph{et al.}~\cite{mahboubi}, who followed Ap\'{e}ry's original proof, but were ultimately forced to find a more elementary way to prove the asymptotics of $\text{lcm}\{1\ldots n\}$ than the Prime Number Theorem. First, we will require some auxiliary material before we get started with the actual proof. \ (* TODO: A lot of this (if not all of it) should really be in the library *) subsection \Auxiliary facts about polynomials\ lemma degree_higher_pderiv: "degree ((pderiv ^^ n) p) = degree p - n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" by (induction n arbitrary: p) (auto simp: degree_pderiv) lemma pcompose_power_left: "pcompose (p ^ n) q = pcompose p q ^ n" by (induction n) (auto simp: pcompose_mult pcompose_1) lemma pderiv_sum: "pderiv (\x\A. f x) = (\x\A. pderiv (f x))" by (induction A rule: infinite_finite_induct) (auto simp: pderiv_add) lemma higher_pderiv_minus: "(pderiv ^^ n) (-p :: 'a :: idom poly) = -(pderiv ^^ n) p" by (induction n) (auto simp: pderiv_minus) lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n - 1)) * pderiv p" by (cases n) (simp_all add: pderiv_power_Suc del: power_Suc) lemma pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (simp add: monom_altdef pderiv_smult pderiv_power pderiv_pCons mult_ac) lemma higher_pderiv_monom: "k \ n \ (pderiv ^^ k) (monom c n) = monom (of_nat (pochhammer (n - k + 1) k) * c) (n - k)" by (induction k) (auto simp: pderiv_monom pochhammer_rec Suc_diff_le Suc_diff_Suc mult_ac) lemma higher_pderiv_mult: "(pderiv ^^ n) (p * q) = (\k\n. Polynomial.smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (n - k)) q))" proof (induction n) case (Suc n) have eq: "(Suc n choose k) = (n choose k) + (n choose (k-1))" if "k > 0" for k using that by (cases k) auto have "(pderiv ^^ Suc n) (p * q) = (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q))" by (simp add: Suc pderiv_sum pderiv_smult pderiv_mult sum.distrib smult_add_right algebra_simps Suc_diff_le) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) = (\k\insert 0 {1..n}. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q" by (subst sum.insert) (auto simp: add_ac) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q)) = (\k=1..n+1. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto also have "\ = (\k\insert (n+1) {1..n}. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (pderiv ^^ Suc n) p * q" by (subst sum.insert) (auto simp: add_ac) also have "(\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + \ = (\k=1..n. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + (pderiv ^^ Suc n) p * q" by (simp add: sum.distrib algebra_simps smult_add_right eq smult_add_left) also have "\ = (\k\{1..n}\{0,Suc n}. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (subst sum.union_disjoint) (auto simp: algebra_simps) also have "{1..n}\{0,Suc n} = {..Suc n}" by auto finally show ?case . qed auto subsection \Auxiliary facts about integrals\ theorem (in pair_sigma_finite) Fubini_set_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "set_borel_measurable (M1 \\<^sub>M M2) (A \ B) f" and integ1: "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2)" and integ2: "AE x\A in M1. set_integrable M2 B (\y. f (x, y))" shows "set_integrable (M1 \\<^sub>M M2) (A \ B) f" unfolding set_integrable_def proof (rule Fubini_integrable) note integ1 also have "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2) \ integrable M1 (\x. LINT y|M2. norm (indicat_real (A \ B) (x, y) *\<^sub>R f (x, y)))" unfolding set_integrable_def by (intro integrable_cong) (auto simp: indicator_def set_lebesgue_integral_def) finally show \ . next from integ2 show "AE x in M1. integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof eventually_elim case (elim x) show "integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof (cases "x \ A") case True with elim have "set_integrable M2 B (\y. f (x, y))" by simp also have "?this \ ?thesis" unfolding set_integrable_def using True by (intro integrable_cong) (auto simp: indicator_def) finally show ?thesis . qed auto qed qed (insert assms, auto simp: set_borel_measurable_def) lemma (in pair_sigma_finite) set_integral_fst': fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\x\A. (\y\B. f (x, y) \M2) \M1)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\x. \y. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M2 \M1)" using assms by (subst integral_fst' [symmetric]) (auto simp: set_integrable_def) also have "\ = (\x\A. (\y\B. f (x,y) \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) set_integral_snd: fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\y\B. (\x\A. f (x, y) \M1) \M2)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\y. \x. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M1 \M2)" using assms by (subst integral_snd) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\y\B. (\x\A. f (x,y) \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed proposition (in pair_sigma_finite) Fubini_set_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "set_integrable (M1 \\<^sub>M M2) (A \ B) (case_prod f)" shows "(\y\B. (\x\A. f x y \M1) \M2) = (\x\A. (\y\B. f x y \M2) \M1)" proof - have "(\y\B. (\x\A. f x y \M1) \M2) = (\y. (\x. indicator (A \ B) (x, y) *\<^sub>R f x y \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) also have "\ = (\x. (\y. indicator (A \ B) (x, y) *\<^sub>R f x y \M2) \M1)" using assms by (intro Fubini_integral) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\x\A. (\y\B. f x y \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) nn_integral_swap: assumes [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+x. f x \(M1 \\<^sub>M M2)) = (\\<^sup>+(y,x). f (x,y) \(M2 \\<^sub>M M1))" by (subst distr_pair_swap, subst nn_integral_distr) (auto simp: case_prod_unfold) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "set_integrable M A f \ set_borel_measurable M A g \ (AE x in M. x \ A \ norm (g x) \ norm (f x)) \ set_integrable M A g" unfolding set_integrable_def apply (erule Bochner_Integration.integrable_bound) apply (simp add: set_borel_measurable_def) apply (erule eventually_mono) apply (auto simp: indicator_def) done lemma set_integrableI_nonneg: fixes f :: "'a \ real" assumes "set_borel_measurable M A f" assumes "AE x in M. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \M) < \" shows "set_integrable M A f" unfolding set_integrable_def proof (rule integrableI_nonneg) from assms show "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(2) show "AE x in M. 0 \ indicat_real A x *\<^sub>R f x" by eventually_elim (auto simp: indicator_def) have "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) = (\\<^sup>+x\A. f x \M)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ < \" by fact finally show "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) < \" . qed lemma set_integral_eq_nn_integral: assumes "set_borel_measurable M A f" assumes "set_nn_integral M A f = ennreal x" "x \ 0" assumes "AE x in M. x \ A \ f x \ 0" shows "set_integrable M A f" and "set_lebesgue_integral M A f = x" proof - have eq: "(\x. (if x \ A then 1 else 0) * f x) = (\x. if x \ A then f x else 0)" "(\x. if x \ A then ennreal (f x) else 0) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" "(\x. ennreal (f x * (if x \ A then 1 else 0))) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" by auto from assms show *: "set_integrable M A f" by (intro set_integrableI_nonneg) auto have "set_lebesgue_integral M A f = enn2real (set_nn_integral M A f)" unfolding set_lebesgue_integral_def using assms(1,4) * eq by (subst integral_eq_nn_integral) - (auto intro!: nn_integral_cong simp: indicator_def set_integrable_def mult_ac) + (auto intro!: nn_integral_cong simp: indicator_def of_bool_def set_integrable_def mult_ac) also have "\ = x" using assms by simp finally show "set_lebesgue_integral M A f = x" . qed lemma set_integral_0 [simp, intro]: "set_integrable M A (\y. 0)" by (simp add: set_integrable_def) lemma set_integrable_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_integrable M A (\y. \x\B. f x y)" using assms by (induction rule: finite_induct) (auto intro!: set_integral_add) lemma set_integral_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_lebesgue_integral M A (\y. \x\B. f x y) = (\x\B. set_lebesgue_integral M A (f x))" using assms apply (induction rule: finite_induct) apply simp apply simp apply (subst set_integral_add) apply (auto intro!: set_integrable_sum) done lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_nn_integral_mono: assumes "\x. x \ space M \ A \ f x \ g x" shows "set_nn_integral M A f \ set_nn_integral M A g" using assms by (intro nn_integral_mono) (auto simp: indicator_def) lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ (F has_field_derivative f x) (at x within {a..b})" assumes cont: "continuous_on {a..b} f" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_field_derivative_iff_has_vector_derivative[symmetric] using deriv by auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 cont] integral_FTC_Icc[OF \a \ b\ 1 cont] by (auto simp: mult.commute) qed lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" proof - have "integrable lborel (\x. indicator {a..b} x *\<^sub>R ((F x) * (g x) + (f x) * (G x)))" by (intro borel_integrable_compact continuous_intros assms) (auto intro!: DERIV_continuous_on assms) thus ?thesis by (simp add: mult_ac) qed lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: assms DERIV_continuous_on) have [continuous_intros]: "continuous_on {a..b} F" by (rule DERIV_continuous_on assms)+ have [continuous_intros]: "continuous_on {a..b} G" by (rule DERIV_continuous_on assms)+ have "(\x. indicator {a..b} x *\<^sub>R (F x * g x + f x * G x) \lborel) = (\x. indicator {a..b} x *\<^sub>R(F x * g x) \lborel) + \x. indicator {a..b} x *\<^sub>R (f x * G x) \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (simp add: algebra_simps) done thus ?thesis using 0 by (simp add: algebra_simps) qed lemma interval_lebesgue_integral_by_parts: assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(LBINT x=a..b. F x * g x) = F b * G b - F a * G a - (LBINT x=a..b. f x * G x)" using integral_by_parts[of a b f g F G] assms by (simp add: interval_integral_Icc set_lebesgue_integral_def mult_ac) lemma interval_lebesgue_integral_by_parts_01: assumes cont_f[intro]: "continuous_on {0..1} f" assumes cont_g[intro]: "continuous_on {0..1} g" assumes [intro]: "\x. x \ {0..1} \ (F has_field_derivative f x) (at x within {0..1})" assumes [intro]: "\x. x \ {0..1} \ (G has_field_derivative g x) (at x within {0..1})" shows "(LBINT x=0..1. F x * g x) = F 1 * G 1 - F 0 * G 0 - (LBINT x=0..1. f x * G x)" using interval_lebesgue_integral_by_parts[of 0 1 f g F G] assms by (simp add: zero_ereal_def one_ereal_def) lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ real" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed subsection \Shifted Legendre polynomials\ text \ The first ingredient we need to show Apéry's theorem is the \<^emph>\shifted Legendre polynomials\ \[P_n(X) = \frac{1}{n!} \frac{\partial^n}{\partial X^n} (X^n(1-X)^n)\] and the auxiliary polynomials \[Q_{n,k}(X) = \frac{\partial^k}{\partial X^k} (X^n(1-X)^n)\ .\] Note that $P_n$ is in fact an \emph{integer} polynomial. Only some very basic properties of these will be proven, since that is all we will need. \ context fixes n :: nat begin definition gen_shleg_poly :: "nat \ int poly" where "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" definition shleg_poly where "shleg_poly = gen_shleg_poly n div [:fact n:]" text \ We can easily prove the following more explicit formula for $Q_{n,k}$: \[Q_{n,k}(X) = \sum_{i=0}^k (-1)^{k-1} {k \choose i} n^{\underline{i}}\, n^{\underline{k-i}}\, X^{n-i}\, (1-X)^{n-k+i}\] \ lemma gen_shleg_poly_altdef: assumes "k \ n" shows "gen_shleg_poly k = (\i\k. smult ((-1)^(k-i) * of_nat (k choose i) * pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n-k+i)))" proof - have *: "(pderiv ^^ i) (x \\<^sub>p [:1, -1:]) = smult ((-1) ^ i) ((pderiv ^^ i) x \\<^sub>p [:1, -1:])" for i and x :: "int poly" by (induction i arbitrary: x) (auto simp: pderiv_smult pderiv_pcompose funpow_Suc_right pderiv_pCons higher_pderiv_minus simp del: funpow.simps(2)) have "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" by (simp add: gen_shleg_poly_def) also have "[:0, 1, -1::int:] = [:0, 1:] * [:1, -1:]" by simp also have "\ ^ n = [:0, 1:] ^ n * [:1, -1:] ^ n" by (simp flip: power_mult_distrib) also have "(pderiv ^^ k) \ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) ([:0, 1:] ^ n) * (pderiv ^^ (k - i)) ([:1, -1:] ^ n)))" by (simp add: higher_pderiv_mult) also have "\ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * (pderiv ^^ (k - i)) (monom 1 n \\<^sub>p [:1, -1:])))" by (simp add: monom_altdef pcompose_pCons pcompose_power_left) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * ((pderiv ^^ (k - i)) (monom 1 n) \\<^sub>p [:1, -1:])))" by (simp add: * mult_ac) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) (monom (pochhammer (n - i + 1) i) (n - i) * monom (pochhammer (n - k + i + 1) (k - i)) (n - k + i) \\<^sub>p [:1, -1:]))" using assms by (simp add: higher_pderiv_monom) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i) * pochhammer (n - i + 1) i * pochhammer (n - k + i + 1) (k - i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n - k + i)))" by (simp add: monom_altdef algebra_simps pcompose_smult pcompose_power_left pcompose_pCons) finally show ?thesis . qed lemma degree_gen_shleg_poly [simp]: "degree (gen_shleg_poly k) = 2 * n - k" by (simp add: gen_shleg_poly_def degree_higher_pderiv degree_power_eq) lemma gen_shleg_poly_n: "gen_shleg_poly n = smult (fact n) shleg_poly" proof - obtain r where r: "gen_shleg_poly n = [:fact n:] * r" unfolding gen_shleg_poly_def using fact_dvd_higher_pderiv[of n "[:0,1,-1:]^n"] by blast have "smult (fact n) shleg_poly = smult (fact n) (gen_shleg_poly n div [:fact n:])" by (simp add: shleg_poly_def) also note r also have "[:fact n:] * r div [:fact n:] = r" by (rule nonzero_mult_div_cancel_left) auto finally show ?thesis by (simp add: r) qed lemma degree_shleg_poly [simp]: "degree shleg_poly = n" using degree_gen_shleg_poly[of n] by (simp add: gen_shleg_poly_n) lemma pderiv_gen_shleg_poly [simp]: "pderiv (gen_shleg_poly k) = gen_shleg_poly (Suc k)" by (simp add: gen_shleg_poly_def) text \ The following functions are the interpretation of the shifted Legendre polynomials and the auxiliary polynomials as a function from reals to reals. \ definition Gen_Shleg :: "nat \ real \ real" where "Gen_Shleg k x = poly (of_int_poly (gen_shleg_poly k)) x" definition Shleg :: "real \ real" where "Shleg = poly (of_int_poly shleg_poly)" lemma Gen_Shleg_altdef: assumes "k \ n" shows "Gen_Shleg k x = (\i\k. (-1)^(k-i) * of_nat (k choose i) * of_int (pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) * x ^ (n - i) * (1 - x) ^ (n-k+i))" using assms by (simp add: Gen_Shleg_def gen_shleg_poly_altdef poly_sum mult_ac) lemma Gen_Shleg_0 [simp]: "k < n \ Gen_Shleg k 0 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_1 [simp]: "k < n \ Gen_Shleg k 1 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_n_0 [simp]: "Gen_Shleg n 0 = fact n" proof - have "Gen_Shleg n 0 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{n}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (intro sum.mono_neutral_right) auto also have "\ = fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Gen_Shleg_n_1 [simp]: "Gen_Shleg n 1 = (-1) ^ n * fact n" proof - have "Gen_Shleg n 1 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{0}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (intro sum.mono_neutral_right) auto also have "\ = (-1) ^ n * fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Shleg_altdef: "Shleg x = Gen_Shleg n x / fact n" by (simp add: Shleg_def Gen_Shleg_def gen_shleg_poly_n) lemma Shleg_0 [simp]: "Shleg 0 = 1" and Shleg_1 [simp]: "Shleg 1 = (-1) ^ n" by (simp_all add: Shleg_altdef) lemma Gen_Shleg_0_left: "Gen_Shleg 0 x = x ^ n * (1 - x) ^ n" by (simp add: Gen_Shleg_def gen_shleg_poly_def power_mult_distrib) lemma has_field_derivative_Gen_Shleg: "(Gen_Shleg k has_field_derivative Gen_Shleg (Suc k) x) (at x)" proof - note [derivative_intros] = poly_DERIV show ?thesis unfolding Gen_Shleg_def by (rule derivative_eq_intros) auto qed lemma continuous_on_Gen_Shleg: "continuous_on A (Gen_Shleg k)" by (auto simp: Gen_Shleg_def intro!: continuous_intros) lemma continuous_on_Gen_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Gen_Shleg k (f x))" by (rule continuous_on_compose2[OF continuous_on_Gen_Shleg[of UNIV]]) auto lemma continuous_on_Shleg: "continuous_on A Shleg" by (auto simp: Shleg_def intro!: continuous_intros) lemma continuous_on_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Shleg (f x))" by (rule continuous_on_compose2[OF continuous_on_Shleg[of UNIV]]) auto lemma measurable_Gen_Shleg [measurable]: "Gen_Shleg n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Gen_Shleg) lemma measurable_Shleg [measurable]: "Shleg \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Shleg) end subsection \Auxiliary facts about the \\\ function\ lemma Re_zeta_ge_1: assumes "x > 1" shows "Re (zeta (of_real x)) \ 1" proof - have *: "(\n. real (Suc n) powr -x) sums Re (zeta (complex_of_real x))" using sums_Re[OF sums_zeta[of "of_real x"]] assms by (simp add: powr_Reals_eq) show "Re (zeta (of_real x)) \ 1" proof (rule sums_le[OF _ _ *]) show "(\n. if n = 0 then 1 else 0) sums 1" by (rule sums_single) qed auto qed lemma sums_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. 1 / (k + 1) ^ n) sums zeta (of_nat n)" using sums_zeta[of "of_nat n"] n by (simp add: powr_minus field_simps flip: of_nat_Suc) from sums_split_initial_segment[OF this, of r] have "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\kkk=1..r. 1 / k ^ n)" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show ?thesis . qed lemma sums_Re_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (Re (zeta (of_nat n)) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. Re (1 / (r + k + 1) ^ n)) sums (Re (zeta (of_nat n) - (\k=1..r. 1 / k ^ n)))" by (intro sums_Re sums_zeta_of_nat_offset assms) thus ?thesis by simp qed subsection \Divisor of a sum of rationals\ text \ A finite sum of rationals of the form $\frac{a_1}{b_1} + \ldots + \frac{a_n}{b_n}$ can be brought into the form $\frac{c}{d}$, where $d$ is the LCM of the $b_i$ (or some integer multiple thereof). \ lemma sum_rationals_common_divisor: fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" using assms proof (induction rule: finite_induct) case empty thus ?case by auto next case (insert x A) define d where "d = (LCM x\A. g x)" from insert have [simp]: "d \ 0" by (auto simp: d_def Lcm_0_iff) from insert have [simp]: "g x \ 0" by auto from insert obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d" by (auto simp: d_def) define e1 where "e1 = lcm d (g x) div d" define e2 where "e2 = lcm d (g x) div g x" have "(\y\insert x A. f y / g y) = c / d + f x / g x" using insert c by simp also have "c / d = (c * e1) / lcm d (g x)" by (simp add: e1_def real_of_int_div) also have "f x / g x = (f x * e2) / lcm d (g x)" by (simp add: e2_def real_of_int_div) also have "(c * e1) / lcm d (g x) + \ = (c * e1 + f x * e2) / (LCM x\insert x A. g x)" using insert by (simp add: add_divide_distrib lcm.commute d_def) finally show ?case .. qed lemma sum_rationals_common_divisor': fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" and "(\x. x \ A \ g x dvd d)" and "d \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / real_of_int d" proof - define d' where "d' = (LCM x\A. g x)" have "d' dvd d" unfolding d'_def using assms(3) by (auto simp: Lcm_dvd_iff) then obtain e where e: "d = d' * e" by blast have "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" by (rule sum_rationals_common_divisor) fact+ then obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d'" unfolding d'_def .. also have "\ = real_of_int (c * e) / real_of_int d" using \d \ 0\ by (simp add: e) finally show ?thesis .. qed subsection \The first double integral\ text \ We shall now investigate the double integral \[I_1 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy}\,x^r y^s\, \text{d}x\,\text{d}y\ .\] Since everything is non-negative for now, we can work over the extended non-negative real numbers and the issues of integrability or summability do not arise at all. \ definition beukers_nn_integral1 :: "nat \ nat \ ennreal" where "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" definition beukers_integral1 :: "nat \ nat \ real" where "beukers_integral1 r s = (\(x,y)\{0<..<1}\{0<..<1}. (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" lemma fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0..1}" shows beukers_denom_ineq: "(1 - x * y) * z < 1" and beukers_denom_neq: "(1 - x * y) * z \ 1" proof - from xyz have *: "x * y < 1 * 1" by (intro mult_strict_mono) auto from * have "(1 - x * y) * z \ (1 - x * y) * 1" using xyz by (intro mult_left_mono) auto also have "\ < 1 * 1" using xyz by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * z < 1" "(1 - x * y) * z \ 1" by simp_all qed text \ We first evaluate the improper integral \[\int_0^1 -\ln x \cdot x^e\,\text{d}x = \frac{1}{(e+1)^2}\ .\] for any $e>-1$. \ lemma integral_0_1_ln_times_powr: assumes "e > -1" shows "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" and "interval_lebesgue_integrable lborel 0 1 (\x. -ln x * x powr e)" proof - define f where "f = (\x. -ln x * x powr e)" define F where "F = (\x. x powr (e + 1) * (1 - (e + 1) * ln x) / (e + 1) ^ 2)" have 0: "isCont f x" if "x \ {0<..<1}" for x using that by (auto intro!: continuous_intros simp: f_def) have 1: "(F has_real_derivative f x) (at x)" if "x \ {0<..<1}" for x proof - show "(F has_real_derivative f x) (at x)" unfolding F_def f_def using that assms apply (insert that assms) apply (rule derivative_eq_intros refl | simp)+ apply (simp add: divide_simps) apply (simp add: power2_eq_square algebra_simps powr_add power_numeral_reduce) done qed have 2: "AE x in lborel. ereal 0 < ereal x \ ereal x < ereal 1 \ 0 \ f x" by (intro AE_I2) (auto simp: f_def mult_nonpos_nonneg) have 3: "((F \ real_of_ereal) \ 0) (at_right (ereal 0))" unfolding ereal_tendsto_simps F_def using assms by real_asymp have 4: "((F \ real_of_ereal) \ F 1) (at_left (ereal 1))" unfolding ereal_tendsto_simps F_def using assms by real_asymp (simp add: field_simps) have "(LBINT x=ereal 0..ereal 1. f x) = F 1 - 0" by (rule interval_integral_FTC_nonneg[where F = F]) (use 0 1 2 3 4 in auto) thus "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" by (simp add: F_def zero_ereal_def one_ereal_def f_def) have "set_integrable lborel (einterval (ereal 0) (ereal 1)) f" by (rule interval_integral_FTC_nonneg) (use 0 1 2 3 4 in auto) thus "interval_lebesgue_integrable lborel 0 1 f" by (simp add: interval_lebesgue_integrable_def einterval_def) qed lemma interval_lebesgue_integral_lborel_01_cong: assumes "\x. x \ {0<..<1} \ f x = g x" shows "interval_lebesgue_integral lborel 0 1 f = interval_lebesgue_integral lborel 0 1 g" using assms by (subst (1 2) interval_integral_Ioo) (auto intro!: set_lebesgue_integral_cong assms) lemma nn_integral_0_1_ln_times_powr: assumes "e > -1" shows "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = ennreal (1 / (e + 1)\<^sup>2)" proof - have *: "(LBINT x=0..1. - ln x * x powr e = 1 / (e + 1)\<^sup>2)" "interval_lebesgue_integrable lborel 0 1 (\x. - ln x * x powr e)" using integral_0_1_ln_times_powr[OF assms] by simp_all have eq: "(\y. (if 0 < y \ y < 1 then 1 else 0) * ln y * y powr e) = (\y. if 0 < y \ y < 1 then ln y * y powr e else 0)" by auto have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = (\\<^sup>+y. ennreal (-ln y * y powr e * indicator {0<..<1} y) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal (1 / (e + 1)\<^sup>2)" using * eq by (subst nn_integral_eq_integral) (auto intro!: AE_I2 simp: indicator_def interval_lebesgue_integrable_def set_integrable_def one_ereal_def zero_ereal_def interval_integral_Ioo mult_ac mult_nonpos_nonneg set_lebesgue_integral_def) finally show ?thesis . qed lemma nn_integral_0_1_ln_times_power: "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = ennreal (1 / (n + 1)\<^sup>2)" proof - have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr real n) \lborel)" by (intro set_nn_integral_cong) (auto simp: powr_realpow) also have "\ = ennreal (1 / (n + 1)^2)" by (subst nn_integral_0_1_ln_times_powr) auto finally show ?thesis by simp qed text \ Next, we also evaluate the more trivial integral \[\int_0^1 x^n\ \text{d}x\ .\] \ lemma nn_integral_0_1_power: "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = ennreal (1 / (n + 1))" proof - have *: "((\a. a ^ (n + 1) / real (n + 1)) has_real_derivative x ^ n) (at x)" for x by (rule derivative_eq_intros refl | simp)+ have "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = (\\<^sup>+y\{0..1}. ennreal (y ^ n) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: indicator_def emeasure_lborel_countable) also have "\ = ennreal (1 ^ (n + 1) / (n + 1) - 0 ^ (n + 1) / (n + 1))" using * by (intro nn_integral_FTC_Icc) auto also have "\ = ennreal (1 / (n + 1))" by simp finally show ?thesis by simp qed text \ $I_1$ can alternatively be written as the triple integral \[\int_0^1\int_0^1\int_0^1 \frac{x^r y^s}{1-(1-xy)w}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ lemma beukers_nn_integral1_altdef: "beukers_nn_integral1 r s = (\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel)" proof - have "(\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\\<^sup>+w\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) \lborel)" by (subst lborel_prod [symmetric], subst lborel_pair.nn_integral_snd [symmetric]) (auto simp: case_prod_unfold indicator_def simp flip: lborel_prod intro!: nn_integral_cong) also have "\ = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y)/(1-x*y) * x^r * y^s) \lborel)" proof (intro nn_integral_cong, clarify) fix x y :: real have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (-ln (x*y)*x^r*y^s/(1-x*y))" if xy: "(x, y) \ {0<..<1} \ {0<..<1}" proof - from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have deriv: "((\w. -ln (1-(1-x*y)*w) / (1-x*y)) has_real_derivative 1/(1-(1-x*y)*w)) (at w)" if w: "w \ {0..1}" for w by (insert xy w \x*y<1\ beukers_denom_ineq[of x y w]) (rule derivative_eq_intros refl | simp add: divide_simps)+ have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (x^r*y^s) * (\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" using xy by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = (\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: emeasure_lborel_countable indicator_def) also have "(\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = ennreal (-ln (1-(1-x*y)*1)/(1-x*y) - (-ln (1-(1-x*y)*0)/(1-x*y)))" using xy deriv less_imp_le[OF beukers_denom_ineq[of x y]] by (intro nn_integral_FTC_Icc) auto finally show ?thesis using xy by (simp flip: ennreal_mult' ennreal_mult'' add: mult_ac) qed thus "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) * indicator ({0<..<1}\{0<..<1}) (x, y) = ennreal (-ln (x*y)/(1-x*y)*x^r*y^s) * indicator ({0<..<1}\{0<..<1}) (x, y)" by (auto simp: indicator_def) qed also have "\ = beukers_nn_integral1 r s" by (simp add: beukers_nn_integral1_def) finally show ?thesis .. qed context fixes r s :: nat and I1 I2' :: real and I2 :: ennreal and D :: "(real \ real \ real) set" assumes rs: "s \ r" defines "D \ {0<..<1}\{0<..<1}\{0<..<1}" begin text \ By unfolding the geometric series, pulling the summation out and evaluating the integrals, we find that \[I_1 = \sum_{k=0}^\infty \frac{1}{(k+r+1)^2(k+s+1)} + \frac{1}{(k+r+1)(k+s+1)^2}\ .\] \ lemma beukers_nn_integral1_series: "beukers_nn_integral1 r s = (\k. ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2)))" proof - have "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel)" unfolding beukers_nn_integral1_def proof (intro set_nn_integral_cong refl, clarify) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have "(\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) = ennreal (-ln (x*y) * x^r * y^s) * (\k. ennreal ((x*y)^k))" using xy by (subst ennreal_suminf_cmult [symmetric], subst ennreal_mult'' [symmetric]) (auto simp: power_add mult_ac power_mult_distrib) also have "(\k. ennreal ((x*y)^k)) = ennreal (1 / (1 - x*y))" using geometric_sums[of "x*y"] \x * y < 1\ xy by (intro suminf_ennreal_eq) auto also have "ennreal (-ln (x*y) * x^r * y^s) * \ = ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s)" using \x * y < 1\ by (subst ennreal_mult'' [symmetric]) auto finally show "ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) = (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)))" .. qed also have "\ = (\k. (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel))" unfolding case_prod_unfold by (subst nn_integral_suminf [symmetric]) (auto simp flip: borel_prod) also have "\ = (\k. ennreal (1 / ((k+r+1)^2*(k+s+1)) + 1 / ((k+r+1)*(k+s+1)^2)))" proof (rule suminf_cong) fix k :: nat define F where "F = (\x y::real. x + y)" have "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+x\{0<..<1}. (\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) \lborel)" unfolding case_prod_unfold lborel_prod [symmetric] by (subst lborel.nn_integral_fst [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2) \lborel)" proof (intro set_nn_integral_cong refl, clarify) fix x :: real assume x: "x \ {0<..<1}" have "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+y\{0<..<1}. (ennreal (-ln x * x^(k+r) * y^(k+s)) + ennreal (-ln y * x^(k+r) * y^(k+s))) \lborel)" by (intro set_nn_integral_cong) (use x in \auto simp: ln_mult ring_distribs mult_nonpos_nonneg simp flip: ennreal_plus\) also have "\ = (\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) + (\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult'') also have "(\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel) = ennreal (1/(k+s+1))" by (subst nn_integral_0_1_power) simp also have "ennreal (-ln x * x^(k+r)) * \ = ennreal (-ln x * x^(k+r) / (k+s+1))" by (subst ennreal_mult'' [symmetric]) auto also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel) = ennreal (x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (use x in \auto intro!: nn_integral_cong simp: indicator_def mult_ac simp flip: ennreal_mult'\) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel) = ennreal (1 / (k + s + 1)\<^sup>2)" by (subst nn_integral_0_1_ln_times_power) simp also have "ennreal (x ^ (k + r)) * \ = ennreal (x ^ (k + r) / (k + s + 1) ^ 2)" by (subst ennreal_mult'' [symmetric]) auto also have "ennreal (- ln x * x ^ (k + r) / (k + s + 1)) + \ = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" using x by (subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) finally show "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" . qed also have "\ = (\\<^sup>+x\{0<..<1}. (ennreal (-ln x * x^(k+r) / (k+s+1)) + ennreal (x^(k+r)/(k+s+1)^2)) \lborel)" by (intro set_nn_integral_cong refl, subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) + (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) = ennreal (1 / (k+s+1)) * (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1 / ((k+s+1) * (k+r+1)^2))" by (subst nn_integral_0_1_ln_times_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "(\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel) = ennreal (1/(k+s+1)^2) * (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1/((k+r+1)*(k+s+1)^2))" by (subst nn_integral_0_1_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "ennreal (1 / ((k+s+1) * (k+r+1)^2)) + \ = ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2))" by (subst ennreal_plus [symmetric]) (auto simp: algebra_simps) finally show "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = \" . qed finally show ?thesis . qed text \ Remembering that $\zeta(3) = \sum k^{-3}$, it is easy to see that if $r = s$, this sum is simply \[2\left(\zeta(3) - \sum_{k=1}^r \frac{1}{k^3}\right)\ .\] \ lemma beukers_nn_integral1_same: assumes "r = s" shows "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" and "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" proof - from assms have [simp]: "s = r" by simp have *: "Suc 2 = 3" by simp have "beukers_nn_integral1 r s = (\k. ennreal (2 / (r + k + 1) ^ 3))" unfolding beukers_nn_integral1_series by (simp only: assms power_Suc [symmetric] mult.commute[of "x ^ 2" for x] * times_divide_eq_right mult_1_right add_ac flip: mult_2) also have **: "(\k. 2 / (r + k + 1) ^ 3) sums (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" using sums_mult[OF sums_Re_zeta_of_nat_offset[of 3], of 2] by simp hence "(\k. ennreal (2 / (r + k + 1) ^ 3)) = ennreal \" by (intro suminf_ennreal_eq) auto finally show "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" . show "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" by (rule sums_le[OF _ sums_zero **]) auto qed lemma beukers_integral1_same: assumes "r = s" shows "beukers_integral1 r s = 2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3))" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_same[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed text \ In contrast, for \r > s\, we find that \[I_1 = \frac{1}{r-s} \sum_{k=s+1}^r \frac{1}{k^2}\ .\] \ lemma beukers_nn_integral1_different: assumes "r > s" shows "beukers_nn_integral1 r s = ennreal ((\k\{s<..r}. 1 / k ^ 2) / (r - s))" proof - have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) sums (1 / (r - s) * ((Re (zeta (of_nat 2)) - (\k=1..s. 1/k^2)) - (Re (zeta (of_nat 2)) - (\k=1..r. 1/k^2))))" (is "_ sums ?S") by (intro sums_mult sums_diff sums_Re_zeta_of_nat_offset) auto also have "?S = ((\k=1..r. 1/k^2) - (\k=1..s. 1/k^2)) / (r - s)" by (simp add: algebra_simps diff_divide_distrib) also have "(\k=1..r. 1/k^2) - (\k=1..s. 1/k^2) = (\k\{1..r}-{1..s}. 1/k^2)" using assms by (subst Groups_Big.sum_diff) auto also have "{1..r} - {1..s} = {s<..r}" by auto also have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) = (\k. 1 / ((k+r+1) * (k+s+1)^2) + 1 / ((k+r+1)^2 * (k+s+1)))" proof (intro ext, goal_cases) case (1 k) define x where "x = real (k + r + 1)" define y where "y = real (k + s + 1)" have [simp]: "x \ 0" "y \ 0" by (auto simp: x_def y_def) have "(x\<^sup>2 * y + x * y\<^sup>2) * (real r - real s) = x * y * (x\<^sup>2 - y\<^sup>2)" by (simp add: algebra_simps power2_eq_square x_def y_def) hence "1 / (x*y^2) + 1 / (x^2*y) = 1 / (r - s) * (1 / y^2 - 1 / x^2)" using assms by (simp add: divide_simps of_nat_diff) thus ?case by (simp add: x_def y_def algebra_simps) qed finally show ?thesis unfolding beukers_nn_integral1_series by (intro suminf_ennreal_eq) (auto simp: add_ac) qed lemma beukers_integral1_different: assumes "r > s" shows "beukers_integral1 r s = (\k\{s<..r}. 1 / k ^ 2) / (r - s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_different[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg intro!: sum_nonneg divide_nonneg_nonneg) qed end text \ It is also easy to see that if we exchange \r\ and \s\, nothing changes. \ lemma beukers_nn_integral1_swap: "beukers_nn_integral1 r s = beukers_nn_integral1 s r" unfolding beukers_nn_integral1_def lborel_prod [symmetric] by (subst lborel_pair.nn_integral_swap, simp) (intro nn_integral_cong, auto simp: indicator_def algebra_simps split: if_splits) lemma beukers_nn_integral1_finite: "beukers_nn_integral1 r s < \" using beukers_nn_integral1_different[of r s] beukers_nn_integral1_different[of s r] by (cases r s rule: linorder_cases) (simp_all add: beukers_nn_integral1_same beukers_nn_integral1_swap) lemma beukers_integral1_integrable: "set_integrable lborel ({0<..<1}\{0<..<1}) (\(x,y). (-ln (x*y) / (1 - x*y) * x^r * y^s :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" have "0 \ ln (x * y) / (1 - x * y) * x ^ r * y ^ s" using mult_strict_mono[of x 1 y 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) (use xy in auto) thus "0 \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s" by simp next show "\\<^sup>+x\{0<..<1}\{0<..<1}. ennreal (case x of (x, y) \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_def case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_integrable': "set_integrable lborel ({0<..<1}\{0<..<1}\{0<..<1}) (\(z,x,y). (x^r * y^s / (1 - (1 - x*y) * z) :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" show "0 \ x^r * y^s / (1 - (1 - x*y) * z)" using mult_strict_mono[of x 1 y 1] xyz beukers_denom_ineq[of x y z] by (intro mult_nonneg_nonneg divide_nonneg_nonneg) auto next show "\\<^sup>+x\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (case x of (z,x,y) \ x ^ r * y ^ s / (1-(1-x*y)*z)) \lborel < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_altdef case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_conv_nn_integral: "beukers_integral1 r s = enn2real (beukers_nn_integral1 r s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using mult_strict_mono[of a 1 b 1] that by (intro divide_nonpos_nonneg mult_nonpos_nonneg) auto thus ?thesis unfolding beukers_integral1_def using beukers_nn_integral1_finite[of r s] by (intro set_integral_eq_nn_integral) (auto simp: case_prod_unfold beukers_nn_integral1_def set_borel_measurable_def simp flip: borel_prod intro!: AE_I2 intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed lemma beukers_integral1_swap: "beukers_integral1 r s = beukers_integral1 s r" by (simp add: beukers_integral1_conv_nn_integral beukers_nn_integral1_swap) subsection \The second double integral\ context fixes n :: nat fixes D :: "(real \ real) set" and D' :: "(real \ real \ real) set" fixes P :: "real \ real" and Q :: "nat \ real \ real" defines "D \ {0<..<1} \ {0<..<1}" and "D' \ {0<..<1} \ {0<..<1} \ {0<..<1}" defines "Q \ Gen_Shleg n" and "P \ Shleg n" begin text \ The next integral to consider is the following variant of $I_1$: \[I_2 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy} P_n(x) P_n(y)\,\text{d}x\,\text{d}y\ .\] \ definition beukers_integral2 :: real where "beukers_integral2 = (\(x,y)\D. (-ln (x*y) / (1-x*y) * P x * P y) \lborel)" text \ $I_2$ is simply a sum of integrals of type $I_1$, so using our results for $I_1$, we can write $I_2$ in the form $A \zeta(3) + \frac{B}{\text{lcm}\{1\ldots n\}^3}$ where $A$ and $B$ are integers and $A > 0$: \ lemma beukers_integral2_conv_int_combination: obtains A B :: int where "A > 0" and "beukers_integral2 = of_int A * Re (zeta 3) + of_int B / of_nat (Lcm {1..n} ^ 3)" proof - let ?I1 = "(\i. (i, i)) ` {..n}" let ?I2 = "Set.filter (\(i,j). i \ j) ({..n}\{..n})" let ?I3 = "Set.filter (\(i,j). i < j) ({..n}\{..n})" let ?I4 = "Set.filter (\(i,j). i > j) ({..n}\{..n})" define p where "p = shleg_poly n" define I where "I = (SIGMA i:{..n}. {1..i})" define J where "J = (SIGMA (i,j):?I4. {j<..i})" define h where "h = beukers_integral1" define A :: int where "A = (\i\n. 2 * poly.coeff p i ^ 2)" define B1 where "B1 = (\(i,k)\I. real_of_int (-2 * coeff p i ^ 2) / real_of_int (k ^ 3))" define B2 where "B2 = (\((i,j),k)\J. real_of_int (2 * coeff p i * coeff p j) / real_of_int (k^2*(i-j)))" define d where "d = Lcm {1..n} ^ 3" have [simp]: "h i j = h j i" for i j by (simp add: h_def beukers_integral1_swap) have "beukers_integral2 = (\(x,y)\D. (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * -ln (x*y) / (1-x*y) * x ^ i * y ^ j) \lborel)" unfolding beukers_integral2_def by (subst sum.cartesian_product [symmetric]) (simp add: poly_altdef P_def Shleg_def mult_ac case_prod_unfold p_def sum_distrib_left sum_distrib_right sum_negf sum_divide_distrib) also have "\ = (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * h i j)" unfolding case_prod_unfold proof (subst set_integral_sum) fix ij :: "nat \ nat" have "set_integrable lborel D (\(x,y). real_of_int (coeff p (fst ij) * coeff p (snd ij)) * (-ln (x*y) / (1-x*y) * x ^ fst ij * y ^ snd ij))" unfolding case_prod_unfold using beukers_integral1_integrable[of "fst ij" "snd ij"] by (intro set_integrable_mult_right) (auto simp: D_def case_prod_unfold) thus "set_integrable lborel D (\pa. real_of_int (coeff p (fst ij) * coeff p (snd ij)) * -ln (fst pa * snd pa) / (1 - fst pa * snd pa) * fst pa ^ fst ij * snd pa ^ snd ij)" by (simp add: mult_ac case_prod_unfold) qed (auto simp: beukers_integral1_def h_def case_prod_unfold mult.assoc D_def simp flip: set_integral_mult_right) also have "\ = (\(i,j)\?I1\?I2. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I1. coeff p i * coeff p j * h i j) + (\(i,j)\?I2. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I1. coeff p i * coeff p j * h i j) = (\i\n. coeff p i ^ 2 * h i i)" by (subst sum.reindex) (auto intro: inj_onI simp: case_prod_unfold power2_eq_square) also have "\ = (\i\n. coeff p i ^ 2 * 2 * (Re (zeta 3) - (\k=1..i. 1 / k ^ 3)))" unfolding h_def D_def by (intro sum.cong refl, subst beukers_integral1_same) auto also have "\ = of_int A * Re (zeta 3) - (\i\n. 2 * coeff p i ^ 2 * (\k=1..i. 1 / k ^ 3))" by (simp add: sum_subtractf sum_distrib_left sum_distrib_right algebra_simps A_def) also have "\ = of_int A * Re (zeta 3) + B1" unfolding I_def B1_def by (subst sum.Sigma [symmetric]) (auto simp: sum_distrib_left sum_negf) also have "(\(i,j)\?I2. coeff p i * coeff p j * h i j) = (\(i,j)\?I3\?I4. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I3. coeff p i * coeff p j * h i j) + (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I3. coeff p i * coeff p j * h i j) = (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.reindex_bij_witness[of _ "\(i,j). (j,i)" "\(i,j). (j,i)"]) auto also have "\ + \ = 2 * \" by simp also have "\ = (\(i,j)\?I4. \k\{j<..i}. 2 * coeff p i * coeff p j / k ^ 2 / (i - j))" unfolding sum_distrib_left by (intro sum.cong refl) (auto simp: h_def beukers_integral1_different sum_divide_distrib sum_distrib_left mult_ac) also have "\ = B2" unfolding J_def B2_def by (subst sum.Sigma [symmetric]) (auto simp: case_prod_unfold) also have "\B1'. B1 = real_of_int B1' / real_of_int d" unfolding B1_def case_prod_unfold by (rule sum_rationals_common_divisor') (auto simp: d_def I_def) then obtain B1' where "B1 = real_of_int B1' / real_of_int d" .. also have "\B2'. B2 = real_of_int B2' / real_of_int d" unfolding B2_def case_prod_unfold J_def proof (rule sum_rationals_common_divisor'; clarsimp?) fix i j k :: nat assume ijk: "i \ n" "j < k" "k \ i" have "int (k ^ 2 * (i - j)) dvd int (Lcm {1..n} ^ 2 * Lcm {1..n})" unfolding int_dvd_int_iff using ijk by (intro mult_dvd_mono dvd_power_same dvd_Lcm) auto also have "\ = d" by (simp add: d_def power_numeral_reduce) finally show "int k ^ 2 * int (i - j) dvd int d" by simp qed(auto simp: d_def J_def intro!: Nat.gr0I) then obtain B2' where "B2 = real_of_int B2' / real_of_int d" .. finally have "beukers_integral2 = of_int A * Re (zeta 3) + of_int (B1' + B2') / of_nat (Lcm {1..n} ^ 3)" by (simp add: add_divide_distrib d_def) moreover have "coeff p 0 = P 0" unfolding P_def p_def Shleg_def by (simp add: poly_0_coeff_0) hence "coeff p 0 = 1" by (simp add: P_def) hence "A > 0" unfolding A_def by (intro sum_pos2[of _ 0]) auto ultimately show ?thesis by (intro that[of A "B1' + B2'"]) auto qed lemma beukers_integral2_integrable: "set_integrable lborel D (\(x,y). -ln (x*y) / (1 - x*y) * P x * P y)" proof - have "bounded (P ` {0..1})" unfolding P_def Shleg_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast have [measurable]: "P \ borel_measurable borel" by (simp add: P_def) from C[of 0] have "C \ 0" by simp show ?thesis proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C ^ 2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y :: real assume xy: "(x, y) \ D" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by (simp add: D_def) have "norm (-ln (x*y) / (1 - x*y) * P x * P y) = (-ln (x*y)) / (1 - x*y) * norm (P x) * norm (P y)" using xy \x * y < 1\ by (simp add: abs_mult abs_divide D_def) also have "\ \ (-ln (x*y)) / (1-x*y) * C * C" using xy C[of x] C[of y] \x * y < 1\ \C \ 0\ by (intro mult_mono divide_left_mono) (auto simp: D_def divide_nonpos_nonneg mult_nonpos_nonneg) also have "\ = norm ((-ln (x*y)) / (1-x*y) * C * C)" using xy \x * y < 1\ \C \ 0\ by (simp add: abs_divide abs_mult D_def) finally show "norm (-ln (x*y) / (1 - x*y) * P x * P y) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (auto simp: algebra_simps power2_eq_square abs_mult abs_divide) qed (auto simp: D_def set_borel_measurable_def case_prod_unfold simp flip: lborel_prod) qed subsection \The triple integral\ text \ Lastly, we turn to the triple integral \[I_3 := \int_0^1 \int_0^1 \int_0^1 \frac{(x(1-x)y(1-y)w(1-w))^n}{(1-(1-xy)w)^{n+1}}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ definition beukers_nn_integral3 :: ennreal where "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" definition beukers_integral3 :: real where "beukers_integral3 = (\(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" text \ We first prove the following bound (which is a consequence of the arithmetic--geometric mean inequality) that will help us to bound the triple integral. \ lemma beukers_integral3_integrand_bound: fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" shows "(x*(1-x)*y*(1-y)*z*(1-z)) / (1-(1-x*y)*z) \ 1 / 27" (is "?lhs \ _") proof - have ineq1: "x * (1 - x) \ 1 / 4" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) - 1 / 4 = -((x - 1 / 2) ^ 2)" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by simp finally show ?thesis by simp qed have ineq2: "x * (1 - x) ^ 2 \ 4 / 27" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) ^ 2 - 4 / 27 = (x - 4 / 3) * (x - 1 / 3) ^ 2" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by (rule mult_nonpos_nonneg) (use x in auto) finally show ?thesis by simp qed have "1 - (1-x*y)*z = (1 - z) + x * y * z" by (simp add: algebra_simps) also have "\ \ 2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z" using arith_geo_mean_sqrt[of "1 - z" "x * y * z"] xyz by (auto simp: real_sqrt_mult) finally have *: "?lhs \ (x*(1-x)*y*(1-y)*z*(1-z)) / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z)" using xyz beukers_denom_ineq[of x y z] by (intro divide_left_mono mult_nonneg_nonneg mult_pos_pos) auto have "(x*(1-x)*y*(1-y)*z*(1-z)) = (sqrt x * sqrt x * (1-x) * sqrt y * sqrt y * (1-y) * sqrt z * sqrt z * sqrt (1-z) * sqrt (1-z))" using xyz by simp also have "\ / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z) = sqrt (x * (1 - x) ^ 2) * sqrt (y * (1 - y) ^ 2) * sqrt (z * (1 - z)) / 2" using xyz by (simp add: divide_simps real_sqrt_mult del: real_sqrt_mult_self) also have "\ \ sqrt (4 / 27) * sqrt (4 / 27) * sqrt (1 / 4) / 2" using xyz by (intro divide_right_mono mult_mono real_sqrt_le_mono ineq1 ineq2) auto also have "\ = 1 / 27" by (simp add: real_sqrt_divide) finally show ?thesis using * by argo qed text \ Connecting the above bound with our results of $I_1$, it is easy to see that $I_3 \leq 2 \cdot 27^{-n} \cdot \zeta(3)$: \ lemma beukers_nn_integral3_le: "beukers_nn_integral3 \ ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" by (simp add: beukers_nn_integral3_def) also have "\ \ (\\<^sup>+(w,x,y)\D'. ((1 / 27) ^ n / (1-(1-x*y)*w)) \lborel)" proof (intro set_nn_integral_mono ennreal_leI, clarify, goal_cases) case (1 w x y) have "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) = ((x*(1-x)*y*(1-y)*w*(1-w)) / (1-(1-x*y)*w)) ^ n / (1-(1-x*y)*w)" by (simp add: divide_simps) also have "\ \ (1 / 27) ^ n / (1 - (1 - x * y) * w)" using beukers_denom_ineq[of x y w] 1 by (intro divide_right_mono power_mono beukers_integral3_integrand_bound) (auto simp: D'_def) finally show ?case . qed also have "\ = ennreal ((1 / 27) ^ n) * (\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel)" unfolding lborel_prod [symmetric] case_prod_unfold by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (- (ln (x * y) / (1 - x * y))) \lborel)" using beukers_nn_integral1_altdef[of 0 0] by (simp add: beukers_nn_integral1_def D'_def case_prod_unfold) also have "\ = ennreal (2 * Re (zeta 3))" using beukers_nn_integral1_same[of 0 0] by (simp add: D_def beukers_nn_integral1_def) also have "ennreal ((1 / 27) ^ n) * \ = ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" by (subst ennreal_mult' [symmetric]) (simp_all add: mult_ac) finally show ?thesis . qed lemma beukers_nn_integral3_finite: "beukers_nn_integral3 < \" by (rule le_less_trans, rule beukers_nn_integral3_le) simp_all lemma beukers_integral3_integrable: "set_integrable lborel D' (\(w,x,y). (x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1))" unfolding case_prod_unfold using less_imp_le[OF beukers_denom_ineq] beukers_nn_integral3_finite by (intro set_integrableI_nonneg AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod intro!: divide_nonneg_nonneg mult_nonneg_nonneg) lemma beukers_integral3_conv_nn_integral: "beukers_integral3 = enn2real beukers_nn_integral3" unfolding beukers_integral3_def using beukers_nn_integral3_finite less_imp_le[OF beukers_denom_ineq] by (intro set_integral_eq_nn_integral AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod) lemma beukers_integral3_le: "beukers_integral3 \ 2 * (1 / 27) ^ n * Re (zeta 3)" proof - have "beukers_integral3 = enn2real beukers_nn_integral3" by (rule beukers_integral3_conv_nn_integral) also have "\ \ enn2real (ennreal (2 * (1 / 27) ^ n * Re (zeta 3)))" by (intro enn2real_mono beukers_nn_integral3_le) auto also have "\ = 2 * (1 / 27) ^ n * Re (zeta 3)" using Re_zeta_ge_1[of 3] by (intro enn2real_ennreal mult_nonneg_nonneg) auto finally show ?thesis . qed text \ It is also easy to see that $I_3 > 0$. \ lemma beukers_nn_integral3_pos: "beukers_nn_integral3 > 0" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have *: "\(AE (w,x,y) in lborel. ennreal ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) * indicator D' (w,x,y) \ 0)" (is "\(AE z in lborel. ?P z)") proof - { fix w x y :: real assume xyw: "(w,x,y) \ D'" hence "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) > 0" using beukers_denom_ineq[of x y w] by (intro divide_pos_pos mult_pos_pos zero_less_power) (auto simp: D'_def) with xyw have "\?P (w,x,y)" by (auto simp: indicator_def D'_def) } hence *: "\?P z" if "z \ D'" for z using that by blast hence "{z\space lborel. \?P z} = D'" by auto moreover have "emeasure lborel D' = 1" proof - have "D' = box (0,0,0) (1,1,1)" by (auto simp: D'_def box_def Basis_prod_def) also have "emeasure lborel \ = 1" by (subst emeasure_lborel_box) (auto simp: Basis_prod_def) finally show ?thesis by simp qed ultimately show ?thesis by (subst AE_iff_measurable[of D']) (simp_all flip: borel_prod) qed hence "nn_integral lborel (\_::real\real\real. 0) < beukers_nn_integral3" unfolding beukers_nn_integral3_def by (intro nn_integral_less) (simp_all add: case_prod_unfold flip: lborel_prod) thus ?thesis by simp qed lemma beukers_integral3_pos: "beukers_integral3 > 0" proof - have "0 < enn2real beukers_nn_integral3" using beukers_nn_integral3_pos beukers_nn_integral3_finite by (subst enn2real_positive_iff) auto also have "\ = beukers_integral3" by (rule beukers_integral3_conv_nn_integral [symmetric]) finally show ?thesis . qed subsection \Connecting the double and triple integral\ text \ In this section, we will prove the most technically involved part, namely that $I_2 = I_3$. I will not go into detail about how this works -- the reader is advised to simply look at Filaseta's presentation of the proof. The basic idea is to integrate by parts \n\ times with respect to \y\ to eliminate the factor $P(y)$, then change variables $z = \frac{1-w}{1-(1-xy)w}$, and then apply the same integration by parts \n\ times to \x\ to eliminate $P(x)$. \ text \ The first expand \[-\frac{\ln (xy)}{1-xy} = \int_0^1 \frac{1}{1-(1-xy)z}\,\text{d}z\ .\] \ lemma beukers_aux_ln_conv_integral: fixes x y :: real assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "-ln (x*y) / (1-x*y) = (LBINT z=0..1. 1 / (1-(1-x*y)*z))" proof - have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - x * y) * u < 1" if u: "u \ {0..1}" for u proof - from u \x * y < 1\ have "(1 - x * y) * u \ (1 - x * y) * 1" by (intro mult_left_mono) auto also have "\ < 1 * 1" using xy by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * u < 1" by simp qed have neq: "(1 - x * y) * u \ 1" if "u \ {0..1}" for u using less[of u] that by simp let ?F = "\z. ln (1-(1-x*y)*z)/(x*y-1)" have "(LBINT z=ereal 0..ereal 1. 1 / (1-(1-x*y)*z)) = ?F 1 - ?F 0" proof (rule interval_integral_FTC_finite, goal_cases cont deriv) case cont show ?case using neq by (intro continuous_intros) auto next case (deriv z) show ?case unfolding has_field_derivative_iff_has_vector_derivative [symmetric] by (insert less[of z] xy \x * y < 1\ deriv) (rule derivative_eq_intros refl | simp)+ qed also have "\ = -ln (x*y) / (1-x*y)" using \x * y < 1\ by (simp add: field_simps) finally show ?thesis by (simp add: zero_ereal_def one_ereal_def) qed text \ The first part we shall show is the integration by parts. \ lemma beukers_aux_by_parts_aux: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" and "k \ n" shows "(LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z))) = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" using assms(3) proof (induction k) case (Suc k) note [derivative_intros] = DERIV_chain2[OF has_field_derivative_Gen_Shleg] define G where "G = (\y. -fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1))" define g where "g = (\y. fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))" have less: "(1 - x * y) * z < 1" and neq: "(1 - x * y) * z \ 1" if y: "y \ {0..1}" for y proof - from y xz have "x * y \ x * 1" by (intro mult_left_mono) auto also have "\ < 1" using xz by simp finally have "(1 - x * y) * z \ 1 * z" using xz y by (intro mult_right_mono) auto also have "\ < 1" using xz by simp finally show "(1 - x * y) * z < 1" by simp thus "(1 - x * y) * z \ 1" by simp qed have cont: "continuous_on {0..1} g" using neq by (auto simp: g_def intro!: continuous_intros) have deriv: "(G has_real_derivative g y) (at y within {0..1})" if y: "y \ {0..1}" for y unfolding G_def by (insert neq xz y, (rule derivative_eq_intros refl power_not_zero)+) (auto simp: divide_simps g_def) have deriv2: "(Q (n - Suc k) has_real_derivative Q (n - k) y) (at y within {0..1})" for y using Suc.prems by (auto intro!: derivative_eq_intros simp: Suc_diff_Suc Q_def) have "(LBINT y=0..1. Q (n-Suc k) y * (fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))) = (LBINT y=0..1. Q (n-Suc k) y * g y)" by (simp add: g_def) also have "(LBINT y=0..1. Q (n-Suc k) y * g y) = -(LBINT y=0..1. Q (n-k) y * G y)" using Suc.prems deriv deriv2 cont by (subst interval_lebesgue_integral_by_parts_01[where f = "Q (n-k)" and G = G]) (auto intro!: continuous_intros simp: Q_def) also have "\ = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" by (simp add: G_def flip: interval_lebesgue_integral_uminus) finally show ?case using Suc by simp qed auto lemma beukers_aux_by_parts: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" shows "(LBINT y=0..1. P y / (1-(1-x*y)*z)) = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have "(LBINT y=0..1. P y * (1/(1-(1-x*y)*z))) = 1 / fact n * (LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z)))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: P_def Q_def Shleg_altdef) also have "\ = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" by (subst beukers_aux_by_parts_aux[OF assms, of n], simp, subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: Q_def mult_ac Gen_Shleg_0_left power_mult_distrib) finally show ?thesis by simp qed text \ We then get the following by applying the integration by parts to \y\: \ lemma beukers_aux_integral_transform1: fixes z :: real assumes z: "z \ {0<..<1}" shows "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "1 - (1 - fst p * snd p) * z \ 0" by simp qed hence integrable: "set_integrable lborel (box (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" by (rule set_integrable_subset) (auto simp: box simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "(1 - (1 - fst p * snd p) * z) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (rule set_integrable_subset) (auto simp: box D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT x=0..1. (LBINT y=0..1. P x * P y / (1-(1-x*y)*z)))" unfolding D_def lborel_prod [symmetric] using box integrable by (subst lborel_pair.set_integral_fst') (simp_all add: interval_integral_Ioo lborel_prod) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using z by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) auto also have "\ = (LBINT x=0..1. (LBINT y=0..1. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding D_def lborel_prod [symmetric] using box integrable' by (subst lborel_pair.set_integral_fst') (simp_all add: D_def interval_integral_Ioo lborel_prod) finally show "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = \" . qed text \ We then change variables for \z\: \ lemma beukers_aux_integral_transform2: assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "(LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - define g where "g = (\z. (1 - z) / (1-(1-x*y)*z))" define g' where "g' = (\z. -x*y / (1-(1-x*y)*z)^2)" have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - (x * y)) * w < 1" and neq: "(1 - (x * y)) * w \ 1" if w: "w \ {0..1}" for w proof - have "(1 - (x * y)) * w \ (1 - (x * y)) * 1" using w \x * y < 1\ by (intro mult_left_mono) auto also have "\ < 1" using xy by simp finally show "(1 - (x * y)) * w < 1" . thus "(1 - (x * y)) * w \ 1" by simp qed have deriv: "(g has_real_derivative g' w) (at w within {0..1})" if "w \ {0..1}" for w unfolding g_def g'_def apply (insert that xy neq) apply (rule derivative_eq_intros refl)+ apply (simp_all add: divide_simps power2_eq_square) apply (auto simp: algebra_simps) done have "continuous_on {0..1} (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" using neq by (auto intro!: continuous_intros) moreover have "g ` {0..1} \ {0..1}" proof clarify fix w :: real assume w: "w \ {0..1}" have "(1 - x * y) * w \ 1 * w" using \x * y < 1\ xy w by (intro mult_right_mono) auto thus "g w \ {0..1}" unfolding g_def using less[of w] w by (auto simp: divide_simps) qed ultimately have cont: "continuous_on (g ` {0..1}) (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" by (rule continuous_on_subset) have cont': "continuous_on {0..1} g'" using neq by (auto simp: g'_def intro!: continuous_intros) have "(LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = (1-y)^n * (LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) also have "(LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w)) = -(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w))" by (subst interval_integral_endpoints_reverse)(simp add: g_def zero_ereal_def one_ereal_def) also have "(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w)) = (LBINT w=0..1. g' w * ((1 - g w) ^ n / (1 - (1-x*y) * g w)))" using deriv cont cont' by (subst interval_integral_substitution_finite [symmetric, where g = g and g' = g']) (simp_all add: zero_ereal_def one_ereal_def) also have "-\ = (LBINT z=0..1. ((x*y)^n * z^n / (1-(1-x*y)*z)^(n+1)))" unfolding interval_lebesgue_integral_uminus [symmetric] using xy apply (intro interval_lebesgue_integral_lborel_01_cong) apply (simp add: divide_simps g_def g'_def) apply (auto simp: algebra_simps power_mult_distrib power2_eq_square) done also have "(1-y)^n * \ = (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) finally show "\ = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" .. qed text \ Lastly, we apply the same integration by parts to \x\: \ lemma beukers_aux_integral_transform3: assumes w: "w \ {0<..<1}" shows "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "1 - (1 - fst p * snd p) * w \ 0" by simp qed hence integrable: "set_integrable lborel D (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "(1 - (1 - fst p * snd p) * w) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT y=0..1. (LBINT x=0..1. P x * (1-y)^n / (1-(1-x*y)*w)))" using integrable unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. P x / (1-(1-y*x)*w)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" using w by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (LBINT x=0..1. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" using integrable' unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) finally show ?thesis . qed text \ We need to show the existence of some of these triple integrals. \ lemma beukers_aux_integrable1: "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x,y),z). P x * P y / (1-(1-x*y)*z))" proof - have D [measurable]: "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp flip: borel_prod) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def case_prod_unfold proof (subst lborel_prod [symmetric], intro lborel_pair.Fubini_set_integrable AE_I2 impI; clarsimp?) fix x y :: real assume xy: "x > 0" "x < 1" "y > 0" "y < 1" have "x * y < 1" using xy mult_strict_mono[of x 1 y 1] by simp show "set_integrable lborel {0<..<1} (\z. P x * P y / (1-(1-x*y)*z))" by (rule set_integrable_subset[of _ "{0..1}"], rule borel_integrable_atLeastAtMost') (use \x*y<1\ beukers_denom_neq[of x y] xy in \auto intro!: continuous_intros simp: P_def\) next have "set_integrable lborel D (\(x,y). (\z\{0<..<1}. norm (P x * P y / (1-(1-x*y)*z)) \lborel))" proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C\<^sup>2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y assume xy: "(x, y) \ D" have "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) = norm (LBINT z:{0<..<1}. \P x\ * \P y\ * (1 / (1-(1-x*y)*z)))" proof (intro arg_cong[where f = norm] set_lebesgue_integral_cong allI impI, goal_cases) case (2 z) with beukers_denom_ineq[of x y z] xy show ?case by (auto simp: abs_mult D_def) qed (auto simp: abs_mult D_def) also have "\ = norm (\P x\ * \P y\ * (LBINT z=0..1. (1 / (1-(1-x*y)*z))))" by (subst set_integral_mult_right) (auto simp: interval_integral_Ioo) also have "\ = norm (norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)))" using beukers_aux_ln_conv_integral[of x y] xy by (simp add: D_def) also have "\ = norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y))" using xy mult_strict_mono[of x 1 y 1] by (auto simp: D_def divide_nonpos_nonneg abs_mult) also have "norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)) \ norm (C * C * (- ln (x * y) / (1 - x * y)))" using xy C[of x] C[of y] mult_strict_mono[of x 1 y 1] unfolding norm_mult norm_divide by (intro mult_mono C) (auto simp: D_def divide_nonpos_nonneg) finally show "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (simp add: power2_eq_square mult_ac) next show "set_borel_measurable lborel D (\(x, y). LBINT z:{0<..<1}. norm (P x * P y / (1 - (1 - x * y) * z)))" unfolding lborel_prod [symmetric] set_borel_measurable_def case_prod_unfold set_lebesgue_integral_def P_def by measurable qed thus "set_integrable lborel ({0<..<1} \ {0<..<1}) (\x. LBINT y:{0<..<1}. \P (fst x) * P (snd x)\ / \1 - (1 - fst x * snd x) * y\)" by (simp add: case_prod_unfold D_def) qed (auto simp: case_prod_unfold lborel_prod [symmetric] set_borel_measurable_def P_def) qed lemma beukers_aux_integrable2: "set_integrable lborel D' (\(z,x,y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" have "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = norm (P x) * (1-y)^n * ((x*y*z) / (1-(1-x*y)*z))^n / (1-(1-x*y)*z)" using xyz beukers_denom_ineq[of x y z] by (simp add: abs_mult power_divide mult_ac) also have "(x*y*z) / (1-(1-x*y)*z) = 1/((1-z)/(z*x*y)+1)" using xyz by (simp add: field_simps) also have "norm (P x) * (1-y)^n * \^n / (1-(1-x*y)*z) \ C * 1^n * 1^n / (1-(1-x*y)*z)" using xyz C[of x] beukers_denom_ineq[of x y z] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*z)\" by linarith finally show "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) \ norm (case (z,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed lemma beukers_aux_integrable3: "set_integrable lborel D' (\(w,x,y). P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y w :: real assume xyw: "x \ {0<..<1}" "y \ {0<..<1}" "w \ {0<..<1}" have "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = norm (P x) * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)" using xyw beukers_denom_ineq[of x y w] by (simp add: abs_mult power_divide mult_ac) also have "\ \ C * 1^n * 1^n / (1-(1-x*y)*w)" using xyw C[of x] beukers_denom_ineq[of x y w] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*w)\" by linarith finally show "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) \ norm (case (w,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed text \ Now we only need to put all of these results together: \ lemma beukers_integral2_conv_3: "beukers_integral2 = beukers_integral3" proof - have cont_P: "continuous_on {0..1} P" by (auto simp: P_def intro: continuous_intros) have D [measurable]: "D \ sets borel" "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp_all flip: borel_prod) have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "beukers_integral2 = (LBINT (x,y):D. P x * P y * (LBINT z=0..1. 1 / (1-(1-x*y)*z)))" unfolding beukers_integral2_def case_prod_unfold by (intro set_lebesgue_integral_cong allI impI, measurable) (subst beukers_aux_ln_conv_integral, auto simp: D_def) also have "\ = (LBINT (x,y):D. (LBINT z=0..1. P x * P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) auto also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * P y / (1-(1-x*y)*z)))" by (simp add: interval_integral_Ioo) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)))" proof (subst lborel_pair.Fubini_set_integral [symmetric]) have "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x, y), z). P x * P y / (1 - (1 - x * y) * z))" using beukers_aux_integrable1 by simp also have "?this \ set_integrable (lborel \\<^sub>M lborel) ({0<..<1} \ D) (\(z,x,y). P x * P y / (1 - (1 - x * y) * z))" unfolding set_integrable_def by (subst lborel_pair.integrable_product_swap_iff [symmetric], intro integrable_cong) (auto simp: indicator_def case_prod_unfold lborel_prod D_def) finally show \ . qed (auto simp: case_prod_unfold) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (rule set_lebesgue_integral_cong) (use beukers_aux_integral_transform1 in simp_all) also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using beukers_aux_integrable2 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT (x,y):D. (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" proof (intro set_lebesgue_integral_cong allI impI; clarify?) fix x y :: real assume xy: "(x, y) \ D" have "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = P x * (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = P x * (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" using xy by (subst beukers_aux_integral_transform2) (auto simp: D_def) also have "\ = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) finally show "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" . qed (auto simp: D_def simp flip: borel_prod) also have "\ = (LBINT w:{0<..<1}. (LBINT (x,y):D. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" using beukers_aux_integrable3 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)))" unfolding case_prod_unfold by (subst set_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. (x*y*w*(1-x)*(1-y))^n / (1-(1-x*y)*w)^(n+1)))" by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_integral_transform3) (auto simp: mult_ac power_mult_distrib) also have "\ = (LBINT w=0..1. (LBINT (x,y):D. (x*y*w*(1-x)*(1-y)*(1-w))^n / (1-(1-x*y)*w)^(n+1)))" by (subst set_integral_mult_right [symmetric]) (auto simp: case_prod_unfold mult_ac power_mult_distrib) also have "\ = beukers_integral3" using beukers_integral3_integrable unfolding D'_def D_def beukers_integral3_def by (subst (2) lborel_prod [symmetric], subst lborel_pair.set_integral_fst') (auto simp: case_prod_unfold interval_integral_Ioo lborel_prod algebra_simps) finally show ?thesis . qed subsection \The main result\ text \ Combining all of the results so far, we can derive the key inequalities \[0 < A\zeta(3) + B < 2 \zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\] for integers $A$, $B$ with $A > 0$. \ lemma zeta_3_linear_combination_bounds: obtains A B :: int where "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" proof - define I where "I = beukers_integral2" define d where "d = Lcm {1..n} ^ 3" have "d > 0" by (auto simp: d_def intro!: Nat.gr0I) from beukers_integral2_conv_int_combination obtain A' B :: int where *: "A' > 0" "I = A' * Re (zeta 3) + B / d" unfolding I_def d_def . define A where "A = A' * d" from * have A: "A > 0" "I = (A * Re (zeta 3) + B) / d" using \d > 0\ by (simp_all add: A_def field_simps) have "0 < I" using beukers_integral3_pos by (simp add: I_def beukers_integral2_conv_3) with \d > 0\ have "A * Re (zeta 3) + B > 0" by (simp add: field_simps A) moreover have "I \ 2 * (1 / 27) ^ n * Re (zeta 3)" using beukers_integral2_conv_3 beukers_integral3_le by (simp add: I_def) hence "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * d / 27 ^ n" using \d > 0\ by (simp add: A field_simps) ultimately show ?thesis using A by (intro that[of A B]) (auto simp: d_def) qed text \ If $\zeta(3) = \frac{a}{b}$ for some integers $a$ and $b$, we can thus derive the inequality $2b\zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\geq 1$ for any natural number $n$. \ lemma beukers_key_inequality: fixes a :: int and b :: nat assumes "b > 0" and ab: "Re (zeta 3) = a / b" shows "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" proof - from zeta_3_linear_combination_bounds obtain A B :: int where AB: "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" . from AB have "0 < (A * Re (zeta 3) + B) * b" using \b > 0\ by (intro mult_pos_pos) auto also have "\ = A * (Re (zeta 3) * b) + B * b" using \b > 0\ by (simp add: algebra_simps) also have "Re (zeta 3) * b = a" using \b > 0\ by (simp add: ab) also have "of_int A * of_int a + of_int (B * b) = of_int (A * a + B * b)" by simp finally have "1 \ A * a + B * b" by linarith hence "1 \ real_of_int (A * a + B * b)" by linarith also have "\ = (A * (a / b) + B) * b" using \b > 0\ by (simp add: ring_distribs) also have "a / b = Re (zeta 3)" by (simp add: ab) also have "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n" using AB by simp finally show "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" using \b > 0\ by (simp add: mult_ac) qed end (* TODO: Move *) lemma smallo_power: "n > 0 \ f \ o[F](g) \ (\x. f x ^ n) \ o[F](\x. g x ^ n)" by (induction n rule: nat_induct_non_zero) (auto intro: landau_o.small.mult) text \ This is now a contradiction, since $\text{lcm}\{1\ldots n\} \in o(3^n)$ by the Prime Number Theorem -- hence the main result. \ theorem zeta_3_irrational: "zeta 3 \ \" proof assume "zeta 3 \ \" obtain a :: int and b :: nat where "b > 0" and ab': "zeta 3 = a / b" proof - from \zeta 3 \ \\ obtain r where r: "zeta 3 = of_rat r" by (elim Rats_cases) also have "r = rat_of_int (fst (quotient_of r)) / rat_of_int (snd (quotient_of r))" by (intro quotient_of_div) auto also have "of_rat \ = (of_int (fst (quotient_of r)) / of_int (snd (quotient_of r)) :: complex)" by (simp add: of_rat_divide) also have "of_int (snd (quotient_of r)) = of_nat (nat (snd (quotient_of r)))" using quotient_of_denom_pos'[of r] by auto finally have "zeta 3 = of_int (fst (quotient_of r)) / of_nat (nat (snd (quotient_of r)))" . thus ?thesis using quotient_of_denom_pos'[of r] by (intro that[of "nat (snd (quotient_of r))" "fst (quotient_of r)"]) auto qed hence ab: "Re (zeta 3) = a / b" by simp interpret prime_number_theorem by standard (rule prime_number_theorem) have Lcm_upto_smallo: "(\n. real (Lcm {1..n})) \ o(\n. c ^ n)" if c: "c > exp 1" for c proof - have "0 < exp (1::real)" by simp also note c finally have "c > 0" . have "(\n. real (Lcm {1..n})) = (\n. real (Lcm {1..nat \real n\}))" by simp also have "\ \ o(\n. c powr real n)" using Lcm_upto.smallo' by (rule landau_o.small.compose) (simp_all add: c filterlim_real_sequentially) also have "(\n. c powr real n) = (\n. c ^ n)" using c \c > 0\ by (subst powr_realpow) auto finally show ?thesis . qed have "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ O(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n)" using \b > 0\ Re_zeta_ge_1[of 3] by simp also have "exp 1 < (3 :: real)" using e_approx_32 by (simp add: abs_if split: if_splits) hence "(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\n. (3 ^ n) ^ 3 / 27 ^ n)" unfolding of_nat_power by (intro landau_o.small.divide_right smallo_power Lcm_upto_smallo) auto also have "(\n. (3 ^ n) ^ 3 / 27 ^ n :: real) = (\_. 1)" by (simp add: power_mult [of 3, symmetric] mult.commute[of _ 3] power_mult[of _ 3]) finally have *: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\_. 1)" . have lim: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ 0" using smalloD_tendsto[OF *] by simp moreover have "1 \ real (2 * b) * Re (zeta 3) * real (Lcm {1..n} ^ 3) / 27 ^ n" for n using beukers_key_inequality[of b a] ab \b > 0\ by simp ultimately have "1 \ (0 :: real)" by (intro tendsto_lowerbound[OF lim] always_eventually allI) auto thus False by simp qed end \ No newline at end of file diff --git a/thys/Zeta_Function/Zeta_Library.thy b/thys/Zeta_Function/Zeta_Library.thy --- a/thys/Zeta_Function/Zeta_Library.thy +++ b/thys/Zeta_Function/Zeta_Library.thy @@ -1,897 +1,898 @@ section \Various preliminary material\ theory Zeta_Library imports "HOL-Complex_Analysis.Complex_Analysis" "HOL-Real_Asymp.Real_Asymp" "Dirichlet_Series.Dirichlet_Series_Analysis" begin subsection \Facts about limits\ lemma at_within_altdef: "at x within A = (INF S\{S. open S \ x \ S}. principal (S \ (A - {x})))" unfolding at_within_def nhds_def inf_principal [symmetric] by (subst INF_inf_distrib [symmetric]) (auto simp: INF_constant) lemma tendsto_at_left_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \dx\{d<.. A m" proof (rule ccontr) assume "\ (\m. \dx\{d<.. A m)" then obtain m where **: "\d. d < c \ \x\{d<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n < c) \ X (Suc n) > c - max 0 ((c - X n) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c - 1"] show ?case by auto next case (2 x n) with **[of "c - max 0 (c - x) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n < c" "\n. X (Suc n) > c - max 0 ((c - X n) / 2)" by auto have X_ge: "X n \ c - (c - X 0) / 2 ^ n" for n proof (induction n) case (Suc n) have "c - (c - X 0) / 2 ^ Suc n = c - (c - (c - (c - X 0) / 2 ^ n)) / 2" by simp also have "c - (c - (c - (c - X 0) / 2 ^ n)) / 2 \ c - (c - X n) / 2" by (intro diff_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c - max 0 ((c - X n) / 2)" using X[of n] by (simp add: max_def) also have "\ < X (Suc n)" using X[of n] by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c - (c - X 0) / 2 ^ n) sequentially" using X_ge by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_left c) sequentially" by (rule tendsto_imp_filterlim_at_left) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m < c" "x \ {d m<.. f x \ A m" for m x by metis have ***: "at_left c = (INF S\{S. open S \ c \ S}. principal (S \ {..m. {d m<..}"]) auto qed lemma shows at_right_PInf [simp]: "at_right (\ :: ereal) = bot" and at_left_MInf [simp]: "at_left (-\ :: ereal) = bot" proof - have "{(\::ereal)<..} = {}" "{..<-(\::ereal)} = {}" by auto thus "at_right (\ :: ereal) = bot" "at_left (-\ :: ereal) = bot" by (simp_all add: at_within_def) qed lemma tendsto_at_left_erealI_sequentially: fixes f :: "ereal \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_left c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_left c)" proof (cases c) case [simp]: PInf have "((\x. f (ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_left_PInf filterlim_filtermap) next case [simp]: MInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_left c')" proof (intro tendsto_at_left_realI_sequentially assms) fix X assume *: "filterlim X (at_left c') sequentially" show "filterlim (\n. ereal (X n)) (at_left c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_left) qed thus ?thesis by (simp add: at_left_ereal filterlim_filtermap) qed lemma tendsto_at_right_realI_sequentially: fixes f :: "real \ 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof - obtain A where A: "decseq A" "open (A n)" "y \ A n" "nhds y = (INF n. principal (A n))" for n by (rule nhds_countable[of y]) (rule that) have "\m. \d>c. \x\{c<.. A m" proof (rule ccontr) assume "\ (\m. \d>c. \x\{c<.. A m)" then obtain m where **: "\d. d > c \ \x\{c<.. A m" by auto have "\X. \n. (f (X n) \ A m \ X n > c) \ X (Suc n) < c + max 0 ((X n - c) / 2)" proof (intro dependent_nat_choice, goal_cases) case 1 from **[of "c + 1"] show ?case by auto next case (2 x n) with **[of "c + max 0 (x - c) / 2"] show ?case by force qed then obtain X where X: "\n. f (X n) \ A m" "\n. X n > c" "\n. X (Suc n) < c + max 0 ((X n - c) / 2)" by auto have X_le: "X n \ c + (X 0 - c) / 2 ^ n" for n proof (induction n) case (Suc n) have "X (Suc n) < c + max 0 ((X n - c) / 2)" by (intro X) also have "\ = c + (X n - c) / 2" using X[of n] by (simp add: field_simps max_def) also have "\ \ c + (c + (X 0 - c) / 2 ^ n - c) / 2" by (intro add_left_mono divide_right_mono Suc diff_right_mono) auto also have "\ = c + (X 0 - c) / 2 ^ Suc n" by simp finally show ?case by linarith qed auto have "X \ c" proof (rule tendsto_sandwich) show "eventually (\n. X n \ c) sequentially" using X by (intro always_eventually) (auto intro!: less_imp_le) show "eventually (\n. X n \ c + (X 0 - c) / 2 ^ n) sequentially" using X_le by (intro always_eventually) auto qed real_asymp+ hence "filterlim X (at_right c) sequentially" by (rule tendsto_imp_filterlim_at_right) (use X in \auto intro!: always_eventually less_imp_le\) from topological_tendstoD[OF *[OF this] A(2, 3), of m] X(1) show False by auto qed then obtain d where d: "d m > c" "x \ {c<.. f x \ A m" for m x by metis have ***: "at_right c = (INF S\{S. open S \ c \ S}. principal (S \ {c<..}))" by (simp add: at_within_altdef) from d show ?thesis unfolding *** A using A(1,2) by (intro filterlim_base[of _ "\m. {.. 'b::first_countable_topology" assumes *: "\X. filterlim X (at_right c) sequentially \ (\n. f (X n)) \ y" shows "(f \ y) (at_right c)" proof (cases c) case [simp]: MInf have "((\x. f (-ereal x)) \ y) at_top" using assms by (intro tendsto_at_topI_sequentially assms) (simp_all flip: uminus_ereal.simps ereal_tendsto_simps add: o_def filterlim_at) thus ?thesis by (simp add: at_right_MInf filterlim_filtermap at_top_mirror) next case [simp]: PInf thus ?thesis by auto next case [simp]: (real c') have "((\x. f (ereal x)) \ y) (at_right c')" proof (intro tendsto_at_right_realI_sequentially assms) fix X assume *: "filterlim X (at_right c') sequentially" show "filterlim (\n. ereal (X n)) (at_right c) sequentially" by (rule filterlim_compose[OF _ *]) (simp add: sequentially_imp_eventually_within tendsto_imp_filterlim_at_right) qed thus ?thesis by (simp add: at_right_ereal filterlim_filtermap) qed proposition analytic_continuation': assumes hol: "f holomorphic_on S" "g holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = g z" and "w \ S" shows "f w = g w" using analytic_continuation[OF holomorphic_on_diff[OF hol] assms(3-7) _ assms(9)] assms(8) by simp subsection \Various facts about integrals\ lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ 'b :: euclidean_space" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" assumes "set_integrable M A f" "set_borel_measurable M A g" assumes "AE x in M. x \ A \ norm (g x) \ norm (f x)" shows "set_integrable M A g" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound) from assms(1) show "integrable M (\x. indicator A x *\<^sub>R f x)" by (simp add: set_integrable_def) from assms(2) show "(\x. indicat_real A x *\<^sub>R g x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(3) show "AE x in M. norm (indicat_real A x *\<^sub>R g x) \ norm (indicat_real A x *\<^sub>R f x)" by eventually_elim (simp add: indicator_def) qed (* TODO replace version in library. Better name? *) lemma nn_integral_has_integral_lebesgue: fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. indicator \ x * f x) = I" proof - from I have "(\x. indicator \ x *\<^sub>R f x) \ lebesgue \\<^sub>M borel" by (rule has_integral_implies_lebesgue_measurable) then obtain f' :: "'a \ real" where [measurable]: "f' \ borel \\<^sub>M borel" and eq: "AE x in lborel. indicator \ x * f x = f' x" by (auto dest: completion_ex_borel_measurable_real) from I have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV" - using nonneg by (simp add: indicator_def if_distrib[of "\x. x * f y" for y] cong: if_cong) + using nonneg by (simp add: indicator_def of_bool_def if_distrib[of "\x. x * f y" for y] cong: if_cong) also have "((\x. abs (indicator \ x * f x)) has_integral I) UNIV \ ((\x. abs (f' x)) has_integral I) UNIV" using eq by (intro has_integral_AE) auto finally have "integral\<^sup>N lborel (\x. abs (f' x)) = I" by (rule nn_integral_has_integral_lborel[rotated 2]) auto also have "integral\<^sup>N lborel (\x. abs (f' x)) = integral\<^sup>N lborel (\x. abs (indicator \ x * f x))" using eq by (intro nn_integral_cong_AE) auto also have "(\x. abs (indicator \ x * f x)) = (\x. indicator \ x * f x)" using nonneg by (auto simp: indicator_def fun_eq_iff) finally show ?thesis . qed subsection \Uniform convergence of integrals\ lemma has_absolute_integral_change_of_variables_1': fixes f :: "real \ real" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "(\x. \g' x\ *\<^sub>R vec (f(g x)) :: real ^ 1) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R vec (f(g x))) = (vec b :: real ^ 1) \ (\x. vec (f x) :: real ^ 1) absolutely_integrable_on (g ` S) \ integral (g ` S) (\x. vec (f x)) = (vec b :: real ^ 1)" using assms unfolding has_field_derivative_iff_has_vector_derivative by (intro has_absolute_integral_change_of_variables_1 assms) auto thus ?thesis by (simp add: absolutely_integrable_on_1_iff integral_on_1_eq) qed lemma set_nn_integral_lborel_eq_integral: fixes f::"'a::euclidean_space \ real" assumes "set_borel_measurable borel A f" assumes "\x. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \lborel) < \" shows "(\\<^sup>+x\A. f x \lborel) = integral A f" proof - have eq: "(\\<^sup>+x\A. f x \lborel) = (\\<^sup>+x. ennreal (indicator A x * f x) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = integral UNIV (\x. indicator A x * f x)" using assms eq by (intro nn_integral_lborel_eq_integral) (auto simp: indicator_def set_borel_measurable_def) also have "integral UNIV (\x. indicator A x * f x) = integral A (\x. indicator A x * f x)" - by (rule integral_spike_set) (auto simp: indicator_def) + by (rule integral_spike_set) (auto intro: empty_imp_negligible) + also have "\ = integral A f" by (rule integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma nn_integral_has_integral_lebesgue': fixes f :: "'a::euclidean_space \ real" assumes nonneg: "\x. x \ \ \ 0 \ f x" and I: "(f has_integral I) \" shows "integral\<^sup>N lborel (\x. ennreal (f x) * indicator \ x) = ennreal I" proof - have "integral\<^sup>N lborel (\x. ennreal (f x) * indicator \ x) = integral\<^sup>N lborel (\x. ennreal (indicator \ x * f x))" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal I" using assms by (intro nn_integral_has_integral_lebesgue) finally show ?thesis . qed lemma uniform_limit_set_lebesgue_integral: fixes f :: "'a \ 'b :: euclidean_space \ 'c :: {banach, second_countable_topology}" assumes "set_integrable lborel X' g" assumes [measurable]: "X' \ sets borel" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel X' (f y)" assumes "\y. y \ Y \ (AE t\X' in lborel. norm (f y t) \ g t)" assumes "eventually (\x. X x \ sets borel \ X x \ X') F" assumes "filterlim (\x. set_lebesgue_integral lborel (X x) g) (nhds (set_lebesgue_integral lborel X' g)) F" shows "uniform_limit Y (\x y. set_lebesgue_integral lborel (X x) (f y)) (\y. set_lebesgue_integral lborel X' (f y)) F" proof (rule uniform_limitI, goal_cases) case (1 \) have integrable_g: "set_integrable lborel U g" if "U \ sets borel" "U \ X'" for U by (rule set_integrable_subset[OF assms(1)]) (use that in auto) have "eventually (\x. dist (set_lebesgue_integral lborel (X x) g) (set_lebesgue_integral lborel X' g) < \) F" using \\ > 0\ assms by (auto simp: tendsto_iff) from this show ?case using \eventually (\_. _ \ _) F\ proof eventually_elim case (elim x) hence [measurable]:"X x \ sets borel" and "X x \ X'" by auto have integrable: "set_integrable lborel U (f y)" if "y \ Y" "U \ sets borel" "U \ X'" for y U apply (rule set_integrable_subset) apply (rule set_integrable_bound[OF assms(1)]) apply (use assms(3) that in \simp add: set_borel_measurable_def\) using assms(4)[OF \y \ Y\] apply eventually_elim apply force using that apply simp_all done show ?case proof fix y assume "y \ Y" have "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) = norm (set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y))" by (simp add: dist_norm norm_minus_commute) also have "set_lebesgue_integral lborel X' (f y) - set_lebesgue_integral lborel (X x) (f y) = set_lebesgue_integral lborel (X' - X x) (f y)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable; (fact | simp)) apply (rule integrable; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "norm \ \ (\t\X'-X x. norm (f y t) \lborel)" by (intro set_integral_norm_bound integrable) (fact | simp)+ also have "AE t\X' - X x in lborel. norm (f y t) \ g t" using assms(4)[OF \y \ Y\] by eventually_elim auto with \y \ Y\ have "(\t\X'-X x. norm (f y t) \lborel) \ (\t\X'-X x. g t \lborel)" by (intro set_integral_mono_AE set_integrable_norm integrable integrable_g) auto also have "\ = (\t\X'. g t \lborel) - (\t\X x. g t \lborel)" unfolding set_lebesgue_integral_def apply (subst Bochner_Integration.integral_diff [symmetric]) unfolding set_integrable_def [symmetric] apply (rule integrable_g; (fact | simp)) apply (rule integrable_g; fact) apply (intro Bochner_Integration.integral_cong) apply (use \X x \ X'\ in \auto simp: indicator_def\) done also have "\ \ dist (\t\X x. g t \lborel) (\t\X'. g t \lborel)" by (simp add: dist_norm) also have "\ < \" by fact finally show "dist (set_lebesgue_integral lborel (X x) (f y)) (set_lebesgue_integral lborel X' (f y)) < \" . qed qed qed lemma integral_dominated_convergence_at_right: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_right c)" assumes bound: "\\<^sub>F i in at_right c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_right c)" proof (rule tendsto_at_right_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_right c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_right c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma integral_dominated_convergence_at_left: fixes s :: "real \ 'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" and f :: "'a \ 'b" and M and c :: real assumes "f \ borel_measurable M" "\t. s t \ borel_measurable M" "integrable M w" assumes lim: "AE x in M. ((\i. s i x) \ f x) (at_left c)" assumes bound: "\\<^sub>F i in at_left c. AE x in M. norm (s i x) \ w x" shows "((\t. integral\<^sup>L M (s t)) \ integral\<^sup>L M f) (at_left c)" proof (rule tendsto_at_left_realI_sequentially) fix X :: "nat \ real" assume X: "filterlim X (at_left c) sequentially" from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound] obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) \ f x) (at_left c)" then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ qed lemma uniform_limit_interval_integral_right: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\b' y. LBINT t=a..b'. f y t) (\y. LBINT t=a..b. f y t) (at_left b)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_leftI) with \a < b\ have "?thesis \ uniform_limit Y (\b' y. \t\einterval a (min b b'). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_left b)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F b' in at_left b. einterval a (min b b') \ sets borel \ einterval a (min b b') \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\b'. set_lebesgue_integral lborel (einterval a (min b b')) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_left b)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_left_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval a (min b (X n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_left b) sequentially" show "AE x in lborel. (\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval a (min b (X t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval a (min b (X t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {max a x<..t. X t \ {max a x<..t. indicator (einterval a (min b (X t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval a (min b (X n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval a (min b (X n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_left: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes "a < b" shows "uniform_limit Y (\a' y. LBINT t=a'..b. f y t) (\y. LBINT t=a..b. f y t) (at_right a)" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\b'. b' \ {a<..a < b\ by (intro eventually_at_rightI) with \a < b\ have "?thesis \ uniform_limit Y (\a' y. \t\einterval (max a a') b. f y t \lborel) (\y. \t\einterval a b. f y t \lborel) (at_right a)" by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F a' in at_right a. einterval (max a a') b \ sets borel \ einterval (max a a') b \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\a'. set_lebesgue_integral lborel (einterval (max a a') b) g) \ set_lebesgue_integral lborel (einterval a b) g) (at_right a)" unfolding set_lebesgue_integral_def proof (intro tendsto_at_right_erealI_sequentially integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix X :: "nat \ ereal" and n :: nat have "set_borel_measurable borel (einterval (max a (X n)) b) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next fix X :: "nat \ ereal" assume X: "filterlim X (at_right a) sequentially" show "AE x in lborel. (\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (X t)) b) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (X t)) b"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..t. X t \ {a<..t. indicator (einterval (max a (X t)) b) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (X n)) b) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma uniform_limit_interval_integral_sequentially: fixes f :: "'a \ real \ 'c :: {banach, second_countable_topology}" assumes "interval_lebesgue_integrable lborel a b g" assumes [measurable]: "\y. y \ Y \ set_borel_measurable borel (einterval a b) (f y)" assumes "\y. y \ Y \ (AE t\einterval a b in lborel. norm (f y t) \ g t)" assumes a': "filterlim a' (at_right a) sequentially" assumes b': "filterlim b' (at_left b) sequentially" assumes "a < b" shows "uniform_limit Y (\n y. LBINT t=a' n..b' n. f y t) (\y. LBINT t=a..b. f y t) sequentially" proof (cases "Y = {}") case False have g_nonneg: "AE t\einterval a b in lborel. g t \ 0" proof - from \Y \ {}\ obtain y where "y \ Y" by auto from assms(3)[OF this] show ?thesis by eventually_elim (auto elim: order.trans[rotated]) qed have ev: "eventually (\n. a < a' n \ a' n < b' n \ b' n < b) sequentially" proof - from ereal_dense2[OF \a < b\] obtain t where t: "a < ereal t" "ereal t < b" by blast from t have "eventually (\n. a' n \ {a<..n. b' n \ {t<..n. a < a' n \ a' n < b' n \ b' n < b) sequentially" by eventually_elim auto qed have "?thesis \ uniform_limit Y (\n y. \t\einterval (max a (a' n)) (min b (b' n)). f y t \lborel) (\y. \t\einterval a b. f y t \lborel) sequentially" using \a < b\ by (intro filterlim_cong arg_cong2[where f = uniformly_on]) (auto simp: interval_lebesgue_integral_def fun_eq_iff min_def max_def intro!: eventually_mono[OF ev]) also have \ proof (rule uniform_limit_set_lebesgue_integral[where g = g], goal_cases) show "\\<^sub>F n in sequentially. einterval (max a (a' n)) (min b (b' n)) \ sets borel \ einterval (max a (a' n)) (min b (b' n)) \ einterval a b" using ev by eventually_elim (auto simp: einterval_def) next show "((\n. set_lebesgue_integral lborel (einterval (max a (a' n)) (min b (b' n))) g) \ set_lebesgue_integral lborel (einterval a b) g) sequentially" unfolding set_lebesgue_integral_def proof (intro integral_dominated_convergence) have *: "set_borel_measurable borel (einterval a b) g" using assms(1) less_imp_le[OF \a < b\] by (simp add: interval_lebesgue_integrable_def set_integrable_def set_borel_measurable_def) show "(\x. indicat_real (einterval a b) x *\<^sub>R g x) \ borel_measurable lborel" using * by (simp add: set_borel_measurable_def) fix n :: nat have "set_borel_measurable borel (einterval (max a (a' n)) (min b (b' n))) g" by (rule set_borel_measurable_subset[OF *]) (auto simp: einterval_def) thus "(\x. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ borel_measurable lborel" by (simp add: set_borel_measurable_def) next show "AE x in lborel. (\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" proof (rule AE_I2) fix x :: real have "(\t. indicator (einterval (max a (a' t)) (min b (b' t))) x :: real) \ indicator (einterval a b) x" proof (cases "x \ einterval a b") case False hence "x \ einterval (max a (a' t)) (min b (b' t))"for t by (auto simp: einterval_def) with False show ?thesis by (simp add: indicator_def) next case True with \a < b\ have "eventually (\t. t \ {a<..n. x \ {a' n<..n. a' n \ {a<..n. b' n \ {x<..n. x \ {a' n<..t. indicator (einterval (max a (a' t)) (min b (b' t))) x = (1 :: real)) sequentially" by eventually_elim (use True in \auto simp: indicator_def einterval_def\) from tendsto_eventually[OF this] and True show ?thesis by (simp add: indicator_def) qed thus "(\n. indicat_real (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicat_real (einterval a b) x *\<^sub>R g x" by (intro tendsto_intros) qed next fix X :: "nat \ ereal" and n :: nat show "AE x in lborel. norm (indicator (einterval (max a (a' n)) (min b (b' n))) x *\<^sub>R g x) \ indicator (einterval a b) x *\<^sub>R g x" using g_nonneg by eventually_elim (auto simp: indicator_def einterval_def) qed (use assms less_imp_le[OF \a < b\] in \auto simp: interval_lebesgue_integrable_def set_integrable_def\) qed (use assms in \auto simp: interval_lebesgue_integrable_def\) finally show ?thesis . qed auto lemma interval_lebesgue_integrable_combine: assumes "interval_lebesgue_integrable lborel A B f" assumes "interval_lebesgue_integrable lborel B C f" assumes "set_borel_measurable borel (einterval A C) f" assumes "A \ B" "B \ C" shows "interval_lebesgue_integrable lborel A C f" proof - have meas: "set_borel_measurable borel (einterval A B \ einterval B C) f" by (rule set_borel_measurable_subset[OF assms(3)]) (use assms in \auto simp: einterval_def\) have "set_integrable lborel (einterval A B \ einterval B C) f" using assms by (intro set_integrable_Un) (auto simp: interval_lebesgue_integrable_def) also have "?this \ set_integrable lborel (einterval A C) f" proof (cases "B \ {\, -\}") case True with assms have "einterval A B \ einterval B C = einterval A C" by (auto simp: einterval_def) thus ?thesis by simp next case False then obtain B' where [simp]: "B = ereal B'" by (cases B) auto have "indicator (einterval A C) x = (indicator (einterval A B \ einterval B C) x :: real)" if "x \ B'" for x using assms(4,5) that by (cases A; cases C) (auto simp: einterval_def indicator_def) hence "{x \ space lborel. indicat_real (einterval A B \ einterval B C) x *\<^sub>R f x \ indicat_real (einterval A C) x *\<^sub>R f x} \ {B'}" by force thus ?thesis unfolding set_integrable_def using meas assms by (intro integrable_cong_AE AE_I[of _ _ "{B'}"]) (simp_all add: set_borel_measurable_def) qed also have "\ \ ?thesis" using order.trans[OF assms(4,5)] by (simp add: interval_lebesgue_integrable_def) finally show ?thesis . qed lemma interval_lebesgue_integrable_bigo_right: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_left B](g)" assumes cont: "continuous_on {A.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_left B)" by (elim landau_o.bigE) then obtain B' where B': "B' < B" "\x. x \ {B'<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_left[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel A (max A B') f" using B' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {max A B'<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {max A B'<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {max A B'<.. c * norm (g x)" by (intro B') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal (max A B')) (ereal B) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \B' < B\ assms by simp qed (use B' assms in auto) qed lemma interval_lebesgue_integrable_bigo_left: fixes A B :: real fixes f :: "real \ real" assumes "f \ O[at_right A](g)" assumes cont: "continuous_on {A<..B} f" assumes meas: "set_borel_measurable borel {A<.. 0" "eventually (\x. norm (f x) \ c * norm (g x)) (at_right A)" by (elim landau_o.bigE) then obtain A' where A': "A' > A" "\x. x \ {A<.. norm (f x) \ c * norm (g x)" using \A < B\ by (auto simp: Topological_Spaces.eventually_at_right[of A]) show ?thesis proof (rule interval_lebesgue_integrable_combine) show "interval_lebesgue_integrable lborel (min B A') B f" using A' assms by (intro interval_integrable_continuous_on continuous_on_subset[OF cont]) auto show "set_borel_measurable borel (einterval (ereal A) (ereal B)) f" using assms by simp have meas': "set_borel_measurable borel {A<..x. c * g x)" using assms by (simp add: interval_lebesgue_integrable_def) thus "set_integrable lborel {A<..x. c * g x)" by (rule set_integrable_subset) auto next fix x assume "x \ {A<.. c * norm (g x)" by (intro A') auto also have "\ \ norm (c * g x)" unfolding norm_mult by (intro mult_right_mono) auto finally show "norm (f x) \ norm (c * g x)" . qed (use meas' in \simp_all add: set_borel_measurable_def\) thus "interval_lebesgue_integrable lborel (ereal A) (ereal (min B A')) f" unfolding interval_lebesgue_integrable_def einterval_eq_Icc using \A' > A\ assms by simp qed (use A' assms in auto) qed subsection \Other material\ (* TODO: Library *) lemma summable_comparison_test_bigo: fixes f :: "nat \ real" assumes "summable (\n. norm (g n))" "f \ O(g)" shows "summable f" proof - from \f \ O(g)\ obtain C where C: "eventually (\x. norm (f x) \ C * norm (g x)) at_top" by (auto elim: landau_o.bigE) thus ?thesis by (rule summable_comparison_test_ev) (insert assms, auto intro: summable_mult) qed lemma fps_expansion_cong: assumes "eventually (\x. g x = h x) (nhds x)" shows "fps_expansion g x = fps_expansion h x" proof - have "(deriv ^^ n) g x = (deriv ^^ n) h x" for n by (intro higher_deriv_cong_ev assms refl) thus ?thesis by (simp add: fps_expansion_def) qed lemma fps_expansion_eq_zero_iff: assumes "g holomorphic_on ball z r" "r > 0" shows "fps_expansion g z = 0 \ (\z\ball z r. g z = 0)" proof assume *: "\z\ball z r. g z = 0" have "eventually (\w. w \ ball z r) (nhds z)" using assms by (intro eventually_nhds_in_open) auto hence "eventually (\z. g z = 0) (nhds z)" by eventually_elim (use * in auto) hence "fps_expansion g z = fps_expansion (\_. 0) z" by (intro fps_expansion_cong) thus "fps_expansion g z = 0" by (simp add: fps_expansion_def fps_zero_def) next assume *: "fps_expansion g z = 0" have "g w = 0" if "w \ ball z r" for w by (rule holomorphic_fun_eq_0_on_ball[OF assms(1) that]) (use * in \auto simp: fps_expansion_def fps_eq_iff\) thus "\w\ball z r. g w = 0" by blast qed lemma fds_nth_higher_deriv: "fds_nth ((fds_deriv ^^ k) F) = (\n. (-1) ^ k * of_real (ln n) ^ k * fds_nth F n)" by (induction k) (auto simp: fds_nth_deriv fun_eq_iff simp flip: scaleR_conv_of_real) lemma binomial_n_n_minus_one [simp]: "n > 0 \ n choose (n - Suc 0) = n" by (cases n) auto lemma has_field_derivative_complex_powr_right: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z within A)" by (rule DERIV_subset, rule has_field_derivative_powr_right) auto lemmas has_field_derivative_complex_powr_right' = has_field_derivative_complex_powr_right[THEN DERIV_chain2] end \ No newline at end of file