diff --git a/thys/DiscretePricing/Disc_Cond_Expect.thy b/thys/DiscretePricing/Disc_Cond_Expect.thy --- a/thys/DiscretePricing/Disc_Cond_Expect.thy +++ b/thys/DiscretePricing/Disc_Cond_Expect.thy @@ -1,1336 +1,1336 @@ (* Title: Disc_Cond_Expect.thy Author: Mnacho Echenim, Univ. Grenoble Alpes *) section \Discrete Conditional Expectation\ theory Disc_Cond_Expect imports "HOL-Probability.Probability" Generated_Subalgebra begin subsection \Preliminary measurability results\ text \These are some useful results, in particular when working with functions that have a countable codomain.\ definition disc_fct where "disc_fct f \ countable (range f)" definition point_measurable where "point_measurable M S f \ (f`(space M)\ S) \ (\ r \ (range f) \ S . f-`{r} \ (space M) \ sets M)" lemma singl_meas_if: assumes "f \ space M \ space N" and "\r\ range f\ space N. \A\ sets N. range f\ A = {r}" shows "point_measurable (fct_gen_subalgebra M N f) (space N) f" unfolding point_measurable_def proof show "f`space (fct_gen_subalgebra M N f)\ space N" using assms by (simp add: Pi_iff fct_gen_subalgebra_space image_subsetI) show "(\r\range f \ space N. f -` {r} \ space (fct_gen_subalgebra M N f) \ sets (fct_gen_subalgebra M N f))" proof fix r assume "r\ range f \ space N" hence "\A\ sets N. range f\ A = {r}" using assms by blast from this obtain A where "A\ sets N" and "range f \ A = {r}" by auto note Aprops = this hence "f-`A = f-`{r}" by auto hence "f-`A \ space M = f-`{r} \ space (fct_gen_subalgebra M N f)" by (simp add: fct_gen_subalgebra_space) thus "f -` {r} \ space (fct_gen_subalgebra M N f) \ sets (fct_gen_subalgebra M N f)" using Aprops fct_gen_subalgebra_sets_mem[of A N f M] by simp qed qed lemma meas_single_meas: assumes "f\ measurable M N" and "\r\ range f\ space N. \A\ sets N. range f\ A = {r}" shows "point_measurable M (space N) f" proof - have "subalgebra M (fct_gen_subalgebra M N f) " using assms fct_gen_subalgebra_is_subalgebra by blast hence "sets (fct_gen_subalgebra M N f) \ sets M" by (simp add: subalgebra_def) moreover have "point_measurable (fct_gen_subalgebra M N f) (space N) f" using assms singl_meas_if by (metis (no_types, lifting) Pi_iff measurable_space) ultimately show ?thesis proof - obtain bb :: "'a measure \ 'b set \ ('a \ 'b) \ 'b" where f1: "\m B f. (\ point_measurable m B f \ f ` space m \ B \ (\b. b \ range f \ B \ f -` {b} \ space m \ sets m)) \ (\ f ` space m \ B \ bb m B f \ range f \ B \ f -` {bb m B f} \ space m \ sets m \ point_measurable m B f)" by (metis (no_types) point_measurable_def) moreover { assume "f -` {bb M (space N) f} \ space (fct_gen_subalgebra M N f) \ sets (fct_gen_subalgebra M N f)" then have "f -` {bb M (space N) f} \ space M \ sets (fct_gen_subalgebra M N f)" by (metis \subalgebra M (fct_gen_subalgebra M N f)\ subalgebra_def) then have "f -` {bb M (space N) f} \ space M \ sets M" using \sets (fct_gen_subalgebra M N f) \ sets M\ by blast then have "f ` space M \ space N \ f -` {bb M (space N) f} \ space M \ sets M" using f1 by (metis \point_measurable (fct_gen_subalgebra M N f) (space N) f\ \subalgebra M (fct_gen_subalgebra M N f)\ subalgebra_def) then have ?thesis using f1 by metis } ultimately show ?thesis by (metis (no_types) \point_measurable (fct_gen_subalgebra M N f) (space N) f\ \subalgebra M (fct_gen_subalgebra M N f)\ subalgebra_def) qed qed definition countable_preimages where "countable_preimages B Y = (\n. if ((infinite B) \ (finite B \ n < card B)) then Y -` {(from_nat_into B) n} else {})" lemma count_pre_disj: fixes i::nat assumes "countable B" and "i \ j" shows "(countable_preimages B Y) i \ (countable_preimages B Y) j = {}" proof (cases "(countable_preimages B Y) i = {} \ (countable_preimages B Y) j = {}") case True thus ?thesis by auto next case False hence "Y -` {(from_nat_into B) i} \ {} \ Y -` {(from_nat_into B) j} \ {}" unfolding countable_preimages_def by meson have "(infinite B) \ (finite B \ i < card B \ j < card B)" using False unfolding countable_preimages_def by meson have "(from_nat_into B) i \ (from_nat_into B) j" by (metis False assms(1) assms(2) bij_betw_def countable_preimages_def from_nat_into_inj from_nat_into_inj_infinite lessThan_iff to_nat_on_finite) thus ?thesis proof - have f1: "\A f n. if infinite A \ finite A \ n < card A then countable_preimages A f n = f -` {from_nat_into A n::'a} else countable_preimages A f n = ({}::'b set)" by (meson countable_preimages_def) then have f2: "infinite B \ finite B \ i < card B" by (metis (no_types) False) have "infinite B \ finite B \ j < card B" using f1 by (meson False) then show ?thesis using f2 f1 \from_nat_into B i \ from_nat_into B j\ by fastforce qed qed lemma count_pre_surj: assumes "countable B" and "w \ Y -`B" shows "\i. w \ (countable_preimages B Y) i" proof (cases "finite B") case True have "\ i < card B. (from_nat_into B) i = Y w" by (metis True assms(1) assms(2) bij_betw_def from_nat_into_to_nat_on image_eqI lessThan_iff to_nat_on_finite vimageE) from this obtain i where "i< card B" and "(from_nat_into B) i = Y w" by blast hence "w \ (countable_preimages B Y) i" by (simp add: countable_preimages_def) thus "\i. w \ (countable_preimages B Y) i" by auto next case False hence "\ i. (from_nat_into B) i = Y w" by (meson assms(1) assms(2) from_nat_into_to_nat_on vimageE) from this obtain i where "(from_nat_into B) i = Y w" by blast hence "w \ (countable_preimages B Y) i" by (simp add: False countable_preimages_def) thus "\i. w \ (countable_preimages B Y) i" by auto qed lemma count_pre_img: assumes "x \ (countable_preimages B Y) n" shows "Y x = (from_nat_into B) n" proof - have "x\ Y -` {(from_nat_into B) n}" using assms unfolding countable_preimages_def by (meson empty_iff) thus ?thesis by simp qed lemma count_pre_union_img: assumes "countable B" shows "Y -`B = (\ i. (countable_preimages B Y) i)" proof (cases "B = {}") case False have "Y -`B \ (\ i. (countable_preimages B Y) i)" by (simp add: assms count_pre_surj subset_eq) moreover have "(\ i. (countable_preimages B Y) i) \ Y -`B" proof - have f1: "\b A f n. (b::'b) \ countable_preimages A f n \ (f b::'a) = from_nat_into A n" by (meson count_pre_img) have "range (from_nat_into B) = B" by (meson False assms range_from_nat_into) then show ?thesis using f1 by blast qed ultimately show ?thesis by simp next case True hence "\ i. (countable_preimages B Y) i = {}" unfolding countable_preimages_def by simp hence "(\ i. (countable_preimages B Y) i) = {}" by auto moreover have "Y -`B = {}" using True by simp ultimately show ?thesis by simp qed lemma count_pre_meas: assumes "point_measurable M (space N) Y" and "B\ space N" and "countable B" shows "\i. (countable_preimages B Y) i \ space M \ sets M" proof fix i have "Y -`B = (\ i. (countable_preimages B Y) i)" using assms by (simp add: count_pre_union_img) show "countable_preimages B Y i \ space M \ sets M" proof (cases "countable_preimages B Y i = {}") case True thus ?thesis by simp next case False from this obtain y where "y \ countable_preimages B Y i" by auto hence "countable_preimages B Y i = Y -`{Y y}" by (metis False count_pre_img countable_preimages_def) have "Y y = from_nat_into B i" by (meson \y \ countable_preimages B Y i\ count_pre_img) hence "Y y \ space N" by (metis UNIV_I UN_I \y \ countable_preimages B Y i\ \Y -`B = (\ i. (countable_preimages B Y) i)\ assms(2) empty_iff from_nat_into subsetCE vimage_empty) moreover have "Y y \ range Y" by simp thus ?thesis by (metis IntI \countable_preimages B Y i = Y -` {Y y}\ assms(1) calculation point_measurable_def) qed qed lemma disct_fct_point_measurable: assumes "disc_fct f" and "point_measurable M (space N) f" shows "f\ measurable M N" unfolding measurable_def proof show "f \ space M \ space N \ (\y\sets N. f -` y \ space M \ sets M)" proof show "f \ space M \ space N" using assms unfolding point_measurable_def by auto show "\y\sets N. f -` y \ space M \ sets M" proof fix y assume "y\ sets N" let ?imY = "range f \ y" have "f-`y = f-`?imY" by auto moreover have "countable ?imY" using assms unfolding disc_fct_def by auto ultimately have "f -`y = (\ i. (countable_preimages ?imY f) i)" using assms count_pre_union_img by metis hence yeq: "f -` y \ space M = (\ i. ((countable_preimages ?imY f) i) \ space M)" by auto have "\i. countable_preimages ?imY f i \ space M \ sets M" by (metis \countable (range f \ y)\ \y \ sets N\ assms(2) inf_le2 le_inf_iff count_pre_meas sets.Int_space_eq1) hence "(\ i. ((countable_preimages ?imY f) i) \ space M) \ sets M" by blast thus "f -` y \ space M \ sets M" using yeq by simp qed qed qed lemma set_point_measurable: assumes "point_measurable M (space N) Y" and "B \ space N" and "countable B" shows "(Y -`B) \ space M \ sets M" proof - have "Y -`B = (\ i. (countable_preimages B Y) i)" using assms by (simp add: count_pre_union_img) hence "Y -`B \ space M = (\ i. ((countable_preimages B Y) i \ space M))" by auto have "\i. (countable_preimages B Y) i \ space M \ sets M" using assms by (simp add: count_pre_meas) hence "(\ i. ((countable_preimages B Y) i \ space M)) \ sets M" by blast show ?thesis using \(\i. countable_preimages B Y i \ space M) \ sets M\ \Y -` B \ space M = (\i. countable_preimages B Y i \ space M)\ by auto qed subsection \Definition of explicit conditional expectation\ text \This section is devoted to an explicit computation of a conditional expectation for random variables that have a countable codomain. More precisely, the computed random variable is almost everywhere equal to a conditional expectation of the random variable under consideration.\ definition img_dce where "img_dce M Y X = (\ y. if measure M ((Y -` {y}) \ space M) = 0 then 0 else ((integral\<^sup>L M (\w. ((X w) * (indicator ((Y -`{y})\ space M) w))))/(measure M ((Y -` {y}) \ space M))))" definition expl_cond_expect where "expl_cond_expect M Y X = (img_dce M Y X) \ Y" lemma nn_expl_cond_expect_pos: assumes "\w \ space M. 0 \ X w" shows "\ w\ space M. 0 \ (expl_cond_expect M Y X) w" proof fix w assume space: "w\ space M" show "0 \ (expl_cond_expect M Y X) w" proof (cases "measure M ((Y -` {Y w})\ space M) = 0") case True thus "0 \ (expl_cond_expect M Y X) w" unfolding expl_cond_expect_def img_dce_def by simp next case False hence "Y -`{Y w} \ space M \ sets M" using measure_notin_sets by blast let ?indA = "((\ x. indicator ((Y -`{Y w})\ space M) x))" have "\w \ space M. 0 \ (X w) * (?indA w)" by (simp add: assms) hence "0 \ (integral\<^sup>L M (\w. ((X w) * (?indA w))))" by simp moreover have "(expl_cond_expect M Y X) w = (integral\<^sup>L M (\w. ((X w) * (?indA w)))) / (measure M ((Y -` {Y w})\ space M))" unfolding expl_cond_expect_def img_dce_def using False by simp moreover have "0 < measure M ((Y -` {Y w}) \ space M)" using False by (simp add: zero_less_measure_iff) ultimately show "0 \ (expl_cond_expect M Y X) w" by simp qed qed lemma expl_cond_expect_const: assumes "Y w = Y y" shows "expl_cond_expect M Y X w = expl_cond_expect M Y X y" unfolding expl_cond_expect_def img_dce_def by (simp add: assms) lemma expl_cond_exp_cong: assumes "\w\space M. X w = Z w" shows "\w\ space M. expl_cond_expect M Y X w = expl_cond_expect M Y Z w" unfolding expl_cond_expect_def img_dce_def by (metis (no_types, lifting) Bochner_Integration.integral_cong assms(1) o_apply) (* example of a proof that IMO takes too long in Isabelle *) lemma expl_cond_exp_add: assumes "integrable M X" and "integrable M Z" shows "\w\ space M. expl_cond_expect M Y (\x. X x + Z x) w = expl_cond_expect M Y X w + expl_cond_expect M Y Z w" proof fix w assume "w\ space M" define prY where "prY = measure M ((Y -` {Y w}) \ space M)" show "expl_cond_expect M Y (\x. X x + Z x) w = expl_cond_expect M Y X w + expl_cond_expect M Y Z w" proof (cases "prY = 0") case True thus ?thesis unfolding expl_cond_expect_def img_dce_def prY_def by simp next case False hence "(Y -` {Y w}) \ space M \ sets M" unfolding prY_def using measure_notin_sets by blast let ?indA = "indicator ((Y -` {Y w}) \ space M)::('a\real)" have "integrable M (\x. X x * ?indA x)" using \Y -` {Y w} \ space M \ sets M\ assms(1) integrable_real_mult_indicator by blast moreover have "integrable M (\x. Z x * ?indA x)" using \Y -` {Y w} \ space M \ sets M\ assms(2) integrable_real_mult_indicator by blast ultimately have "integral\<^sup>L M (\x. X x * ?indA x + Z x * ?indA x) = integral\<^sup>L M (\x. X x * ?indA x) + integral\<^sup>L M (\x. Z x * ?indA x)" using Bochner_Integration.integral_add by blast moreover have "\x\ space M. X x * ?indA x + Z x * ?indA x = (X x + Z x) * ?indA x" by (simp add: indicator_def) ultimately have fsteq: "integral\<^sup>L M (\x. (X x + Z x) * ?indA x) = integral\<^sup>L M (\x. X x * ?indA x) + integral\<^sup>L M (\x. Z x * ?indA x)" by (metis (no_types, lifting) Bochner_Integration.integral_cong) have "integral\<^sup>L M (\x. (X x + Z x) * ?indA x/prY) = integral\<^sup>L M (\x. (X x + Z x) * ?indA x)/prY" by simp also have "... = integral\<^sup>L M (\x. X x * ?indA x)/prY + integral\<^sup>L M (\x. Z x * ?indA x)/prY" using fsteq by (simp add: add_divide_distrib) also have "... = integral\<^sup>L M (\x. X x * ?indA x/prY) + integral\<^sup>L M (\x. Z x * ?indA x/prY)" by auto finally have "integral\<^sup>L M (\x. (X x + Z x) * ?indA x/prY) = integral\<^sup>L M (\x. X x * ?indA x/prY) + integral\<^sup>L M (\x. Z x * ?indA x/prY)" . thus ?thesis using False unfolding expl_cond_expect_def img_dce_def by (simp add: add_divide_distrib fsteq) qed qed lemma expl_cond_exp_diff: assumes "integrable M X" and "integrable M Z" shows "\w\ space M. expl_cond_expect M Y (\x. X x - Z x) w = expl_cond_expect M Y X w - expl_cond_expect M Y Z w" proof fix w assume "w\ space M" define prY where "prY = measure M ((Y -` {Y w}) \ space M)" show "expl_cond_expect M Y (\x. X x - Z x) w = expl_cond_expect M Y X w - expl_cond_expect M Y Z w" proof (cases "prY = 0") case True thus ?thesis unfolding expl_cond_expect_def img_dce_def prY_def by simp next case False hence "(Y -` {Y w}) \ space M \ sets M" unfolding prY_def using measure_notin_sets by blast let ?indA = "indicator ((Y -` {Y w}) \ space M)::('a\real)" have "integrable M (\x. X x * ?indA x)" using \Y -` {Y w} \ space M \ sets M\ assms(1) integrable_real_mult_indicator by blast moreover have "integrable M (\x. Z x * ?indA x)" using \Y -` {Y w} \ space M \ sets M\ assms(2) integrable_real_mult_indicator by blast ultimately have "integral\<^sup>L M (\x. X x * ?indA x - Z x * ?indA x) = integral\<^sup>L M (\x. X x * ?indA x) - integral\<^sup>L M (\x. Z x * ?indA x)" using Bochner_Integration.integral_diff by blast moreover have "\x\ space M. X x * ?indA x - Z x * ?indA x = (X x - Z x) * ?indA x" by (simp add: indicator_def) ultimately have fsteq: "integral\<^sup>L M (\x. (X x - Z x) * ?indA x) = integral\<^sup>L M (\x. X x * ?indA x) - integral\<^sup>L M (\x. Z x * ?indA x)" by (metis (no_types, lifting) Bochner_Integration.integral_cong) have "integral\<^sup>L M (\x. (X x - Z x) * ?indA x/prY) = integral\<^sup>L M (\x. (X x - Z x) * ?indA x)/prY" by simp also have "... = integral\<^sup>L M (\x. X x * ?indA x)/prY - integral\<^sup>L M (\x. Z x * ?indA x)/prY" using fsteq by (simp add: diff_divide_distrib) also have "... = integral\<^sup>L M (\x. X x * ?indA x/prY) - integral\<^sup>L M (\x. Z x * ?indA x/prY)" by auto finally have "integral\<^sup>L M (\x. (X x - Z x) * ?indA x/prY) = integral\<^sup>L M (\x. X x * ?indA x/prY) - integral\<^sup>L M (\x. Z x * ?indA x/prY)" . thus ?thesis using False unfolding expl_cond_expect_def img_dce_def by (simp add: diff_divide_distrib fsteq) qed qed lemma expl_cond_expect_prop_sets: assumes "disc_fct Y" and "point_measurable M (space N) Y" and "D = {w\ space M. Y w \ space N \ (P (expl_cond_expect M Y X w))}" shows "D\ sets M" proof - let ?C = "{y \ (Y`(space M)) \ (space N). P (img_dce M Y X y)}" have "space M \ UNIV" by simp hence "Y`(space M) \ range Y" by auto hence "countable (Y`(space M))" using assms countable_subset unfolding disc_fct_def by auto hence "countable ?C" using assms unfolding disc_fct_def by auto have eqset: "D = (\ b\ ?C. Y-`{b})\ space M" proof show "D\ (\ b\ ?C. Y-`{b})\ space M" proof fix w assume "w\ D" hence "w\ space M \ Y w \ (space N) \ (P (expl_cond_expect M Y X w))" by (simp add: assms) hence "P (img_dce M Y X (Y w))" by (simp add: expl_cond_expect_def) hence "Y w \ ?C" using \w \ space M \ Y w \ space N \ P (expl_cond_expect M Y X w)\ by blast thus "w\ (\ b\ ?C. Y-`{b})\ space M" using \w \ space M \ Y w \ space N \ P (expl_cond_expect M Y X w)\ by blast qed show "(\ b\ ?C. Y-`{b})\ space M \ D" proof fix w assume "w\ (\ b\ ?C. Y-`{b})\ space M" from this obtain b where "b\ ?C \ w\ Y-`{b}" by auto note bprops = this hence "Y w = b" by auto hence "Y w\ space N" using bprops by simp show "w \ D" by (metis (mono_tags, lifting) IntE \Y w = b\ \w \ (\b\?C. Y -` {b}) \ space M\ assms(3) bprops mem_Collect_eq o_apply expl_cond_expect_def) qed qed also have "... = (\ b\ ?C. Y-`{b}\ space M)" by blast finally have "D = (\ b\ ?C. Y-`{b}\ space M)". have "\b\ ?C. Y-`{b} \ space M \ sets M" using assms unfolding point_measurable_def by auto hence "(\ b\ ?C. Y-`{b}\ space M) \ sets M" using \countable ?C\ by blast thus ?thesis using \D = (\b\?C. Y -` {b} \ space M)\ by blast qed lemma expl_cond_expect_prop_sets2: assumes "disc_fct Y" and "point_measurable (fct_gen_subalgebra M N Y) (space N) Y" and "D = {w\ space M. Y w \ space N \ (P (expl_cond_expect M Y X w))}" shows "D\ sets (fct_gen_subalgebra M N Y)" proof - let ?C = "{y \ (Y`(space M)) \ (space N). P (img_dce M Y X y)}" have "space M \ UNIV" by simp hence "Y`(space M) \ range Y" by auto hence "countable (Y`(space M))" using assms countable_subset unfolding disc_fct_def by auto hence "countable ?C" using assms unfolding disc_fct_def by auto have eqset: "D = (\ b\ ?C. Y-`{b})\ space M" proof show "D\ (\ b\ ?C. Y-`{b})\ space M" proof fix w assume "w\ D" hence "w\ space M \ Y w \ (space N) \ (P (expl_cond_expect M Y X w))" by (simp add: assms) hence "P (img_dce M Y X (Y w))" by (simp add: expl_cond_expect_def) hence "Y w \ ?C" using \w \ space M \ Y w \ space N \ P (expl_cond_expect M Y X w)\ by blast thus "w\ (\ b\ ?C. Y-`{b})\ space M" using \w \ space M \ Y w \ space N \ P (expl_cond_expect M Y X w)\ by blast qed show "(\ b\ ?C. Y-`{b})\ space M \ D" proof fix w assume "w\ (\ b\ ?C. Y-`{b})\ space M" from this obtain b where "b\ ?C \ w\ Y-`{b}" by auto note bprops = this hence "Y w = b" by auto hence "Y w\ space N" using bprops by simp show "w \ D" by (metis (mono_tags, lifting) IntE \Y w = b\ \w \ (\b\?C. Y -` {b}) \ space M\ assms(3) bprops mem_Collect_eq o_apply expl_cond_expect_def) qed qed also have "... = (\ b\ ?C. Y-`{b}\ space M)" by blast finally have "D = (\ b\ ?C. Y-`{b}\ space M)". have "space M = space (fct_gen_subalgebra M N Y)" by (simp add: fct_gen_subalgebra_space) hence "\b\ ?C. Y-`{b} \ space M \ sets (fct_gen_subalgebra M N Y)" using assms unfolding point_measurable_def by auto hence "(\ b\ ?C. Y-`{b}\ space M) \ sets (fct_gen_subalgebra M N Y)" using \countable ?C\ by blast thus ?thesis using \D = (\b\?C. Y -` {b} \ space M)\ by blast qed lemma expl_cond_expect_disc_fct: assumes "disc_fct Y" shows "disc_fct (expl_cond_expect M Y X)" using assms unfolding disc_fct_def expl_cond_expect_def by (metis countable_image image_comp) lemma expl_cond_expect_point_meas: assumes "disc_fct Y" and "point_measurable M (space N) Y" shows "point_measurable M UNIV (expl_cond_expect M Y X)" proof - have "disc_fct (expl_cond_expect M Y X)" using assms by (simp add: expl_cond_expect_disc_fct) show ?thesis unfolding point_measurable_def proof show "(expl_cond_expect M Y X)`space M \ UNIV" by simp show "\r\range (expl_cond_expect M Y X) \ UNIV. expl_cond_expect M Y X -` {r} \ space M \ sets M" proof fix r assume "r\ range (expl_cond_expect M Y X) \ UNIV" let ?D = "{w \ space M. Y w \ space N \ (expl_cond_expect M Y X w) = r}" have "?D \ sets M" using expl_cond_expect_prop_sets[of Y M N ?D "\x. x = r" X] using assms by simp moreover have "expl_cond_expect M Y X -`{r}\ space M = ?D" proof show "expl_cond_expect M Y X -`{r}\ space M \ ?D" proof fix w assume "w\ expl_cond_expect M Y X -`{r}\ space M" hence "Y w \ space N" by (meson IntD2 assms(1) assms(2) disct_fct_point_measurable measurable_space) thus "w \ ?D" using \w \ expl_cond_expect M Y X -` {r} \ space M\ by blast qed show "?D \ expl_cond_expect M Y X -`{r}\ space M" proof fix w assume "w\ ?D" thus "w\ expl_cond_expect M Y X -`{r}\ space M" by blast qed qed ultimately show "expl_cond_expect M Y X -` {r} \ space M \ sets M" by simp qed qed qed lemma expl_cond_expect_borel_measurable: assumes "disc_fct Y" and "point_measurable M (space N) Y" shows "(expl_cond_expect M Y X) \ borel_measurable M" using expl_cond_expect_point_meas[of Y M] assms disct_fct_point_measurable[of "expl_cond_expect M Y X"] by (simp add: expl_cond_expect_disc_fct) lemma expl_cond_exp_borel: assumes "Y \ space M \ space N" and "disc_fct Y" and "\r\ range Y\ space N. \A\ sets N. range Y\ A = {r}" shows "(expl_cond_expect M Y X) \ borel_measurable (fct_gen_subalgebra M N Y)" proof (intro borel_measurableI) fix S::"real set" assume "open S" show "expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y) \ sets (fct_gen_subalgebra M N Y)" proof (rule expl_cond_expect_prop_sets2) show "disc_fct Y" using assms by simp show "point_measurable (fct_gen_subalgebra M N Y) (space N) Y" using assms by (simp add: singl_meas_if) show "expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y) = {w \ space M. Y w \ space N \ (expl_cond_expect M Y X w) \ S}" proof show " expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y) \ {w \ space M. Y w \ space N \ expl_cond_expect M Y X w \ S}" proof fix x assume asm: "x \ expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y)" hence "expl_cond_expect M Y X x \ S" by auto moreover have "x\ space M" using asm by (simp add:fct_gen_subalgebra_space) ultimately show "x \{w \ space M. Y w \ space N \ expl_cond_expect M Y X w \ S}" using assms by auto qed show "{w \ space M. Y w \ space N \ expl_cond_expect M Y X w \ S} \ expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y)" proof fix x assume asm2: "x \ {w \ space M. Y w \ space N \ expl_cond_expect M Y X w \ S}" hence "x\ space (fct_gen_subalgebra M N Y)" by (simp add:fct_gen_subalgebra_space) moreover have "x \ expl_cond_expect M Y X -`S" using asm2 by simp ultimately show "x \ expl_cond_expect M Y X -` S \ space (fct_gen_subalgebra M N Y)" by simp qed qed qed qed lemma expl_cond_expect_indic_borel_measurable: assumes "disc_fct Y" and "point_measurable M (space N) Y" and "B\ space N" and "countable B" shows "(\w. expl_cond_expect M Y X w * indicator (countable_preimages B Y n \ space M) w)\ borel_measurable M" proof - have "countable_preimages B Y n \ space M \ sets M" using assms by (auto simp add: count_pre_meas) have "(expl_cond_expect M Y X)\ borel_measurable M" using expl_cond_expect_point_meas[of Y M N X] assms disct_fct_point_measurable[of "expl_cond_expect M Y X"] by (simp add: expl_cond_expect_disc_fct) moreover have "(indicator (countable_preimages B Y n \ space M))\ borel_measurable M" using \countable_preimages B Y n \ space M \ sets M\ borel_measurable_indicator by blast ultimately show ?thesis using borel_measurable_times by blast qed lemma (in finite_measure) dce_prod: assumes "point_measurable M (space N) Y" and "integrable M X" and "\ w\ space M. 0 \ X w" shows "\ w. (Y w) \ space N \ (expl_cond_expect M Y X) w * measure M ((Y -` {Y w})\ space M) = integral\<^sup>L M (\y. (X y) * (indicator ((Y -`{Y w})\ space M) y))" proof (intro allI impI) fix w assume "Y w\ space N" let ?indY = "(\y. indicator ((Y -`{Y w})\ space M) y)::'a \ real" show "expl_cond_expect M Y X w * measure M ((Y -` {Y w})\ space M) = integral\<^sup>L M (\y. (X y) * ?indY y) " proof (cases "AE y in M. ?indY y = 0") case True (* Very long proof, Sledgehammer was lost. Everything had to be detailed *) hence "emeasure M ((Y -` {Y w})\ space M) = 0" proof - have "AE y in M. y \ Y -` {Y w} \ space M" using True eventually_elim2 by auto hence "\N\ null_sets M.{x\ space M. \(x\ Y -` {Y w} \ space M)} \ N" using eventually_ae_filter[of "\x. x \ Y -` {Y w} \ space M" M] by simp hence "\N\ null_sets M. {x\ space M. x\ Y -` {Y w} \ space M} \ N" by simp from this obtain N where "N\ null_sets M" and "{x\ space M. x\ Y -` {Y w} \ space M} \ N" by auto note Nprops = this have "{x\ space M. x\ Y -` {Y w}} \ N" using Nprops by auto hence "emeasure M {x\ space M. x\ Y -` {Y w}} \ emeasure M N" by (simp add: emeasure_mono Nprops(1) null_setsD2) thus ?thesis by (metis (no_types, lifting) Collect_cong Int_def Nprops(1) le_zero_eq null_setsD1) qed hence "enn2real (emeasure M ((Y -` {Y w})\ space M)) = 0" by simp hence "measure M ((Y -` {Y w})\ space M) = 0" unfolding measure_def by simp hence lhs: "expl_cond_expect M Y X w = 0" unfolding expl_cond_expect_def img_dce_def by simp have zer: "AE y in M. (X y) * ?indY y = (\y. 0) y" using True by auto hence rhs: "integral\<^sup>L M (\y. (X y) * ?indY y) = 0" proof - have "\ w\ space M. 0 \ X w * ?indY w" using assms by simp have "integrable M (\y. (X y) * ?indY y)" using assms - by (metis (mono_tags, lifting) IntI UNIV_I \Y w \ space N\ image_eqI integrable_cong integrable_real_mult_indicator point_measurable_def) + by (metis (mono_tags, lifting) IntI UNIV_I \Y w \ space N\ image_eqI Bochner_Integration.integrable_cong integrable_real_mult_indicator point_measurable_def) hence "(\y. (X y) * ?indY y) \ borel_measurable M" by blast thus ?thesis using zer integral_cong_AE[of "(\y. (X y) * ?indY y)" M "\y. 0"] by simp qed thus "expl_cond_expect M Y X w*measure M ((Y -` {Y w})\ space M) = integral\<^sup>L M (\y. (X y) * ?indY y)" using lhs rhs by simp next case False hence "\(AE y in M. y \ (Y -`{Y w})\ space M)" by (simp add: indicator_eq_0_iff) hence "emeasure M ((Y -` {Y w})\ space M) \ 0" proof - have "(Y -` {Y w})\ space M\ sets M" by (meson IntI UNIV_I \Y w \ space N\ assms(1) image_eqI point_measurable_def) have "(Y -` {Y w})\ space M \ null_sets M" using \\ (AE y in M. y \ Y -` {Y w} \ space M)\ eventually_ae_filter by blast thus ?thesis using \Y -` {Y w} \ space M \ sets M\ by blast qed hence "measure M ((Y -` {Y w})\ space M) \ 0" by (simp add: emeasure_eq_measure) thus "expl_cond_expect M Y X w* measure M ((Y -` {Y w})\ space M) = integral\<^sup>L M (\y. (X y) * ?indY y)" unfolding expl_cond_expect_def img_dce_def using o_apply by auto qed qed lemma expl_cond_expect_const_exp: shows "integral\<^sup>L M (\y. expl_cond_expect M Y X w * (indicator (Y -` {Y w} \ space M)) y) = integral\<^sup>L M (\y. expl_cond_expect M Y X y * (indicator (Y -` {Y w} \ space M)) y)" proof - let ?ind = "(indicator (Y -` {Y w} \ space M))" have "\ y\ space M. expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof fix y assume "y\ space M" show "expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof (cases "y\ Y -` {Y w} \ space M") case False thus ?thesis by simp next case True hence "Y w = Y y" by auto hence "expl_cond_expect M Y X w = expl_cond_expect M Y X y" using expl_cond_expect_const[of Y w y M X] by simp thus ?thesis by simp qed qed thus ?thesis by (meson Bochner_Integration.integral_cong) qed lemma nn_expl_cond_expect_const_exp: assumes "\w\ space M. 0 \ X w" shows "integral\<^sup>N M (\y. expl_cond_expect M Y X w * (indicator (Y -` {Y w} \ space M)) y) = integral\<^sup>N M (\y. expl_cond_expect M Y X y * (indicator (Y -` {Y w} \ space M)) y)" proof - let ?ind = "(indicator (Y -` {Y w} \ space M))" have forall: "\ y\ space M. expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof fix y assume "y\ space M" show "expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof (cases "y\ Y -` {Y w} \ space M") case False thus ?thesis by simp next case True hence "Y w = Y y" by auto hence "expl_cond_expect M Y X w = expl_cond_expect M Y X y" using expl_cond_expect_const[of Y] by blast thus ?thesis by simp qed qed show ?thesis by (metis (no_types, lifting) forall nn_integral_cong) qed lemma (in finite_measure) nn_expl_cond_bounded: assumes "\w\ space M. 0 \ X w" and "integrable M X" and "point_measurable M (space N) Y" and "w\ space M" and "Y w\ space N" shows "integral\<^sup>N M (\y. expl_cond_expect M Y X y * (indicator (Y -` {Y w} \ space M)) y) < \" proof - let ?ind = "(indicator (Y -` {Y w} \ space M))::'a\real" have "0 \ expl_cond_expect M Y X w" using assms nn_expl_cond_expect_pos[of M X Y] by simp have "integrable M (\y. expl_cond_expect M Y X w * ?ind y)" proof - have eq: "(Y -`{Y w} \ space M) \ space M = (Y -`{Y w} \ space M)" by auto have "(Y -` {Y w} \ space M) \ sets M" using assms by (simp add: point_measurable_def) moreover have "emeasure M (Y -`{Y w} \ space M) < \" by (simp add: inf.strict_order_iff) ultimately have "integrable M (\y. ?ind y)" using integrable_indicator_iff[of M "(Y -`{Y w} \ space M)"] by simp thus ?thesis using integrable_mult_left_iff[of M "expl_cond_expect M Y X w" "?ind"] by blast qed have "\y\ space M. 0 \ expl_cond_expect M Y X w * ?ind y" using \0 \ expl_cond_expect M Y X w\ mult_nonneg_nonneg by blast hence "\y\ space M. expl_cond_expect M Y X w * ?ind y = norm (expl_cond_expect M Y X w * ?ind y)" by auto hence inf: "integral\<^sup>N M (\y. expl_cond_expect M Y X w * ?ind y) < \" using integrable_iff_bounded[of M "(\y. expl_cond_expect M Y X w * ?ind y)"] \0 \ expl_cond_expect M Y X w\ real_norm_def nn_integral_cong by (metis (no_types, lifting) \integrable M (\y. expl_cond_expect M Y X w * indicator (Y -` {Y w} \ space M) y)\) have "integral\<^sup>N M (\y. expl_cond_expect M Y X y * ?ind y) = integral\<^sup>N M (\y. expl_cond_expect M Y X w * ?ind y)" using assms by (simp add: nn_expl_cond_expect_const_exp) also have "... < \" using inf by simp finally show ?thesis . qed lemma (in finite_measure) count_prod: fixes Y::"'a\'b" assumes "B\ space N" and "point_measurable M (space N) Y" and "integrable M X" and "\ w \ space M. 0 \ X w" shows "\i. integral\<^sup>L M (\y. (X y) * (indicator (countable_preimages B Y i \ space M)) y) = integral\<^sup>L M (\y. (expl_cond_expect M Y X y) * (indicator (countable_preimages B Y i \ space M)) y)" proof fix i show "integral\<^sup>L M (\y. X y * indicator (countable_preimages B Y i \ space M) y) = integral\<^sup>L M (\y. expl_cond_expect M Y X y * indicator (countable_preimages B Y i \ space M) y)" proof (cases "countable_preimages B Y i \ space M = {}") case True thus ?thesis by simp next case False from this obtain w where "w\ countable_preimages B Y i" by auto hence "Y w = (from_nat_into B) i" by (meson count_pre_img) hence "Y w \ B" proof (cases "infinite B") case True thus ?thesis by (simp add: \Y w = from_nat_into B i\ from_nat_into infinite_imp_nonempty) next case False thus ?thesis by (metis Finite_Set.card_0_eq \Y w = from_nat_into B i\ \w \ countable_preimages B Y i\ countable_preimages_def equals0D from_nat_into gr_implies_not0) qed let ?ind = "(indicator (Y -` {Y w} \ space M))::'a\real" have "integral\<^sup>L M (\y. (X y) * (indicator (countable_preimages B Y i \ space M)) y) = integral\<^sup>L M (\y. X y * ?ind y)" by (metis (no_types, opaque_lifting) \Y w = from_nat_into B i\ \\thesis. (\w. w \ countable_preimages B Y i \ thesis) \ thesis\ countable_preimages_def empty_iff) also have "... = expl_cond_expect M Y X w * measure M (Y -` {Y w} \ space M)" using dce_prod[of N Y X] by (metis (no_types, lifting) \Y w \ B\ assms subsetCE) also have "... = expl_cond_expect M Y X w * (integral\<^sup>L M ?ind)" by auto also have "... = integral\<^sup>L M (\y. expl_cond_expect M Y X w * ?ind y)" by auto also have "... = integral\<^sup>L M (\y. expl_cond_expect M Y X y * ?ind y)" proof - have "\ y\ space M. expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof fix y assume "y\ space M" show "expl_cond_expect M Y X w * ?ind y = expl_cond_expect M Y X y * ?ind y" proof (cases "y\ Y -` {Y w} \ space M") case False thus ?thesis by simp next case True hence "Y w = Y y" by auto hence "expl_cond_expect M Y X w = expl_cond_expect M Y X y" using expl_cond_expect_const[of Y] by blast thus ?thesis by simp qed qed thus ?thesis by (meson Bochner_Integration.integral_cong) qed also have "... = integral\<^sup>L M (\y. expl_cond_expect M Y X y * indicator (countable_preimages B Y i \ space M) y)" by (metis (no_types, opaque_lifting) \Y w = from_nat_into B i\ \\thesis. (\w. w \ countable_preimages B Y i \ thesis) \ thesis\ countable_preimages_def empty_iff) finally show ?thesis . qed qed lemma (in finite_measure) count_pre_integrable: assumes "point_measurable M (space N) Y" and "disc_fct Y" and "B\ space N" and "countable B" and "integrable M X" and "\ w \ space M. 0 \ X w" shows "integrable M (\w. expl_cond_expect M Y X w * indicator (countable_preimages B Y n \ space M) w)" proof- have "integral\<^sup>L M (\y. (X y) * (indicator (countable_preimages B Y n \ space M)) y) = integral\<^sup>L M (\y. (expl_cond_expect M Y X y) * (indicator (countable_preimages B Y n \ space M)) y)" using assms count_prod by blast have "\w \ space M. 0 \ (expl_cond_expect M Y X w) * (indicator (countable_preimages B Y n \ space M)) w" by (simp add: assms nn_expl_cond_expect_pos) have "countable_preimages B Y n \ space M \ sets M" using count_pre_meas[of M] assms by auto hence "integrable M (\w. X w * indicator (countable_preimages B Y n \ space M) w)" using assms integrable_real_mult_indicator by blast show ?thesis proof (rule integrableI_nonneg) show "(\w. expl_cond_expect M Y X w * indicator (countable_preimages B Y n \ space M) w)\ borel_measurable M" proof - have "(expl_cond_expect M Y X)\ borel_measurable M" using expl_cond_expect_point_meas[of Y M N X] assms disct_fct_point_measurable[of "expl_cond_expect M Y X"] by (simp add: expl_cond_expect_disc_fct) moreover have "(indicator (countable_preimages B Y n \ space M))\ borel_measurable M" using \countable_preimages B Y n \ space M \ sets M\ borel_measurable_indicator by blast ultimately show ?thesis using borel_measurable_times by blast qed show "AE x in M. 0 \ expl_cond_expect M Y X x * indicator (countable_preimages B Y n \ space M) x" by (simp add: \\w\space M. 0 \ expl_cond_expect M Y X w * indicator (countable_preimages B Y n \ space M) w\) show "(\\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indicator (countable_preimages B Y n \ space M) x) \M) < \" proof (cases "countable_preimages B Y n \ space M = {}") case True thus ?thesis by simp next case False from this obtain w where "w\ countable_preimages B Y n\ space M" by auto hence "countable_preimages B Y n = Y -`{Y w}" by (metis IntD1 count_pre_img countable_preimages_def equals0D) have "w\ space M" using False using \w \ countable_preimages B Y n \ space M\ by blast moreover have "Y w \ space N" by (meson \w \ space M\ assms(1) assms(2) disct_fct_point_measurable measurable_space) ultimately show ?thesis using assms nn_expl_cond_bounded[of X N Y] using \countable_preimages B Y n = Y -` {Y w}\ by presburger qed qed qed lemma (in finite_measure) nn_cond_expl_is_cond_exp_tmp: assumes "\ w\ space M. 0 \ X w" and "integrable M X" and "disc_fct Y" and "point_measurable M (space M') Y" shows "\ A \ sets M'. integrable M (\w. ((expl_cond_expect M Y X) w) * (indicator ((Y -`A)\ (space M)) w)) \ integral\<^sup>L M (\w. (X w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator ((Y -`A) \ (space M))) w)" proof fix A assume "A \ sets M'" let ?imA = "A \ (range Y)" have "countable ?imA" using assms disc_fct_def by blast have "Y -`A = Y -`?imA" by auto define prY where "prY = countable_preimages ?imA Y" have un: "Y -`?imA = (\ i. prY i)" using \countable ?imA\ by (metis count_pre_union_img prY_def) have "(Y -`?imA) \ (space M) = (\ i. prY i) \ (space M)" using \Y -`A = Y -`?imA\ un by simp also have "... = (\ i. (prY i) \ (space M))" by blast finally have eq2: "(Y -`?imA) \ (space M) = (\ i. (prY i) \ (space M))". define indpre::"nat \ 'a \ real" where "indpre = (\ i x. (indicator ((prY i) \ (space M))) x)" have "\ i. indpre i \ borel_measurable M" proof fix i show "indpre i \ borel_measurable M" unfolding indpre_def prY_def proof (rule borel_measurable_indicator, cases "countable_preimages (A \ range Y) Y i \ space M = {}") case True thus "countable_preimages (A \ range Y) Y i \ space M \ sets M" by simp next case False from this obtain x where "x\ countable_preimages (A \ range Y) Y i \ space M" by blast hence "Y x \ space M'" by (metis Int_iff UN_I \A \ sets M'\ \prY \ countable_preimages (A \ range Y) Y\ imageE rangeI sets.sets_into_space subset_eq un vimage_eq) thus "countable_preimages (A \ range Y) Y i \ space M \ sets M" by (metis IntE IntI \x \ countable_preimages (A \ range Y) Y i \ space M\ assms(4) count_pre_img countable_preimages_def empty_iff point_measurable_def rangeI) qed qed have "\i. integrable M (\w. (X w) * indpre i w)" proof fix i show "integrable M (\w. (X w) * indpre i w)" unfolding indpre_def prY_def proof (rule integrable_real_mult_indicator) show "countable_preimages (A \ range Y) Y i \ space M \ sets M" proof (cases "countable_preimages (A \ range Y) Y i = {}") case True thus "countable_preimages (A \ range Y) Y i \ space M \ sets M" by (simp add: True) next case False hence "Y -` {(from_nat_into (A \ range Y)) i} \ {}" unfolding countable_preimages_def by meson have "(infinite (A \ range Y)) \ (finite (A \ range Y) \ i < card (A \ range Y))" using False unfolding countable_preimages_def by meson show ?thesis by (metis \A \ sets M'\ \countable (A \ range Y)\ assms(4) count_pre_meas le_inf_iff range_from_nat_into sets.Int_space_eq1 sets.empty_sets sets.sets_into_space subset_range_from_nat_into) qed show "integrable M X" using assms by simp qed qed hence prod_bm: "\ i. (\w. (X w) * indpre i w) \ borel_measurable M" by (simp add: assms(2) borel_measurable_integrable borel_measurable_times) have posprod: "\ i w. 0 \ (X w) * indpre i w" proof (intro allI) fix i fix w show "0 \ (X w) * indpre i w" by (metis IntE assms(1) indicator_pos_le indicator_simps(2) indpre_def mult_eq_0_iff mult_sign_intros(1)) qed let ?indA = "indicator ((Y -`(A \ range Y))\ (space M))::'a\real" have "\ i j. i \ j \ ((prY i) \ (space M)) \ ((prY j) \ (space M)) = {}" by (simp add: \countable (A \ range Y)\ \prY \ countable_preimages (A \ range Y) Y\ count_pre_disj inf_commute inf_sup_aci(3)) hence sumind: "\x. (\i. indpre i x) sums ?indA x" using \countable ?imA\ eq2 unfolding prY_def indpre_def by (metis indicator_sums) hence sumxlim: "\x. (\i. (X x) * indpre i x::real) sums ((X x) * indicator ((Y -`?imA) \ (space M)) x)" using \countable ?imA\ unfolding prY_def using sums_mult by blast hence sum: "\x. (\ i.((X x) * indpre i x)::real) = (X x) * indicator ((Y -`?imA) \ (space M)) x" by (metis sums_unique) hence b: "\ w. 0 \ (\ i.((X w) * indpre i w))" using suminf_nonneg by (metis \\x. (\i. X x * indpre i x) sums (X x * indicator (Y -` (A \ range Y) \ space M) x)\ posprod summable_def) have sumcondlim: "\x. (\i. (expl_cond_expect M Y X x) * indpre i x::real) sums ((expl_cond_expect M Y X x) * ?indA x)" using \countable ?imA\ unfolding prY_def using sums_mult sumind by blast have "integrable M (\w. (X w) * ?indA w)" proof (rule integrable_real_mult_indicator) show "Y -` (A\ range Y) \ space M \ sets M" using \A \ sets M'\ assms(3) assms(4) disct_fct_point_measurable measurable_sets by (metis \Y -` A = Y -` (A \ range Y)\) show "integrable M X" using assms by simp qed hence intsum: "integrable M (\w. (\i. ((X w) * indpre i w)))" using sum - integrable_cong[of M M "\ w.(X w) * (indicator ((Y -`A)\ (space M)) w)" "\w. (\ i.((X w) * indpre i w))"] + Bochner_Integration.integrable_cong[of M M "\ w.(X w) * (indicator ((Y -`A)\ (space M)) w)" "\w. (\ i.((X w) * indpre i w))"] using \Y -` A = Y -` (A \ range Y)\ by presburger have "integral\<^sup>L M (\w. (X w) * ?indA w) = integral\<^sup>L M (\w. (\ i.((X w) * indpre i w)))" using \Y -` A = Y -` (A \ range Y)\ sum by auto also have "... = \\<^sup>+ w. ((\ i. ((X w) * indpre i w))) \M" using nn_integral_eq_integral by (metis (mono_tags, lifting) AE_I2 intsum b nn_integral_cong) also have "(\\<^sup>+ w. ((\ i. ((X w) * indpre i w))) \M) = \\<^sup>+ w. ((\ i. ennreal ((X w) * indpre i w))) \M" using suminf_ennreal2 summable_def posprod sum sumxlim proof - { fix aa :: 'a have "\a. ennreal (\n. X a * indpre n a) = (\n. ennreal (X a * indpre n a))" by (metis (full_types) posprod suminf_ennreal2 summable_def sumxlim) then have "(\\<^sup>+ a. ennreal (\n. X a * indpre n a) \M) = (\\<^sup>+ a. (\n. ennreal (X a * indpre n a)) \M) \ ennreal (\n. X aa * indpre n aa) = (\n. ennreal (X aa * indpre n aa))" by metis } then show ?thesis by presburger qed also have "... = (\i. integral\<^sup>N M ((\i w. (X w) * indpre i w) i))" proof (intro nn_integral_suminf) fix i show "(\x. ennreal (X x * indpre i x))\ borel_measurable M" using measurable_compose_rev measurable_ennreal prod_bm by blast qed also have "... = (\i. integral\<^sup>N M ((\i w. (expl_cond_expect M Y X w) * indpre i w) i))" proof (intro suminf_cong) fix n have "A \ range Y \ space M'" using \A \ sets M'\ sets.Int_space_eq1 by auto have "integral\<^sup>N M (\w. (X w) * indpre n w) = integral\<^sup>L M (\w. (X w) * indpre n w)" using nn_integral_eq_integral[of M "\w. (X w) * indpre n w"] by (simp add: \\i. integrable M (\w. X w * indpre i w)\ posprod) also have "... = integral\<^sup>L M (\w. (expl_cond_expect M Y X) w * indpre n w)" proof - have "integral\<^sup>L M (\w. X w * indicator (countable_preimages (A \ range Y) Y n \ space M) w) = integral\<^sup>L M (\w. expl_cond_expect M Y X w * indicator (countable_preimages (A \ range Y) Y n \ space M) w)" using count_prod[of "A\ range Y" "M'" Y X ] assms \A \ range Y \ space M'\ by blast thus ?thesis using \indpre \ \i. indicator (prY i \ space M)\ prY_def by presburger qed also have "... = integral\<^sup>N M (\w. (expl_cond_expect M Y X) w * indpre n w)" proof - have "integrable M (\w. (expl_cond_expect M Y X) w * indpre n w)" unfolding indpre_def prY_def using count_pre_integrable assms \A \ range Y \ space M'\ \countable (A \ range Y)\ by blast moreover have "AE w in M. 0 \ (expl_cond_expect M Y X) w * indpre n w" by (simp add: \indpre \ \i. indicator (prY i \ space M)\ assms(1) nn_expl_cond_expect_pos) ultimately show ?thesis by (simp add:nn_integral_eq_integral) qed finally show "integral\<^sup>N M (\w. (X w) * indpre n w) = integral\<^sup>N M (\w. (expl_cond_expect M Y X) w * indpre n w)" . qed also have "... = integral\<^sup>N M (\w. \i. ((expl_cond_expect M Y X w) * indpre i w))" proof - have "(\x. (\i. ennreal (expl_cond_expect M Y X x * indpre i x))) = (\x. ennreal (\i. (expl_cond_expect M Y X x * indpre i x)))" proof- have posex: "\ i x. 0 \ (expl_cond_expect M Y X x) * (indpre i x)" by (metis IntE \indpre \ \i. indicator (prY i \ space M)\ assms(1) indicator_pos_le indicator_simps(2) mult_eq_0_iff mult_sign_intros(1) nn_expl_cond_expect_pos) have "\x. (\i. ennreal (expl_cond_expect M Y X x * indpre i x)) = (ennreal (\i. (expl_cond_expect M Y X x * indpre i x)))" proof fix x show "(\i. ennreal (expl_cond_expect M Y X x * indpre i x)) = (ennreal (\i. (expl_cond_expect M Y X x * indpre i x)))" using suminf_ennreal2[of "\i. (expl_cond_expect M Y X x * indpre i x)"] sumcondlim summable_def posex proof - have f1: "summable (\n. expl_cond_expect M Y X x * indpre n x)" using sumcondlim summable_def by blast obtain nn :: nat where "\ 0 \ expl_cond_expect M Y X x * indpre nn x \ \ summable (\n. expl_cond_expect M Y X x * indpre n x) \ ennreal (\n. expl_cond_expect M Y X x * indpre n x) = (\n. ennreal (expl_cond_expect M Y X x * indpre n x))" by (metis (full_types) \\\i. 0 \ expl_cond_expect M Y X x * indpre i x; summable (\i. expl_cond_expect M Y X x * indpre i x)\ \ (\i. ennreal (expl_cond_expect M Y X x * indpre i x)) = ennreal (\i. expl_cond_expect M Y X x * indpre i x)\) then show ?thesis using f1 posex by presburger qed qed thus ?thesis by simp qed have "\i. (\w. (expl_cond_expect M Y X w) * indpre i w) \ borel_measurable M" proof - show ?thesis using \\i. (indpre i)\ borel_measurable M\ assms(3) assms(4) borel_measurable_times expl_cond_expect_borel_measurable by blast qed hence "\i. (\x. ennreal (expl_cond_expect M Y X x * indpre i x))\ borel_measurable M" using measurable_compose_rev measurable_ennreal by blast thus ?thesis using nn_integral_suminf[of "(\i w. (expl_cond_expect M Y X w) * indpre i w)" M, symmetric] using \(\x. \i. ennreal (expl_cond_expect M Y X x * indpre i x)) = (\x. ennreal (\i. expl_cond_expect M Y X x * indpre i x))\ by auto qed also have "... = integral\<^sup>N M (\w. (expl_cond_expect M Y X w) * ?indA w)" using sumcondlim by (metis (no_types, lifting) sums_unique) also have "... = integral\<^sup>L M (\w. (expl_cond_expect M Y X w) * ?indA w)" proof - have scdint: "integrable M (\w. (expl_cond_expect M Y X w) * ?indA w)" proof - have rv: "(\w. (expl_cond_expect M Y X w) * indicator ((Y -`?imA) \ (space M)) w) \ borel_measurable M" proof - have "expl_cond_expect M Y X \ borel_measurable M" using expl_cond_expect_borel_measurable using assms by blast moreover have "(Y -`?imA) \ (space M) \ sets M" by (metis \A \ sets M'\ \Y -` A = Y -` (A \ range Y)\ assms(3) assms(4) disct_fct_point_measurable measurable_sets) ultimately show ?thesis using borel_measurable_indicator_iff borel_measurable_times by blast qed moreover have born: "integral\<^sup>N M (\w. ennreal (norm (expl_cond_expect M Y X w * ?indA w))) < \" proof - have "integral\<^sup>N M (\w. ennreal (norm (expl_cond_expect M Y X w * ?indA w))) = integral\<^sup>N M (\w. ennreal (expl_cond_expect M Y X w * ?indA w))" proof - have "\w\ space M. norm (expl_cond_expect M Y X w * ?indA w) = expl_cond_expect M Y X w * ?indA w" using nn_expl_cond_expect_pos by (simp add: nn_expl_cond_expect_pos assms(1)) thus ?thesis by (metis (no_types, lifting) nn_integral_cong) qed thus ?thesis by (metis (no_types, lifting) \(\i. \\<^sup>+ x. ennreal (X x * indpre i x) \M) = (\i. \\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indpre i x) \M)\ \(\i. \\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indpre i x) \M) = (\\<^sup>+ x. ennreal (\i. expl_cond_expect M Y X x * indpre i x) \M)\ \(\\<^sup>+ w. (\i. ennreal (X w * indpre i w)) \M) = (\i. \\<^sup>+ x. ennreal (X x * indpre i x) \M)\ \(\\<^sup>+ x. ennreal (\i. X x * indpre i x) \M) = (\\<^sup>+ w. (\i. ennreal (X w * indpre i w)) \M)\ \(\\<^sup>+ x. ennreal (\i. expl_cond_expect M Y X x * indpre i x) \M) = (\\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x) \M)\ \ennreal (integral\<^sup>L M (\w. \i. X w * indpre i w)) = (\\<^sup>+ x. ennreal (\i. X x * indpre i x) \M)\ ennreal_less_top infinity_ennreal_def) qed show "integrable M (\w. (expl_cond_expect M Y X w) * ?indA w)" proof (rule iffD2[OF integrable_iff_bounded]) show "((\w. expl_cond_expect M Y X w * indicator (Y -` (A \ range Y) \ space M) w) \ borel_measurable M) \ ((\\<^sup>+ x. (ennreal (norm (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x))) \M) < \)" proof show "(\w. expl_cond_expect M Y X w * indicator (Y -` (A \ range Y) \ space M) w)\ borel_measurable M" using rv by simp show "(\\<^sup>+ x. ennreal (norm (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x)) \M) < \" using born by simp qed qed qed moreover have "\w\ space M. 0 \ (expl_cond_expect M Y X w) * indicator ((Y -`?imA) \ (space M)) w" by (simp add: assms(1) nn_expl_cond_expect_pos) ultimately show ?thesis using nn_integral_eq_integral by (metis (mono_tags, lifting) AE_I2 nn_integral_cong) qed finally have myeq: "ennreal (integral\<^sup>L M (\w. (X w) * ?indA w)) = integral\<^sup>L M (\w. (expl_cond_expect M Y X w) * ?indA w)" . thus "integrable M (\w. expl_cond_expect M Y X w * indicator (Y -` A \ space M) w) \ integral\<^sup>L M (\w. X w * indicator (Y -` A \ space M) w) = integral\<^sup>L M (\w. expl_cond_expect M Y X w * indicator (Y -` A \ space M) w)" proof - have "0 \ integral\<^sup>L M (\w. X w * indicator (Y -` A \ space M) w)" using \Y -` A = Y -` (A \ range Y)\ b sum by fastforce moreover have "0 \ integral\<^sup>L M (\w. expl_cond_expect M Y X w * indicator (Y -` A \ space M) w)" by (simp add: assms(1) nn_expl_cond_expect_pos) ultimately have expeq: "integral\<^sup>L M (\w. X w * indicator (Y -` A \ space M) w) = integral\<^sup>L M (\w. expl_cond_expect M Y X w * indicator (Y -` A \ space M) w)" by (metis (mono_tags, lifting) Bochner_Integration.integral_cong \Y -` A = Y -` (A \ range Y)\ ennreal_inj myeq) have "integrable M (\w. (expl_cond_expect M Y X w) * ?indA w)" proof - have rv: "(\w. (expl_cond_expect M Y X w) * indicator ((Y -`?imA) \ (space M)) w) \ borel_measurable M" proof - have "expl_cond_expect M Y X \ borel_measurable M" using expl_cond_expect_borel_measurable using assms by blast moreover have "(Y -`?imA) \ (space M) \ sets M" by (metis \A \ sets M'\ \Y -` A = Y -` (A \ range Y)\ assms(3) assms(4) disct_fct_point_measurable measurable_sets) ultimately show ?thesis using borel_measurable_indicator_iff borel_measurable_times by blast qed moreover have born: "integral\<^sup>N M (\w. ennreal (norm (expl_cond_expect M Y X w * ?indA w))) < \" proof - have "integral\<^sup>N M (\w. ennreal (norm (expl_cond_expect M Y X w * ?indA w))) = integral\<^sup>N M (\w. ennreal (expl_cond_expect M Y X w * ?indA w))" proof - have "\w\ space M. norm (expl_cond_expect M Y X w * ?indA w) = expl_cond_expect M Y X w * ?indA w" using nn_expl_cond_expect_pos by (simp add: nn_expl_cond_expect_pos assms(1)) thus ?thesis by (metis (no_types, lifting) nn_integral_cong) qed thus ?thesis by (metis (no_types, lifting) \(\i. \\<^sup>+ x. ennreal (X x * indpre i x) \M) = (\i. \\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indpre i x) \M)\ \(\i. \\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indpre i x) \M) = (\\<^sup>+ x. ennreal (\i. expl_cond_expect M Y X x * indpre i x) \M)\ \(\\<^sup>+ w. (\i. ennreal (X w * indpre i w)) \M) = (\i. \\<^sup>+ x. ennreal (X x * indpre i x) \M)\ \(\\<^sup>+ x. ennreal (\i. X x * indpre i x) \M) = (\\<^sup>+ w. (\i. ennreal (X w * indpre i w)) \M)\ \(\\<^sup>+ x. ennreal (\i. expl_cond_expect M Y X x * indpre i x) \M) = (\\<^sup>+ x. ennreal (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x) \M)\ \ennreal (integral\<^sup>L M (\w. \i. X w * indpre i w)) = (\\<^sup>+ x. ennreal (\i. X x * indpre i x) \M)\ ennreal_less_top infinity_ennreal_def) qed show "integrable M (\w. (expl_cond_expect M Y X w) * ?indA w)" proof (rule iffD2[OF integrable_iff_bounded]) show "((\w. expl_cond_expect M Y X w * indicator (Y -` (A \ range Y) \ space M) w) \ borel_measurable M) \ ((\\<^sup>+ x. (ennreal (norm (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x))) \M) < \)" proof show "(\w. expl_cond_expect M Y X w * indicator (Y -` (A \ range Y) \ space M) w)\ borel_measurable M" using rv by simp show "(\\<^sup>+ x. ennreal (norm (expl_cond_expect M Y X x * indicator (Y -` (A \ range Y) \ space M) x)) \M) < \" using born by simp qed qed qed hence "integrable M (\w. expl_cond_expect M Y X w * indicator (Y -` A \ space M) w)" using \Y -` A = Y -` (A \ range Y)\ by presburger thus ?thesis using expeq by simp qed qed lemma (in finite_measure) nn_expl_cond_exp_integrable: assumes "\ w\ space M. 0 \ X w" and "integrable M X" and "disc_fct Y" and "point_measurable M (space N) Y" shows "integrable M (expl_cond_expect M Y X)" proof - have "Y-`space N \ space M = space M" by (meson assms(3) assms(4) disct_fct_point_measurable inf_absorb2 measurable_space subsetI vimageI) let ?indA = "indicator ((Y -`space N)\ (space M))::'a\real" have "\w\ space M. (?indA w)= (1::real)" by (simp add: \Y -` space N \ space M = space M\) hence "\w\ space M. ((expl_cond_expect M Y X) w) * (?indA w) = (expl_cond_expect M Y X) w" by simp moreover have "integrable M (\w. ((expl_cond_expect M Y X) w) * (?indA w))" using assms nn_cond_expl_is_cond_exp_tmp[of X Y N] by blast - ultimately show ?thesis by (metis (no_types, lifting) integrable_cong) + ultimately show ?thesis by (metis (no_types, lifting) Bochner_Integration.integrable_cong) qed lemma (in finite_measure) nn_cond_expl_is_cond_exp: assumes "\ w\ space M. 0 \ X w" and "integrable M X" and "disc_fct Y" and "point_measurable M (space N) Y" shows "\ A \ sets N. integral\<^sup>L M (\w. (X w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator ((Y -`A) \ (space M))) w)" by (metis (mono_tags, lifting) assms nn_cond_expl_is_cond_exp_tmp) lemma (in finite_measure) expl_cond_exp_integrable: assumes "integrable M X" and "disc_fct Y" and "point_measurable M (space N) Y" shows "integrable M (expl_cond_expect M Y X)" proof - let ?zer = "\w. 0" let ?Xp = "\w. max (?zer w) (X w)" let ?Xn = "\w. max (?zer 0) (-X w)" have "\w. X w = ?Xp w - ?Xn w" by auto have ints: "integrable M ?Xp" "integrable M ?Xn" using integrable_max assms by auto hence "integrable M (expl_cond_expect M Y ?Xp)" using assms nn_expl_cond_exp_integrable by (metis max.cobounded1) moreover have "integrable M (expl_cond_expect M Y ?Xn)" using ints assms nn_expl_cond_exp_integrable by (metis max.cobounded1) ultimately have integr: "integrable M (\w. (expl_cond_expect M Y ?Xp) w - (expl_cond_expect M Y ?Xn) w)" by auto have "\w\ space M. (expl_cond_expect M Y ?Xp) w - (expl_cond_expect M Y ?Xn) w = (expl_cond_expect M Y X) w" proof fix w assume "w\ space M" hence "(expl_cond_expect M Y ?Xp) w - (expl_cond_expect M Y ?Xn) w = (expl_cond_expect M Y (\x. ?Xp x - ?Xn x)) w" using ints by (simp add: expl_cond_exp_diff) also have "... = expl_cond_expect M Y X w" using expl_cond_exp_cong \\w. X w = ?Xp w - ?Xn w\ by auto finally show "(expl_cond_expect M Y ?Xp) w - (expl_cond_expect M Y ?Xn) w = expl_cond_expect M Y X w" . qed thus ?thesis using integr - by (metis (mono_tags, lifting) integrable_cong) + by (metis (mono_tags, lifting) Bochner_Integration.integrable_cong) qed lemma (in finite_measure) is_cond_exp: assumes "integrable M X" and "disc_fct Y" and "point_measurable M (space N) Y" shows "\ A \ sets N. integral\<^sup>L M (\w. (X w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator ((Y -`A) \ (space M))) w)" proof - let ?zer = "\w. 0" let ?Xp = "\w. max (?zer w) (X w)" let ?Xn = "\w. max (?zer 0) (-X w)" have "\w. X w = ?Xp w - ?Xn w" by auto have ints: "integrable M ?Xp" "integrable M ?Xn" using integrable_max assms by auto hence posint: "integrable M (expl_cond_expect M Y ?Xp)" using assms nn_expl_cond_exp_integrable by (metis max.cobounded1) have negint: "integrable M (expl_cond_expect M Y ?Xn)" using ints assms nn_expl_cond_exp_integrable by (metis max.cobounded1) have eqp: "\ A \ sets N. integral\<^sup>L M (\w. (?Xp w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xp) w) * (indicator ((Y -`A) \ (space M))) w)" using nn_cond_expl_is_cond_exp[of ?Xp Y N] assms by auto have eqn: "\ A \ sets N. integral\<^sup>L M (\w. (?Xn w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xn) w) * (indicator ((Y -`A) \ (space M))) w)" using nn_cond_expl_is_cond_exp[of ?Xn Y N] assms by auto show "\ A \ sets N. integral\<^sup>L M (\w. (X w) * (indicator ((Y -`A)\ (space M)) w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator ((Y -`A) \ (space M))) w)" proof fix A assume "A\ sets N" let ?imA = "A \ (range Y)" have "countable ?imA" using assms disc_fct_def by blast have "Y -`A = Y -`?imA" by auto have yev: "Y -` (A\ range Y) \ space M \ sets M" using \A \ sets N\ assms(3) assms(2) disct_fct_point_measurable measurable_sets by (metis \Y -` A = Y -` (A \ range Y)\) let ?indA = "indicator ((Y -`(A \ range Y))\ (space M))::'a\real" have intp: "integrable M (\w. (?Xp w) * ?indA w)" proof (rule integrable_real_mult_indicator) show "Y -` (A\ range Y) \ space M \ sets M" using yev by simp show "integrable M ?Xp" using assms by simp qed have intn: "integrable M (\w. (?Xn w) * ?indA w)" proof (rule integrable_real_mult_indicator) show "Y -` (A\ range Y) \ space M \ sets M" using yev by simp show "integrable M ?Xn" using assms by simp qed have exintp: "integrable M (\w. (expl_cond_expect M Y ?Xp w) * ?indA w)" proof (rule integrable_real_mult_indicator) show "Y -` (A\ range Y) \ space M \ sets M" using yev by simp show "integrable M (expl_cond_expect M Y ?Xp)" using posint by simp qed have exintn: "integrable M (\w. (expl_cond_expect M Y ?Xn w) * ?indA w)" proof (rule integrable_real_mult_indicator) show "Y -` (A\ range Y) \ space M \ sets M" using yev by simp show "integrable M (expl_cond_expect M Y ?Xn)" using negint by simp qed have "integral\<^sup>L M (\w. X w * indicator (Y -` A \ space M) w) = integral\<^sup>L M (\w. (?Xp w - ?Xn w) * indicator (Y -` A \ space M) w)" using \\w. X w =?Xp w - ?Xn w\ by auto also have "... = integral\<^sup>L M (\w. (?Xp w * indicator (Y -` A \ space M) w) - ?Xn w * indicator (Y -` A \ space M) w)" by (simp add: left_diff_distrib) also have "... = integral\<^sup>L M (\w. (?Xp w * indicator (Y -` A \ space M) w)) - integral\<^sup>L M (\w. ?Xn w * indicator (Y -` A \ space M) w)" using \Y -` A = Y -` (A \ range Y)\ intp intn by auto also have "... = integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xp) w) * (indicator ((Y -`A) \ (space M))) w) - integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xn) w) * (indicator ((Y -`A) \ (space M))) w)" using eqp eqn by (simp add: \A \ sets N\) also have "... = integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xp) w) * (indicator ((Y -`A) \ (space M))) w - ((expl_cond_expect M Y ?Xn) w) * (indicator ((Y -`A) \ (space M))) w)" using \Y -` A = Y -` (A \ range Y)\ exintn exintp by auto also have "... = integral\<^sup>L M (\w. ((expl_cond_expect M Y ?Xp) w - (expl_cond_expect M Y ?Xn) w) * (indicator ((Y -`A) \ (space M))) w)" by (simp add: left_diff_distrib) also have "... = integral\<^sup>L M (\w. ((expl_cond_expect M Y (\x. ?Xp x - ?Xn x) w) * (indicator ((Y -`A) \ (space M))) w))" using expl_cond_exp_diff[of M ?Xp ?Xn Y] ints by (metis (mono_tags, lifting) Bochner_Integration.integral_cong) also have "... = integral\<^sup>L M (\w. ((expl_cond_expect M Y X w) * (indicator ((Y -`A) \ (space M))) w))" using \\w. X w = ?Xp w - ?Xn w\ expl_cond_exp_cong[of M X "\x. ?Xp x - ?Xn x" Y] by presburger finally show "integral\<^sup>L M (\w. X w * indicator (Y -` A \ space M) w) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X w) * (indicator ((Y -`A) \ (space M))) w))" . qed qed lemma (in finite_measure) charact_cond_exp: assumes "disc_fct Y" and "integrable M X" and "point_measurable M (space N) Y" and "Y \ space M \ space N" and "\r\ range Y\ space N. \A\ sets N. range Y\ A = {r}" shows "AE w in M. real_cond_exp M (fct_gen_subalgebra M N Y) X w = expl_cond_expect M Y X w" proof (rule sigma_finite_subalgebra.real_cond_exp_charact) have "Y\ measurable M N" by (simp add: assms(1) assms(3) disct_fct_point_measurable) have "point_measurable M (space N) Y" by (simp add: assms(3)) show "sigma_finite_subalgebra M (fct_gen_subalgebra M N Y)" unfolding sigma_finite_subalgebra_def proof show "subalgebra M (fct_gen_subalgebra M N Y)" using \Y\ measurable M N\ by (simp add: fct_gen_subalgebra_is_subalgebra) show "sigma_finite_measure (restr_to_subalg M (fct_gen_subalgebra M N Y))" unfolding sigma_finite_measure_def proof (intro exI conjI) let ?A = "{space M}" show "countable ?A" by simp show "?A \ sets (restr_to_subalg M (fct_gen_subalgebra M N Y))" by (metis empty_subsetI insert_subset sets.top space_restr_to_subalg) show "\ ?A = space (restr_to_subalg M (fct_gen_subalgebra M N Y))" by (simp add: space_restr_to_subalg) show "\a\{space M}. emeasure (restr_to_subalg M (fct_gen_subalgebra M N Y)) a \ \" by (metis \subalgebra M (fct_gen_subalgebra M N Y)\ emeasure_finite emeasure_restr_to_subalg infinity_ennreal_def sets.top singletonD subalgebra_def) qed qed show "integrable M X" using assms by simp show "expl_cond_expect M Y X \ borel_measurable (fct_gen_subalgebra M N Y)" using assms by (simp add:expl_cond_exp_borel) show "integrable M (expl_cond_expect M Y X)" using assms expl_cond_exp_integrable by blast have "\A\ sets M. integral\<^sup>L M (\w. (X w) * (indicator A w)) = set_lebesgue_integral M A X" by (metis (no_types, lifting) Bochner_Integration.integral_cong mult_commute_abs real_scaleR_def set_lebesgue_integral_def) have "\A\ sets M. integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator A w)) = set_lebesgue_integral M A (expl_cond_expect M Y X)" by (metis (no_types, lifting) Bochner_Integration.integral_cong mult_commute_abs real_scaleR_def set_lebesgue_integral_def) have "\A\ sets (fct_gen_subalgebra M N Y). integral\<^sup>L M (\w. (X w) * (indicator A w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator A w))" proof fix A assume "A \ sets (fct_gen_subalgebra M N Y)" hence "A \ {Y -` B \ space M |B. B \ sets N}" using assms by (simp add:fct_gen_subalgebra_sigma_sets) hence "\B \ sets N. A = Y -`B \ space M" by auto from this obtain B where "B\ sets N" and "A = Y -`B\ space M" by auto thus "integral\<^sup>L M (\w. (X w) * (indicator A w)) = integral\<^sup>L M (\w. ((expl_cond_expect M Y X) w) * (indicator A w))" using is_cond_exp using Bochner_Integration.integral_cong \point_measurable M (space N) Y\ assms(1) assms(2) by blast qed thus "\A. A \ sets (fct_gen_subalgebra M N Y) \ set_lebesgue_integral M A X = set_lebesgue_integral M A (expl_cond_expect M Y X)" by (smt Bochner_Integration.integral_cong Groups.mult_ac(2) real_scaleR_def set_lebesgue_integral_def) qed lemma (in finite_measure) charact_cond_exp': assumes "disc_fct Y" and "integrable M X" and "Y\ measurable M N" and "\r\ range Y\ space N. \A\ sets N. range Y\ A = {r}" shows "AE w in M. real_cond_exp M (fct_gen_subalgebra M N Y) X w = expl_cond_expect M Y X w" proof (rule charact_cond_exp) show "disc_fct Y" using assms by simp show "integrable M X" using assms by simp show "\r\range Y \ space N. \A\sets N. range Y \ A = {r}" using assms by simp show "Y\ space M \ space N" by (meson Pi_I assms(3) measurable_space) show "point_measurable M (space N) Y" using assms by (simp add: meas_single_meas) qed end diff --git a/thys/DiscretePricing/Fair_Price.thy b/thys/DiscretePricing/Fair_Price.thy --- a/thys/DiscretePricing/Fair_Price.thy +++ b/thys/DiscretePricing/Fair_Price.thy @@ -1,3523 +1,3522 @@ (* Title: Fair_Price.thy Author: Mnacho Echenim, Univ. Grenoble Alpes *) section \Fair Prices\ text \This section contains the formalization of financial notions, such as markets, price processes, portfolios, arbitrages, fair prices, etc. It also defines risk-neutral probability spaces, and proves the main result about the fair price of a derivative in a risk-neutral probability space, namely that this fair price is equal to the expectation of the discounted value of the derivative's payoff.\ theory Fair_Price imports Filtration Martingale Geometric_Random_Walk begin subsection \Preliminary results\ lemma (in prob_space) finite_borel_measurable_integrable: assumes "f\ borel_measurable M" and "finite (f`(space M))" shows "integrable M f" proof - have "simple_function M f" using assms by (simp add: simple_function_borel_measurable) moreover have "emeasure M {y \ space M. f y \ 0} \ \" by simp ultimately have "Bochner_Integration.simple_bochner_integrable M f" using Bochner_Integration.simple_bochner_integrable.simps by blast hence "has_bochner_integral M f (Bochner_Integration.simple_bochner_integral M f)" using has_bochner_integral_simple_bochner_integrable by auto thus ?thesis using integrable.simps by auto qed subsubsection \On the almost everywhere filter\ lemma AE_eq_trans[trans]: assumes "AE x in M. A x = B x" and "AE x in M. B x = C x" shows "AE x in M. A x = C x" using assms by auto abbreviation AEeq where "AEeq M X Y \ AE w in M. X w = Y w" lemma AE_add: assumes "AE w in M. f w = g w" and "AE w in M. f' w = g' w" shows "AE w in M. f w + f' w = g w + g' w" using assms by auto lemma AE_sum: assumes "finite I" and "\ i\I. AE w in M. f i w = g i w" shows "AE w in M. (\i\ I. f i w) = (\i\ I. g i w)" using assms(1) subset_refl[of I] proof (induct rule: finite_subset_induct) case empty then show ?case by simp next case (insert a F) have "AEeq M (f a) (g a)" using assms(2) insert.hyps(2) by auto have "AE w in M. (\i\ insert a F. f i w) = f a w + (\i\ F. f i w)" by (simp add: insert.hyps(1) insert.hyps(3)) also have "AE w in M. f a w + (\i\ F. f i w) = g a w + (\i\ F. f i w)" using \AEeq M (f a) (g a)\ by auto also have "AE w in M. g a w + (\i\ F. f i w) = g a w + (\i\ F. g i w)" using insert.hyps(4) by auto also have "AE w in M. g a w + (\i\ F. g i w) = (\i\ insert a F. g i w)" by (simp add: insert.hyps(1) insert.hyps(3)) finally show ?case by auto qed lemma AE_eq_cst: assumes "AE w in M. (\w. c) w = (\w. d) w" and "emeasure M (space M) \ 0" shows "c = d" proof (rule ccontr) assume "c \ d" from \AE w in M. (\w. c) w = (\w. d) w\ obtain N where Nprops: "{w\ space M. \(\w. c) w = (\w. d) w} \ N" "N\ sets M" "emeasure M N = 0" by (force elim:AE_E) have "\w\ space M. (\w. c) w \ (\w. d) w" using \c\ d\ by simp hence "{w\ space M. (\w. c) w \ (\w. d) w} = space M" by auto hence "space M\ N" using Nprops by auto thus False using \emeasure M N = 0\ assms by (meson Nprops(2) \emeasure M (space M) \ 0\ \emeasure M N = 0\ \space M \ N\ emeasure_eq_0) qed subsubsection \On conditional expectations\ lemma (in prob_space) subalgebra_sigma_finite: assumes "subalgebra M N" shows "sigma_finite_subalgebra M N" unfolding sigma_finite_subalgebra_def by (simp add: assms prob_space_axioms prob_space_imp_sigma_finite prob_space_restr_to_subalg) lemma (in prob_space) trivial_subalg_cond_expect_AE: assumes "subalgebra M N" and "sets N = {{}, space M}" and "integrable M f" shows "AE x in M. real_cond_exp M N f x = (\x. expectation f) x" proof (intro sigma_finite_subalgebra.real_cond_exp_charact) show "sigma_finite_subalgebra M N" by (simp add: assms(1) subalgebra_sigma_finite) show "integrable M f" using assms by simp show "integrable M (\x. expectation f)" by auto show "(\x. expectation f) \ borel_measurable N" by simp show "\A. A \ sets N \ set_lebesgue_integral M A f = \x\A. expectation f\M" proof - fix A assume "A \ sets N" show "set_lebesgue_integral M A f = \x\A. expectation f\M" proof (cases "A = {}") case True thus ?thesis by (simp add: set_lebesgue_integral_def) next case False hence "A = space M" using assms \A\ sets N\ by auto have "set_lebesgue_integral M A f = expectation f" using \A = space M\ by (metis (mono_tags, lifting) Bochner_Integration.integral_cong indicator_simps(1) scaleR_one set_lebesgue_integral_def) also have "... =\x\A. expectation f\M" using \A = space M\ by (auto simp add:prob_space set_lebesgue_integral_def) finally show ?thesis . qed qed qed lemma (in prob_space) triv_subalg_borel_eq: assumes "subalgebra M F" and "sets F = {{}, space M}" and "AE x in M. f x = (c::'b::{t2_space})" and "f\ borel_measurable F" shows "\x\ space M. f x = c" proof fix x assume "x\ space M" have "space M = space F" using assms by (simp add:subalgebra_def) hence "x\ space F" using \x\ space M\ by simp have "space M \ {}" by (simp add:not_empty) hence "\d. \y\ space F. f y = d" by (metis assms(1) assms(2) assms(4) subalgebra_def triv_measurable_cst) from this obtain d where "\y \space F. f y = d" by auto hence "f x = d" using \x\ space F\ by simp also have "... = c" proof (rule ccontr) assume "d\ c" from \AE x in M. f x= c\ obtain N where Nprops: "{x\ space M. \f x = c} \ N" "N\ sets M" "emeasure M N = 0" by (force elim:AE_E) have "space M \ {x\ space M. \f x = c}" using \\y \space F. f y = d\ \space M = space F\ \d\ c\ by auto hence "space M\ N" using Nprops by auto thus False using \emeasure M N = 0\ emeasure_space_1 Nprops(2) emeasure_mono by fastforce qed finally show "f x = c" . qed lemma (in prob_space) trivial_subalg_cond_expect_eq: assumes "subalgebra M N" and "sets N = {{}, space M}" and "integrable M f" shows "\x\ space M. real_cond_exp M N f x = expectation f" proof (rule triv_subalg_borel_eq) show "subalgebra M N" "sets N = {{}, space M}" using assms by auto show "real_cond_exp M N f \ borel_measurable N" by simp show "AE x in M. real_cond_exp M N f x = expectation f" by (rule trivial_subalg_cond_expect_AE, (auto simp add:assms)) qed lemma (in sigma_finite_subalgebra) real_cond_exp_cong': assumes "\w \ space M. f w = g w" and "f\ borel_measurable M" shows "AE w in M. real_cond_exp M F f w = real_cond_exp M F g w" proof (rule real_cond_exp_cong) show "AE w in M. f w = g w" using assms by simp show "f\ borel_measurable M" using assms by simp show "g\ borel_measurable M" using assms by (metis measurable_cong) qed lemma (in sigma_finite_subalgebra) real_cond_exp_bsum : fixes f::"'b \ 'a \ real" assumes [measurable]: "\i. i\I \ integrable M (f i)" shows "AE x in M. real_cond_exp M F (\x. \i\I. f i x) x = (\i\I. real_cond_exp M F (f i) x)" proof (rule real_cond_exp_charact) fix A assume [measurable]: "A \ sets F" then have A_meas [measurable]: "A \ sets M" by (meson subsetD subalg subalgebra_def) have *: "\i. i \ I \ integrable M (\x. indicator A x * f i x)" using integrable_mult_indicator[OF \A \ sets M\ assms(1)] by auto have **: "\i. i \ I \ integrable M (\x. indicator A x * real_cond_exp M F (f i) x)" using integrable_mult_indicator[OF \A \ sets M\ real_cond_exp_int(1)[OF assms(1)]] by auto have inti: "\i. i \ I \(\x. indicator A x * f i x \M) = (\x. indicator A x * real_cond_exp M F (f i) x \M)" using real_cond_exp_intg(2)[symmetric,of "indicator A" ] using "*" \A \ sets F\ assms borel_measurable_indicator by blast have "(\x\A. (\i\I. f i x)\M) = (\x. (\i\I. indicator A x * f i x)\M)" by (simp add: sum_distrib_left set_lebesgue_integral_def) also have "... = (\i\I. (\x. indicator A x * f i x \M))" using Bochner_Integration.integral_sum[of I M "\i x. indicator A x * f i x"] * by simp also have "... = (\i\I. (\x. indicator A x * real_cond_exp M F (f i) x \M))" using inti by auto also have "... = (\x. (\i\I. indicator A x * real_cond_exp M F (f i) x)\M)" by (rule Bochner_Integration.integral_sum[symmetric], simp add: **) also have "... = (\x\A. (\i\I. real_cond_exp M F (f i) x)\M)" by (simp add: sum_distrib_left set_lebesgue_integral_def) finally show "(\x\A. (\i\I. f i x)\M) = (\x\A. (\i\I. real_cond_exp M F (f i) x)\M)" by auto qed (auto simp add: assms real_cond_exp_int(1)[OF assms(1)]) subsection \Financial formalizations\ subsubsection \Markets\ definition stk_strict_subs::"'c set \ bool" where "stk_strict_subs S \ S \ UNIV" typedef ('a,'c) discrete_market = "{(s::('c set), a::'c \ (nat \ 'a \ real)). stk_strict_subs s}" unfolding stk_strict_subs_def by fastforce definition prices where "prices Mkt = (snd (Rep_discrete_market Mkt))" definition assets where "assets Mkt = UNIV" definition stocks where "stocks Mkt = (fst (Rep_discrete_market Mkt))" definition discrete_market_of where "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" lemma prices_of: shows "prices (discrete_market_of S A) = A" proof - have "stk_strict_subs (if (stk_strict_subs S) then S else {})" proof (cases "stk_strict_subs S") case True thus ?thesis by simp next case False thus ?thesis unfolding stk_strict_subs_def by simp qed hence fct: "((if (stk_strict_subs S) then S else {}), A) \ {(s, a). stk_strict_subs s}" by simp have "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" unfolding discrete_market_of_def by simp hence "Rep_discrete_market (discrete_market_of S A) = (if (stk_strict_subs S) then S else {},A)" using Abs_discrete_market_inverse[of "(if (stk_strict_subs S) then S else {}, A)"] fct by simp thus ?thesis unfolding prices_def by simp qed lemma stocks_of: assumes "UNIV \ S" shows "stocks (discrete_market_of S A) = S" proof - have "stk_strict_subs S" using assms unfolding stk_strict_subs_def by simp hence fct: "((if (stk_strict_subs S) then S else {}), A) \ {(s, a). stk_strict_subs s}" by simp have "discrete_market_of S A = Abs_discrete_market (if (stk_strict_subs S) then S else {}, A)" unfolding discrete_market_of_def by simp hence "Rep_discrete_market (discrete_market_of S A) = (if (stk_strict_subs S) then S else {},A)" using Abs_discrete_market_inverse[of "(if (stk_strict_subs S) then S else {}, A)"] fct by simp thus ?thesis unfolding stocks_def using \stk_strict_subs S\ by simp qed lemma mkt_stocks_assets: shows "stk_strict_subs (stocks Mkt)" unfolding stocks_def prices_def by (metis Rep_discrete_market mem_Collect_eq split_beta') subsubsection \Quantity processes and portfolios\ text \These are functions that assign quantities to assets; each quantity is a stochastic process. Basic operations are defined on these processes.\ paragraph \Basic operations\ definition qty_empty where "qty_empty = (\ (x::'a) (n::nat) w. 0::real)" definition qty_single where "qty_single asset qt_proc = (qty_empty(asset := qt_proc))" definition qty_sum::"('b \ nat \ 'a \ real) \ ('b \ nat \ 'a \ real) \ ('b \ nat \ 'a \ real)" where "qty_sum pf1 pf2 = (\x n w. pf1 x n w + pf2 x n w)" definition qty_mult_comp::"('b \ nat \ 'a \ real) \ (nat \ 'a \ real) \ ('b \ nat \ 'a \ real)" where "qty_mult_comp pf1 qty = (\x n w. (pf1 x n w) * (qty n w))" definition qty_rem_comp::"('b \ nat \ 'a \ real) \ 'b \ ('b \ nat \ 'a \ real)" where "qty_rem_comp pf1 x = pf1(x:=(\n w. 0))" definition qty_replace_comp where "qty_replace_comp pf1 x pf2 = qty_sum (qty_rem_comp pf1 x) (qty_mult_comp pf2 (pf1 x))" paragraph \Support sets\ text \If p x n w is different from 0, this means that this quantity is held on interval ]n-1, n].\ definition support_set::"('b \ nat \ 'a \ real) \ 'b set" where "support_set p = {x. \ n w. p x n w \ 0}" lemma qty_empty_support_set: shows "support_set qty_empty = {}" unfolding support_set_def qty_empty_def by simp lemma sum_support_set: shows "support_set (qty_sum pf1 pf2) \ (support_set pf1) \ (support_set pf2)" proof (intro subsetI, rule ccontr) fix x assume "x\ support_set (qty_sum pf1 pf2)" and "x \ support_set pf1 \ support_set pf2" note xprops = this hence "\ n w. (qty_sum pf1 pf2) x n w \ 0" by (simp add: support_set_def) from this obtain n w where "(qty_sum pf1 pf2) x n w \ 0" by auto note nwprops = this have "pf1 x n w = 0" "pf2 x n w = 0" using xprops by (auto simp add:support_set_def) hence "(qty_sum pf1 pf2) x n w = 0" unfolding qty_sum_def by simp thus False using nwprops by simp qed lemma mult_comp_support_set: shows "support_set (qty_mult_comp pf1 qty) \ (support_set pf1)" proof (intro subsetI, rule ccontr) fix x assume "x\ support_set (qty_mult_comp pf1 qty)" and "x \ support_set pf1" note xprops = this hence "\ n w. (qty_mult_comp pf1 qty) x n w \ 0" by (simp add: support_set_def) from this obtain n w where "qty_mult_comp pf1 qty x n w \ 0" by auto note nwprops = this have "pf1 x n w = 0" using xprops by (simp add:support_set_def) hence "(qty_mult_comp pf1 qty) x n w = 0" unfolding qty_mult_comp_def by simp thus False using nwprops by simp qed lemma remove_comp_support_set: shows "support_set (qty_rem_comp pf1 x) \ ((support_set pf1) - {x})" proof (intro subsetI, rule ccontr) fix y assume "y\ support_set (qty_rem_comp pf1 x)" and "y \ support_set pf1 - {x}" note xprops = this hence "y\ support_set pf1 \ y = x" by simp have "\ n w. (qty_rem_comp pf1 x) y n w \ 0" using xprops by (simp add: support_set_def) from this obtain n w where "(qty_rem_comp pf1 x) y n w \ 0" by auto note nwprops = this show False proof (cases "y\ support_set pf1") case True hence "pf1 y n w = 0" using xprops by (simp add:support_set_def) hence "(qty_rem_comp pf1 x) x n w = 0" unfolding qty_rem_comp_def by simp thus ?thesis using nwprops by (metis \pf1 y n w = 0\ fun_upd_apply qty_rem_comp_def) next case False hence "y = x" using \y\ support_set pf1 \ y = x\ by simp hence "(qty_rem_comp pf1 x) x n w = 0" unfolding qty_rem_comp_def by simp thus ?thesis using nwprops by (simp add: \y = x\) qed qed lemma replace_comp_support_set: shows "support_set (qty_replace_comp pf1 x pf2) \ (support_set pf1 - {x}) \ support_set pf2" proof - have "support_set (qty_replace_comp pf1 x pf2) \ support_set (qty_rem_comp pf1 x) \ support_set (qty_mult_comp pf2 (pf1 x))" unfolding qty_replace_comp_def by (simp add:sum_support_set) also have "... \ (support_set pf1 - {x}) \ support_set pf2" using remove_comp_support_set mult_comp_support_set by (metis sup.mono) finally show ?thesis . qed lemma single_comp_support: shows "support_set (qty_single asset qty) \ {asset}" proof fix x assume "x\ support_set (qty_single asset qty)" show "x\ {asset}" proof (rule ccontr) assume "x\ {asset}" hence "x\ asset" by auto have "\ n w. qty_single asset qty x n w \ 0" using \x\ support_set (qty_single asset qty)\ by (simp add:support_set_def) from this obtain n w where "qty_single asset qty x n w \ 0" by auto thus False using \x\asset\ by (simp add: qty_single_def qty_empty_def) qed qed lemma single_comp_nz_support: assumes "\ n w. qty n w\ 0" shows "support_set (qty_single asset qty) = {asset}" proof show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) have "asset\ support_set (qty_single asset qty)" using assms unfolding support_set_def qty_single_def by simp thus "{asset} \ support_set (qty_single asset qty)" by auto qed paragraph \Portfolios\ definition portfolio where "portfolio p \ finite (support_set p)" definition stock_portfolio :: "('a, 'b) discrete_market \ ('b \ nat \ 'a \ real) \ bool" where "stock_portfolio Mkt p \ portfolio p \ support_set p \ stocks Mkt" lemma sum_portfolio: assumes "portfolio pf1" and "portfolio pf2" shows "portfolio (qty_sum pf1 pf2)" unfolding portfolio_def proof - have "support_set (qty_sum pf1 pf2) \ (support_set pf1) \ (support_set pf2)" by (simp add: sum_support_set) thus "finite (support_set (qty_sum pf1 pf2))" using assms unfolding portfolio_def using infinite_super by auto qed lemma sum_basic_support_set: assumes "stock_portfolio Mkt pf1" and "stock_portfolio Mkt pf2" shows "stock_portfolio Mkt (qty_sum pf1 pf2)" using assms sum_support_set[of pf1 pf2] unfolding stock_portfolio_def by (metis Diff_subset_conv gfp.leq_trans subset_Un_eq sum_portfolio) lemma mult_comp_portfolio: assumes "portfolio pf1" shows "portfolio (qty_mult_comp pf1 qty)" unfolding portfolio_def proof - have "support_set (qty_mult_comp pf1 qty) \ (support_set pf1)" by (simp add: mult_comp_support_set) thus "finite (support_set (qty_mult_comp pf1 qty))" using assms unfolding portfolio_def using infinite_super by auto qed lemma mult_comp_basic_support_set: assumes "stock_portfolio Mkt pf1" shows "stock_portfolio Mkt (qty_mult_comp pf1 qty)" using assms mult_comp_support_set[of pf1] unfolding stock_portfolio_def using mult_comp_portfolio by blast lemma remove_comp_portfolio: assumes "portfolio pf1" shows "portfolio (qty_rem_comp pf1 x)" unfolding portfolio_def proof - have "support_set (qty_rem_comp pf1 x) \ (support_set pf1)" using remove_comp_support_set[of pf1 x] by blast thus "finite (support_set (qty_rem_comp pf1 x))" using assms unfolding portfolio_def using infinite_super by auto qed lemma remove_comp_basic_support_set: assumes "stock_portfolio Mkt pf1" shows "stock_portfolio Mkt (qty_mult_comp pf1 qty)" using assms mult_comp_support_set[of pf1] unfolding stock_portfolio_def using mult_comp_portfolio by blast lemma replace_comp_portfolio: assumes "portfolio pf1" and "portfolio pf2" shows "portfolio (qty_replace_comp pf1 x pf2)" unfolding portfolio_def proof - have "support_set (qty_replace_comp pf1 x pf2) \ (support_set pf1) \ (support_set pf2)" using replace_comp_support_set[of pf1 x pf2] by blast thus "finite (support_set (qty_replace_comp pf1 x pf2))" using assms unfolding portfolio_def using infinite_super by auto qed lemma replace_comp_stocks: assumes "support_set pf1 \ stocks Mkt \ {x}" and "support_set pf2 \ stocks Mkt" shows "support_set (qty_replace_comp pf1 x pf2) \ stocks Mkt" proof - have "support_set (qty_rem_comp pf1 x) \ stocks Mkt" using assms(1) remove_comp_support_set by fastforce moreover have "support_set (qty_mult_comp pf2 (pf1 x)) \ stocks Mkt" using assms mult_comp_support_set by fastforce ultimately show ?thesis unfolding qty_replace_comp_def using sum_support_set by fastforce qed lemma single_comp_portfolio: shows "portfolio (qty_single asset qty)" by (meson finite.emptyI finite.insertI finite_subset portfolio_def single_comp_support) paragraph \Value processes\ definition val_process where "val_process Mkt p = (if (\ (portfolio p)) then (\ n w. 0) else (\ n w . (sum (\x. ((prices Mkt) x n w) * (p x (Suc n) w)) (support_set p))))" lemma subset_val_process': assumes "finite A" and "support_set p \ A" shows "val_process Mkt p n w = (sum (\x. ((prices Mkt) x n w) * (p x (Suc n) w)) A)" proof - have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x n w) * (p x (Suc n) w)) = 0" by simp hence "val_process Mkt p n w = (\x\ (support_set p). ((prices Mkt) x n w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x n w) * (p x (Suc n) w))" unfolding val_process_def using \portfolio p\ by simp also have "... = (\ x\ A. ((prices Mkt) x n w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "val_process Mkt p n w = (\ x\ A. ((prices Mkt) x n w) * (p x (Suc n) w))" . qed lemma sum_val_process: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. val_process Mkt (qty_sum pf1 pf2) n w = (val_process Mkt pf1) n w + (val_process Mkt pf2) n w" proof (intro allI) fix n w have vp1: "val_process Mkt pf1 n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf1 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf1 \ support_set pf1 \ support_set pf2" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_val_process') qed have vp2: "val_process Mkt pf2 n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf2 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf2 \ support_set pf2 \ support_set pf1" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_val_process') qed have pf:"portfolio (qty_sum pf1 pf2)" using assms by (simp add:sum_portfolio) have fin:"finite (support_set pf1 \ support_set pf2)" using assms finite_Un unfolding portfolio_def by auto have "(val_process Mkt pf1) n w + (val_process Mkt pf2) n w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf1 x (Suc n) w)) + (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * (pf2 x (Suc n) w))" using vp1 vp2 by simp also have "... = (\ x\ (support_set pf1)\ (support_set pf2). (((prices Mkt) x n w) * (pf1 x (Suc n) w)) + ((prices Mkt) x n w) * (pf2 x (Suc n) w))" by (simp add: sum.distrib) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * ((pf1 x (Suc n) w) + (pf2 x (Suc n) w)))" by (simp add: distrib_left) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x n w) * ((qty_sum pf1 pf2) x (Suc n) w))" by (simp add: qty_sum_def) also have "... = (\ x\ (support_set (qty_sum pf1 pf2)). ((prices Mkt) x n w) * ((qty_sum pf1 pf2) x (Suc n) w))" using sum_support_set[of pf1 pf2] subset_val_process'[of "support_set pf1\ support_set pf2" "qty_sum pf1 pf2"] pf fin unfolding val_process_def by simp also have "... = val_process Mkt (qty_sum pf1 pf2) n w" by (metis (no_types, lifting) pf sum.cong val_process_def) finally have "(val_process Mkt pf1) n w + (val_process Mkt pf2) n w = val_process Mkt (qty_sum pf1 pf2) n w" . thus "val_process Mkt (qty_sum pf1 pf2) n w = (val_process Mkt pf1) n w + (val_process Mkt pf2) n w" .. qed lemma mult_comp_val_process: assumes "portfolio pf1" shows "\n w. val_process Mkt (qty_mult_comp pf1 qty) n w = ((val_process Mkt pf1) n w) * (qty (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_mult_comp pf1 qty)" using assms by (simp add:mult_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto have "((val_process Mkt pf1) n w) * (qty (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x n w) * (pf1 x (Suc n) w))*(qty (Suc n) w)" unfolding val_process_def using assms by simp also have "... = (\ x\ (support_set pf1). (((prices Mkt) x n w) * (pf1 x (Suc n) w) * (qty (Suc n) w)))" using sum_distrib_right by auto also have "... = (\ x\ (support_set pf1). ((prices Mkt) x n w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" unfolding qty_mult_comp_def by (simp add: mult.commute mult.left_commute) also have "... = (\ x\ (support_set (qty_mult_comp pf1 qty)). ((prices Mkt) x n w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" using mult_comp_support_set[of pf1] subset_val_process'[of "support_set pf1" "qty_mult_comp pf1 qty"] pf fin unfolding val_process_def by simp also have "... = val_process Mkt (qty_mult_comp pf1 qty) n w" by (metis (no_types, lifting) pf sum.cong val_process_def) finally have "(val_process Mkt pf1) n w * (qty (Suc n) w) = val_process Mkt (qty_mult_comp pf1 qty) n w" . thus "val_process Mkt (qty_mult_comp pf1 qty) n w = (val_process Mkt pf1) n w * (qty (Suc n) w)" .. qed lemma remove_comp_values: assumes "x \ y" shows "\n w. pf1 x n w = (qty_rem_comp pf1 y) x n w" proof (intro allI) fix n w show "pf1 x n w = (qty_rem_comp pf1 y) x n w" by (simp add: assms qty_rem_comp_def) qed lemma remove_comp_val_process: assumes "portfolio pf1" shows "\n w. val_process Mkt (qty_rem_comp pf1 y) n w = ((val_process Mkt pf1) n w) - (prices Mkt y n w)* (pf1 y (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_rem_comp pf1 y)" using assms by (simp add:remove_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto hence fin2: "finite (support_set pf1 - {y})" by simp have "((val_process Mkt pf1) n w) = (\ x\ (support_set pf1). ((prices Mkt) x n w) * (pf1 x (Suc n) w))" unfolding val_process_def using assms by simp also have "... = (\ x\ (support_set pf1 - {y}). (((prices Mkt) x n w) * (pf1 x (Suc n) w))) + (prices Mkt y n w)* (pf1 y (Suc n) w)" proof (cases "y\ support_set pf1") case True thus ?thesis by (simp add: fin sum_diff1) next case False hence "pf1 y (Suc n) w = 0" unfolding support_set_def by simp thus ?thesis by (simp add: fin sum_diff1) qed also have "... = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x n w) * ((qty_rem_comp pf1 y) x (Suc n) w)) + (prices Mkt y n w)* (pf1 y (Suc n) w)" proof - have "(\ x\ (support_set pf1 - {y}). (((prices Mkt) x n w) * (pf1 x (Suc n) w))) = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x n w) * ((qty_rem_comp pf1 y) x (Suc n) w))" proof (rule sum.cong,simp) fix x assume "x\ support_set pf1 - {y}" show "prices Mkt x n w * pf1 x (Suc n) w = prices Mkt x n w * qty_rem_comp pf1 y x (Suc n) w" using remove_comp_values by (metis DiffD2 \x \ support_set pf1 - {y}\ singletonI) qed thus ?thesis by simp qed also have "... = (val_process Mkt (qty_rem_comp pf1 y) n w) + (prices Mkt y n w)* (pf1 y (Suc n) w)" using subset_val_process'[of "support_set pf1 - {y}" "qty_rem_comp pf1 y"] fin2 by (simp add: remove_comp_support_set) finally have "(val_process Mkt pf1) n w = (val_process Mkt (qty_rem_comp pf1 y) n w) + (prices Mkt y n w)* (pf1 y (Suc n) w)" . thus "val_process Mkt (qty_rem_comp pf1 y) n w = ((val_process Mkt pf1) n w) - (prices Mkt y n w)* (pf1 y (Suc n) w)" by simp qed lemma replace_comp_val_process: assumes "\n w. prices Mkt x n w = val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" shows "\n w. val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt pf1 n w" proof (intro allI) fix n w have "val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt (qty_rem_comp pf1 x) n w + val_process Mkt (qty_mult_comp pf2 (pf1 x)) n w" unfolding qty_replace_comp_def using assms sum_val_process[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = val_process Mkt pf1 n w - (prices Mkt x n w * pf1 x (Suc n) w) + val_process Mkt pf2 n w * pf1 x (Suc n) w" by (simp add: assms(2) assms(3) mult_comp_val_process remove_comp_val_process) also have "... = val_process Mkt pf1 n w" using assms by simp finally show "val_process Mkt (qty_replace_comp pf1 x pf2) n w = val_process Mkt pf1 n w" . qed lemma qty_single_val_process: shows "val_process Mkt (qty_single asset qty) n w = prices Mkt asset n w * qty (Suc n) w" proof - have "val_process Mkt (qty_single asset qty) n w = (sum (\x. ((prices Mkt) x n w) * ((qty_single asset qty) x (Suc n) w)) {asset})" proof (rule subset_val_process') show "finite {asset}" by simp show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) qed also have "... = prices Mkt asset n w * qty (Suc n) w" unfolding qty_single_def by simp finally show ?thesis . qed subsubsection \Trading strategies\ locale disc_equity_market = triv_init_disc_filtr_prob_space + fixes Mkt::"('a,'b) discrete_market" paragraph \Discrete predictable processes\ paragraph \Trading strategy\ definition (in disc_filtr_prob_space) trading_strategy where "trading_strategy p \ portfolio p \ (\asset \ support_set p. borel_predict_stoch_proc F (p asset))" definition (in disc_filtr_prob_space) support_adapt:: "('a, 'b) discrete_market \ ('b \ nat \ 'a \ real) \ bool" where "support_adapt Mkt pf \ (\ asset \ support_set pf. borel_adapt_stoch_proc F (prices Mkt asset))" lemma (in disc_filtr_prob_space) quantity_adapted: assumes "\ asset \ support_set p. p asset (Suc n) \ borel_measurable (F n)" "\asset \ support_set p. prices Mkt asset n \ borel_measurable (F n)" shows "val_process Mkt p n \ borel_measurable (F n)" proof (cases "portfolio p") case False have "val_process Mkt p = (\ n w. 0)" unfolding val_process_def using False by simp thus "?thesis" by simp next case True hence "val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w)" unfolding val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc n) \ borel_measurable (F n)" using assms unfolding trading_strategy_def adapt_stoch_proc_def by simp moreover have "prices Mkt asset n \ borel_measurable (F n)" using \asset \ support_set p\ assms(2) unfolding support_adapt_def by (simp add: adapt_stoch_proc_def) ultimately show "(\x. prices Mkt asset n x * p asset (Suc n) x) \ borel_measurable (F n)" by simp qed ultimately show "val_process Mkt p n \ borel_measurable (F n)" by simp qed lemma (in disc_filtr_prob_space) trading_strategy_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (val_process Mkt p)" unfolding support_adapt_def proof (cases "portfolio p") case False have "val_process Mkt p = (\ n w. 0)" unfolding val_process_def using False by simp thus "borel_adapt_stoch_proc F (val_process Mkt p)" by (simp add: constant_process_borel_adapted) next case True show ?thesis unfolding adapt_stoch_proc_def proof fix n have "val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w)" unfolding val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x n w * p x (Suc n) w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc n) \ borel_measurable (F n)" using assms unfolding trading_strategy_def predict_stoch_proc_def by simp moreover have "prices Mkt asset n \ borel_measurable (F n)" using \asset \ support_set p\ assms(2) unfolding support_adapt_def by (simp add:adapt_stoch_proc_def) ultimately show "(\x. prices Mkt asset n x * p asset (Suc n) x) \ borel_measurable (F n)" by simp qed ultimately show "val_process Mkt p n \ borel_measurable (F n)" by simp qed qed lemma (in disc_equity_market) ats_val_process_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (val_process Mkt p)" unfolding support_adapt_def by (meson assms(1) assms(2) subsetCE trading_strategy_adapted) lemma (in disc_equity_market) trading_strategy_init: assumes "trading_strategy p" and "support_adapt Mkt p" shows "\c. \w \ space M. val_process Mkt p 0 w = c" using assms adapted_init ats_val_process_adapted by simp definition (in disc_equity_market) initial_value where "initial_value pf = constant_image (val_process Mkt pf 0)" lemma (in disc_equity_market) initial_valueI: assumes "trading_strategy pf" and "support_adapt Mkt pf" shows "\w\ space M. val_process Mkt pf 0 w = initial_value pf" unfolding initial_value_def proof (rule constant_imageI) show "\c. \w\space M. val_process Mkt pf 0 w = c" using trading_strategy_init assms by simp qed lemma (in disc_equity_market) inc_predict_support_trading_strat: assumes "trading_strategy pf1" shows "\ asset \ support_set pf1 \ B. borel_predict_stoch_proc F (pf1 asset)" proof fix asset assume "asset \ support_set pf1 \ B" show "borel_predict_stoch_proc F (pf1 asset)" proof (cases "asset \ support_set pf1") case True thus ?thesis using assms unfolding trading_strategy_def by simp next case False hence "\n w. pf1 asset n w = 0" unfolding support_set_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "pf1 asset 0 \ measurable (F 0) borel" using \\n w. pf1 asset n w = 0\ by (simp add: borel_measurable_const measurable_cong) next show "\n. pf1 asset (Suc n) \ borel_measurable (F n)" proof fix n have "\w. pf1 asset (Suc n) w = 0" using \\n w. pf1 asset n w = 0\ by simp have "0\ space borel" by simp thus "pf1 asset (Suc n) \ measurable (F n) borel" using measurable_const[of 0 borel "F n"] by (metis \0 \ space borel \ (\x. 0) \ borel_measurable (F n)\ \0 \ space borel\ \\n w. pf1 asset n w = 0\ measurable_cong) qed qed qed qed lemma (in disc_equity_market) inc_predict_support_trading_strat': assumes "trading_strategy pf1" and "asset \ support_set pf1\ B" shows "borel_predict_stoch_proc F (pf1 asset)" proof (cases "asset \ support_set pf1") case True thus ?thesis using assms unfolding trading_strategy_def by simp next case False hence "\n w. pf1 asset n w = 0" unfolding support_set_def by simp show ?thesis unfolding predict_stoch_proc_def proof show "pf1 asset 0 \ measurable (F 0) borel" using \\n w. pf1 asset n w = 0\ by (simp add: borel_measurable_const measurable_cong) next show "\n. pf1 asset (Suc n) \ borel_measurable (F n)" proof fix n have "\w. pf1 asset (Suc n) w = 0" using \\n w. pf1 asset n w = 0\ by simp have "0\ space borel" by simp thus "pf1 asset (Suc n) \ measurable (F n) borel" using measurable_const[of 0 borel "F n"] by (metis \0 \ space borel \ (\x. 0) \ borel_measurable (F n)\ \0 \ space borel\ \\n w. pf1 asset n w = 0\ measurable_cong) qed qed qed lemma (in disc_equity_market) inc_support_trading_strat: assumes "trading_strategy pf1" shows "\ asset \ support_set pf1 \ B. borel_adapt_stoch_proc F (pf1 asset)" using assms by (simp add: inc_predict_support_trading_strat predict_imp_adapt) lemma (in disc_equity_market) qty_empty_trading_strat: shows "trading_strategy qty_empty" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio qty_empty" by (metis fun_upd_triv qty_single_def single_comp_portfolio) show "\asset. asset \ support_set qty_empty \ borel_predict_stoch_proc F (qty_empty asset)" using qty_empty_support_set by auto qed lemma (in disc_equity_market) sum_trading_strat: assumes "trading_strategy pf1" and "trading_strategy pf2" shows "trading_strategy (qty_sum pf1 pf2)" proof - have "\ asset \ support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)" using assms by (simp add: inc_predict_support_trading_strat) have "\ asset \ support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)" using assms by (simp add: inc_predict_support_trading_strat) have "\ asset \ support_set pf1 \ support_set pf2. borel_predict_stoch_proc F ((qty_sum pf1 pf2) asset)" proof fix asset assume "asset \ support_set pf1 \ support_set pf2" show "borel_predict_stoch_proc F (qty_sum pf1 pf2 asset)" unfolding predict_stoch_proc_def qty_sum_def proof show "(\w. pf1 asset 0 w + pf2 asset 0 w) \ borel_measurable (F 0)" proof - have "(\w. pf1 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast moreover have "(\w. pf2 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast ultimately show ?thesis by simp qed next show "\n. (\w. pf1 asset (Suc n) w + pf2 asset (Suc n) w) \ borel_measurable (F n)" proof fix n have "(\w. pf1 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1 \ support_set pf2. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast moreover have "(\w. pf2 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf2 \ support_set pf1. borel_predict_stoch_proc F (pf2 asset)\ \asset \ support_set pf1 \ support_set pf2\ predict_stoch_proc_def by blast ultimately show "(\w. pf1 asset (Suc n) w + pf2 asset (Suc n) w) \ borel_measurable (F n)" by simp qed qed qed thus ?thesis unfolding trading_strategy_def using sum_support_set[of pf1 pf2] by (meson assms(1) assms(2) subsetCE sum_portfolio trading_strategy_def) qed lemma (in disc_equity_market) mult_comp_trading_strat: assumes "trading_strategy pf1" and "borel_predict_stoch_proc F qty" shows "trading_strategy (qty_mult_comp pf1 qty)" proof - have "\ asset \ support_set pf1. borel_predict_stoch_proc F (pf1 asset)" using assms unfolding trading_strategy_def by simp have "\ asset \ support_set pf1. borel_predict_stoch_proc F (qty_mult_comp pf1 qty asset)" unfolding predict_stoch_proc_def qty_mult_comp_def proof (intro ballI conjI) fix asset assume "asset \ support_set pf1" show "(\w. pf1 asset 0 w * qty 0 w) \ borel_measurable (F 0)" proof - have "(\w. pf1 asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ predict_stoch_proc_def by auto moreover have "(\w. qty 0 w) \ borel_measurable (F 0)" using assms predict_stoch_proc_def by auto ultimately show "(\w. pf1 asset 0 w * qty 0 w) \ borel_measurable (F 0)" by simp qed show "\n. (\w. pf1 asset (Suc n) w * qty (Suc n) w) \ borel_measurable (F n)" proof fix n have "(\w. pf1 asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ predict_stoch_proc_def by blast moreover have "(\w. qty (Suc n) w) \ borel_measurable (F n)" using assms predict_stoch_proc_def by blast ultimately show "(\w. pf1 asset (Suc n) w * qty (Suc n) w) \ borel_measurable (F n)" by simp qed qed thus ?thesis unfolding trading_strategy_def using mult_comp_support_set[of pf1 qty] by (meson assms(1) mult_comp_portfolio subsetCE trading_strategy_def) qed lemma (in disc_equity_market) remove_comp_trading_strat: assumes "trading_strategy pf1" shows "trading_strategy (qty_rem_comp pf1 x)" proof - have "\ asset \ support_set pf1. borel_predict_stoch_proc F (pf1 asset)" using assms unfolding trading_strategy_def by simp have "\ asset \ support_set pf1. borel_predict_stoch_proc F (qty_rem_comp pf1 x asset)" unfolding predict_stoch_proc_def qty_rem_comp_def proof (intro ballI conjI) fix asset assume "asset \ support_set pf1" show "(pf1(x := \n w. 0)) asset 0 \ borel_measurable (F 0)" proof - show "(\w. (pf1(x := \n w. 0)) asset 0 w) \ borel_measurable (F 0)" proof (cases "asset = x") case True thus ?thesis by simp next case False thus "(\w. (pf1(x := \n w. 0)) asset 0 w) \ borel_measurable (F 0)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ by (simp add: predict_stoch_proc_def) qed qed show "\n. (\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" proof fix n show "(\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" proof (cases "asset = x") case True thus ?thesis by simp next case False thus "(\w. (pf1(x := \n w. 0)) asset (Suc n) w) \ borel_measurable (F n)" using \\asset\support_set pf1. borel_predict_stoch_proc F (pf1 asset)\ \asset \ support_set pf1\ by (simp add: predict_stoch_proc_def) qed qed qed thus ?thesis unfolding trading_strategy_def using remove_comp_support_set[of pf1 x] by (metis Diff_empty assms remove_comp_portfolio subsetCE subset_Diff_insert trading_strategy_def) qed lemma (in disc_equity_market) replace_comp_trading_strat: assumes "trading_strategy pf1" and "trading_strategy pf2" shows "trading_strategy (qty_replace_comp pf1 x pf2)" unfolding qty_replace_comp_def proof (rule sum_trading_strat) show "trading_strategy (qty_rem_comp pf1 x)" using assms by (simp add: remove_comp_trading_strat) show "trading_strategy (qty_mult_comp pf2 (pf1 x))" proof (cases "x\ support_set pf1") case True hence "borel_predict_stoch_proc F (pf1 x)" using assms unfolding trading_strategy_def by simp thus ?thesis using assms by (simp add: mult_comp_trading_strat) next case False thus ?thesis proof - obtain nn :: "'c \ ('c \ nat \ 'a \ real) \ nat" and aa :: "'c \ ('c \ nat \ 'a \ real) \ 'a" where "\x0 x1. (\v2 v3. x1 x0 v2 v3 \ 0) = (x1 x0 (nn x0 x1) (aa x0 x1) \ 0)" by moura then have "\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))" by auto then show ?thesis proof - have "\f c n a. qty_mult_comp f (pf1 x) (c::'c) n a = 0" by (metis False \\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))\ mult.commute mult_zero_left qty_mult_comp_def support_set_def) then show ?thesis by (metis (no_types) \\f c. (c \ {c. \n a. f c n a \ 0} \ f c (nn c f) (aa c f) \ 0) \ (c \ {c. \n a. f c n a \ 0} \ (\n a. f c n a = 0))\ assms(2) mult_comp_portfolio support_set_def trading_strategy_def) qed qed qed qed subsubsection \Self-financing portfolios\ paragraph \Closing value process\ fun up_cl_proc where "up_cl_proc Mkt p 0 = val_process Mkt p 0" | "up_cl_proc Mkt p (Suc n) = (\w. \x\support_set p. prices Mkt x (Suc n) w * p x (Suc n) w)" definition cls_val_process where "cls_val_process Mkt p = (if (\ (portfolio p)) then (\ n w. 0) else (\ n w . up_cl_proc Mkt p n w))" lemma (in disc_filtr_prob_space) quantity_updated_borel: assumes "\n. \ asset \ support_set p. p asset (Suc n) \ borel_measurable (F n)" and "\n. \asset \ support_set p. prices Mkt asset n \ borel_measurable (F n)" shows "\n. cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "portfolio p") case False have "cls_val_process Mkt p = (\ n w. 0)" unfolding cls_val_process_def using False by simp thus "?thesis" by simp next case True show "\n. cls_val_process Mkt p n \ borel_measurable (F n)" proof fix n show "cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "n = 0") case False hence "\m. n = Suc m" using old.nat.exhaust by auto from this obtain m where "n = Suc m" by auto have "cls_val_process Mkt p (Suc m) = (\w. \x\support_set p. prices Mkt x (Suc m) w * p x (Suc m) w)" unfolding cls_val_process_def using True by simp moreover have "(\w. \x\support_set p. prices Mkt x (Suc m) w * p x (Suc m) w) \ borel_measurable (F (Suc m))" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset (Suc m) \ borel_measurable (F m)" using assms unfolding trading_strategy_def adapt_stoch_proc_def by simp hence "p asset (Suc m) \ borel_measurable (F (Suc m))" using Suc_n_not_le_n increasing_measurable_info nat_le_linear by blast moreover have "prices Mkt asset (Suc m) \ borel_measurable (F (Suc m))" using \asset \ support_set p\ assms(2) unfolding support_adapt_def adapt_stoch_proc_def by blast ultimately show "(\x. prices Mkt asset (Suc m) x * p asset (Suc m) x) \ borel_measurable (F (Suc m))" by simp qed ultimately have "cls_val_process Mkt p (Suc m) \ borel_measurable (F (Suc m))" by simp thus ?thesis using \n = Suc m\ by simp next case True thus "cls_val_process Mkt p n \ borel_measurable (F n)" by (metis (no_types, lifting) assms(1) assms(2) quantity_adapted up_cl_proc.simps(1) cls_val_process_def val_process_def) qed qed qed lemma (in disc_equity_market) cls_val_process_adapted: assumes "trading_strategy p" and "support_adapt Mkt p" shows "borel_adapt_stoch_proc F (cls_val_process Mkt p)" proof (cases "portfolio p") case False have "cls_val_process Mkt p = (\ n w. 0)" unfolding cls_val_process_def using False by simp thus "borel_adapt_stoch_proc F (cls_val_process Mkt p)" by (simp add: constant_process_borel_adapted) next case True show ?thesis unfolding adapt_stoch_proc_def proof fix n show "cls_val_process Mkt p n \ borel_measurable (F n)" proof (cases "n = 0") case True thus "cls_val_process Mkt p n \ borel_measurable (F n)" using up_cl_proc.simps(1) assms by (metis (no_types, lifting) adapt_stoch_proc_def ats_val_process_adapted cls_val_process_def val_process_def) next case False hence "\m. Suc m = n" using not0_implies_Suc by blast from this obtain m where "Suc m = n" by auto hence "cls_val_process Mkt p n = up_cl_proc Mkt p n" unfolding cls_val_process_def using True by simp also have "... = (\w. \x\support_set p. prices Mkt x n w * p x n w)" using up_cl_proc.simps(2) \Suc m = n\ by auto finally have "cls_val_process Mkt p n = (\w. \x\support_set p. prices Mkt x n w * p x n w)" . moreover have "(\w. \x\support_set p. prices Mkt x n w * p x n w) \ borel_measurable (F n)" proof (rule borel_measurable_sum) fix asset assume "asset\ support_set p" hence "p asset n \ borel_measurable (F n)" using assms unfolding trading_strategy_def predict_stoch_proc_def using Suc_n_not_le_n \Suc m = n\ increasing_measurable_info nat_le_linear by blast moreover have "prices Mkt asset n \ borel_measurable (F n)" using assms \asset \ support_set p\ unfolding support_adapt_def adapt_stoch_proc_def using stock_portfolio_def by blast ultimately show "(\x. prices Mkt asset n x * p asset n x) \ borel_measurable (F n)" by simp qed ultimately show "cls_val_process Mkt p n \ borel_measurable (F n)" by simp qed qed qed lemma subset_cls_val_process: assumes "finite A" and "support_set p \ A" shows "\n w. cls_val_process Mkt p (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) A)" proof (intro allI) fix n::nat fix w::'b have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) = 0" by simp hence "cls_val_process Mkt p (Suc n) w = (\x\ (support_set p). ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" unfolding cls_val_process_def using \portfolio p\ up_cl_proc.simps(2)[of Mkt p n] by simp also have "... = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "cls_val_process Mkt p (Suc n) w = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" . qed lemma subset_cls_val_process': assumes "finite A" and "support_set p \ A" shows "cls_val_process Mkt p (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) A)" proof - have "portfolio p" using assms unfolding portfolio_def using finite_subset by auto have "\C. (support_set p) \ C = {} \ (support_set p) \ C = A" using assms(2) by auto from this obtain C where "(support_set p) \ C = {}" and "(support_set p) \ C = A" by auto note Cprops = this have "finite C" using assms \(support_set p) \ C = A\ by auto have "\x\ C. p x (Suc n) w = 0" using Cprops(1) support_set_def by fastforce hence "(\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) = 0" by simp hence "cls_val_process Mkt p (Suc n) w = (\x\ (support_set p). ((prices Mkt) x (Suc n) w) * (p x (Suc n) w)) + (\x\ C. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" unfolding cls_val_process_def using \portfolio p\ up_cl_proc.simps(2)[of Mkt p n] by simp also have "... = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" using \portfolio p\ \finite C\ Cprops portfolio_def sum_union_disjoint' by (metis (no_types, lifting)) finally show "cls_val_process Mkt p (Suc n) w = (\ x\ A. ((prices Mkt) x (Suc n) w) * (p x (Suc n) w))" . qed lemma sum_cls_val_process_Suc: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w" proof (intro allI) fix n w have vp1: "cls_val_process Mkt pf1 (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf1 \ support_set pf1 \ support_set pf2" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (simp add: subset_cls_val_process) qed have vp2: "cls_val_process Mkt pf2 (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" proof - have "finite (support_set pf1 \ support_set pf2) \ support_set pf2 \ support_set pf2 \ support_set pf1" by (meson assms(1) assms(2) finite_Un portfolio_def sup.cobounded1) then show ?thesis by (auto simp add: subset_cls_val_process) qed have pf:"portfolio (qty_sum pf1 pf2)" using assms by (simp add:sum_portfolio) have fin:"finite (support_set pf1 \ support_set pf2)" using assms finite_Un unfolding portfolio_def by auto have "(cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w)) + (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" using vp1 vp2 by simp also have "... = (\ x\ (support_set pf1)\ (support_set pf2). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w)) + ((prices Mkt) x (Suc n) w) * (pf2 x (Suc n) w))" by (simp add: sum.distrib) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * ((pf1 x (Suc n) w) + (pf2 x (Suc n) w)))" by (simp add: distrib_left) also have "... = (\ x\ (support_set pf1)\ (support_set pf2). ((prices Mkt) x (Suc n) w) * ((qty_sum pf1 pf2) x (Suc n) w))" by (simp add: qty_sum_def) also have "... = (\ x\ (support_set (qty_sum pf1 pf2)). ((prices Mkt) x (Suc n) w) * ((qty_sum pf1 pf2) x (Suc n) w))" using sum_support_set[of pf1 pf2] subset_cls_val_process[of "support_set pf1\ support_set pf2" "qty_sum pf1 pf2"] pf fin unfolding cls_val_process_def by simp also have "... = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" by (metis (no_types, lifting) pf sum.cong up_cl_proc.simps(2) cls_val_process_def) finally have "(cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" . thus "cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w + (cls_val_process Mkt pf2) (Suc n) w" .. qed lemma sum_cls_val_process0: assumes "portfolio pf1" and "portfolio pf2" shows "\w. cls_val_process Mkt (qty_sum pf1 pf2) 0 w = (cls_val_process Mkt pf1) 0 w + (cls_val_process Mkt pf2) 0 w" unfolding cls_val_process_def by (simp add: sum_val_process assms(1) assms(2) sum_portfolio) lemma sum_cls_val_process: assumes "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_sum pf1 pf2) n w = (cls_val_process Mkt pf1) n w + (cls_val_process Mkt pf2) n w" by (metis (no_types, lifting) assms(1) assms(2) sum_cls_val_process0 sum_cls_val_process_Suc up_cl_proc.elims) lemma mult_comp_cls_val_process0: assumes "portfolio pf1" shows "\w. cls_val_process Mkt (qty_mult_comp pf1 qty) 0 w = ((cls_val_process Mkt pf1) 0 w) * (qty (Suc 0) w)" unfolding cls_val_process_def by (simp add: assms mult_comp_portfolio mult_comp_val_process) lemma mult_comp_cls_val_process_Suc: assumes "portfolio pf1" shows "\n w. cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) * (qty (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_mult_comp pf1 qty)" using assms by (simp add:mult_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto have "((cls_val_process Mkt pf1) (Suc n) w) * (qty (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))*(qty (Suc n) w)" unfolding cls_val_process_def using assms by simp also have "... = (\ x\ (support_set pf1). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w) * (qty (Suc n) w)))" using sum_distrib_right by auto also have "... = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" unfolding qty_mult_comp_def by (simp add: mult.commute mult.left_commute) also have "... = (\ x\ (support_set (qty_mult_comp pf1 qty)). ((prices Mkt) x (Suc n) w) * ((qty_mult_comp pf1 qty) x (Suc n) w))" using mult_comp_support_set[of pf1 qty] subset_cls_val_process[of "support_set pf1" "qty_mult_comp pf1 qty"] pf fin up_cl_proc.simps(2) unfolding cls_val_process_def by simp also have "... = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" by (metis (no_types, lifting) pf sum.cong cls_val_process_def up_cl_proc.simps(2)) finally have "(cls_val_process Mkt pf1) (Suc n) w * (qty (Suc n) w) = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" . thus "cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = (cls_val_process Mkt pf1) (Suc n) w * (qty (Suc n) w)" .. qed lemma remove_comp_cls_val_process0: assumes "portfolio pf1" shows "\w. cls_val_process Mkt (qty_rem_comp pf1 y) 0 w = ((cls_val_process Mkt pf1) 0 w) - (prices Mkt y 0 w)* (pf1 y (Suc 0) w)" unfolding cls_val_process_def by (simp add: assms remove_comp_portfolio remove_comp_val_process) lemma remove_comp_cls_val_process_Suc: assumes "portfolio pf1" shows "\n w. cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) - (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof (intro allI) fix n w have pf:"portfolio (qty_rem_comp pf1 y)" using assms by (simp add:remove_comp_portfolio) have fin:"finite (support_set pf1)" using assms unfolding portfolio_def by auto hence fin2: "finite (support_set pf1 - {y})" by simp have "((cls_val_process Mkt pf1) (Suc n) w) = (\ x\ (support_set pf1). ((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))" unfolding cls_val_process_def using assms by simp also have "... = (\ x\ (support_set pf1 - {y}). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof (cases "y\ support_set pf1") case True thus ?thesis by (simp add: fin sum_diff1) next case False hence "pf1 y (Suc n) w = 0" unfolding support_set_def by simp thus ?thesis by (simp add: fin sum_diff1) qed also have "... = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x (Suc n) w) * ((qty_rem_comp pf1 y) x (Suc n) w)) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" proof - have "(\ x\ (support_set pf1 - {y}). (((prices Mkt) x (Suc n) w) * (pf1 x (Suc n) w))) = (\ x\ (support_set pf1 - {y}). ((prices Mkt) x (Suc n) w) * ((qty_rem_comp pf1 y) x (Suc n) w))" proof (rule sum.cong,simp) fix x assume "x\ support_set pf1 - {y}" show "prices Mkt x (Suc n) w * pf1 x (Suc n) w = prices Mkt x (Suc n) w * qty_rem_comp pf1 y x (Suc n) w" using remove_comp_values by (metis DiffD2 \x \ support_set pf1 - {y}\ singletonI) qed thus ?thesis by simp qed also have "... = (cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" using subset_cls_val_process[of "support_set pf1 - {y}" "qty_rem_comp pf1 y"] fin2 by (simp add: remove_comp_support_set) finally have "(cls_val_process Mkt pf1) (Suc n) w = (cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w) + (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" . thus "cls_val_process Mkt (qty_rem_comp pf1 y) (Suc n) w = ((cls_val_process Mkt pf1) (Suc n) w) - (prices Mkt y (Suc n) w)* (pf1 y (Suc n) w)" by simp qed lemma replace_comp_cls_val_process0: assumes "\w. prices Mkt x 0 w = cls_val_process Mkt pf2 0 w" and "portfolio pf1" and "portfolio pf2" shows "\w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt pf1 0 w" proof fix w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt (qty_rem_comp pf1 x) 0 w + cls_val_process Mkt (qty_mult_comp pf2 (pf1 x)) 0 w" unfolding qty_replace_comp_def using assms sum_cls_val_process0[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = cls_val_process Mkt pf1 0 w - (prices Mkt x 0 w * pf1 x (Suc 0) w) + cls_val_process Mkt pf2 0 w * pf1 x (Suc 0) w" by (simp add: assms(2) assms(3) mult_comp_cls_val_process0 remove_comp_cls_val_process0) also have "... = cls_val_process Mkt pf1 0 w" using assms by simp finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) 0 w = cls_val_process Mkt pf1 0 w" . qed lemma replace_comp_cls_val_process_Suc: assumes "\n w. prices Mkt x (Suc n) w = cls_val_process Mkt pf2 (Suc n) w" and "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" proof (intro allI) fix n w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt (qty_rem_comp pf1 x) (Suc n) w + cls_val_process Mkt (qty_mult_comp pf2 (pf1 x)) (Suc n) w" unfolding qty_replace_comp_def using assms sum_cls_val_process_Suc[of "qty_rem_comp pf1 x" "qty_mult_comp pf2 (pf1 x)"] by (simp add: mult_comp_portfolio remove_comp_portfolio) also have "... = cls_val_process Mkt pf1 (Suc n) w - (prices Mkt x (Suc n) w * pf1 x (Suc n) w) + cls_val_process Mkt pf2 (Suc n) w * pf1 x (Suc n) w" by (simp add: assms(2) assms(3) mult_comp_cls_val_process_Suc remove_comp_cls_val_process_Suc) also have "... = cls_val_process Mkt pf1 (Suc n) w" using assms by simp finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" . qed lemma replace_comp_cls_val_process: assumes "\n w. prices Mkt x n w = cls_val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" shows "\n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) n w = cls_val_process Mkt pf1 n w" by (metis (no_types, lifting) assms replace_comp_cls_val_process0 replace_comp_cls_val_process_Suc up_cl_proc.elims) lemma qty_single_updated: shows "cls_val_process Mkt (qty_single asset qty) (Suc n) w = prices Mkt asset (Suc n) w * qty (Suc n) w" proof - have "cls_val_process Mkt (qty_single asset qty) (Suc n) w = (sum (\x. ((prices Mkt) x (Suc n) w) * ((qty_single asset qty) x (Suc n) w)) {asset})" proof (rule subset_cls_val_process') show "finite {asset}" by simp show "support_set (qty_single asset qty) \ {asset}" by (simp add: single_comp_support) qed also have "... = prices Mkt asset (Suc n) w * qty (Suc n) w" unfolding qty_single_def by simp finally show ?thesis . qed paragraph \Self-financing\ definition self_financing where "self_financing Mkt p \ (\n. val_process Mkt p (Suc n) = cls_val_process Mkt p (Suc n))" lemma self_financingE: assumes "self_financing Mkt p" shows "\n. val_process Mkt p n = cls_val_process Mkt p n" proof fix n show "val_process Mkt p n = cls_val_process Mkt p n" proof (cases "n = 0") case False thus ?thesis using assms unfolding self_financing_def by (metis up_cl_proc.elims) next case True thus ?thesis by (simp add: cls_val_process_def val_process_def) qed qed lemma static_portfolio_self_financing: assumes "\ x \ support_set p. (\w i. p x i w = p x (Suc i) w)" shows "self_financing Mkt p" unfolding self_financing_def proof (intro allI impI) fix n show "val_process Mkt p (Suc n) = cls_val_process Mkt p (Suc n)" proof (cases "portfolio p") case False thus ?thesis unfolding val_process_def cls_val_process_def by simp next case True have "\w. (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = cls_val_process Mkt p (Suc n) w" proof fix w show "(\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = cls_val_process Mkt p (Suc n) w" proof - have "(\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w) = (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc n) w)" proof (rule sum.cong, simp) fix x assume "x\ support_set p" hence "p x (Suc n) w = p x (Suc (Suc n)) w" using assms by blast thus "prices Mkt x (Suc n) w * p x (Suc (Suc n)) w = prices Mkt x (Suc n) w * p x (Suc n) w" by simp qed also have "... = cls_val_process Mkt p (Suc n) w" using up_cl_proc.simps(2)[of Mkt p n] by (metis True cls_val_process_def) finally show ?thesis . qed qed moreover have "\w. val_process Mkt p (Suc n) w = (\x\ support_set p. prices Mkt x (Suc n) w * p x (Suc (Suc n)) w)" unfolding val_process_def using True by simp ultimately show ?thesis by auto qed qed lemma sum_self_financing: assumes "portfolio pf1" and "portfolio pf2" and "self_financing Mkt pf1" and "self_financing Mkt pf2" shows "self_financing Mkt (qty_sum pf1 pf2)" proof - have "\ n w. val_process Mkt (qty_sum pf1 pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" proof (intro allI) fix n w have "val_process Mkt (qty_sum pf1 pf2) (Suc n) w = val_process Mkt pf1 (Suc n) w + val_process Mkt pf2 (Suc n) w" using assms by (simp add:sum_val_process) also have "... = cls_val_process Mkt pf1 (Suc n) w + val_process Mkt pf2 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt pf1 (Suc n) w + cls_val_process Mkt pf2 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" using assms by (simp add: sum_cls_val_process) finally show "val_process Mkt (qty_sum pf1 pf2) (Suc n) w = cls_val_process Mkt (qty_sum pf1 pf2) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed lemma mult_time_constant_self_financing: assumes "portfolio pf1" and "self_financing Mkt pf1" and "\n w. qty n w = qty (Suc n) w" shows "self_financing Mkt (qty_mult_comp pf1 qty)" proof - have "\ n w. val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" proof (intro allI) fix n w have "val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = val_process Mkt pf1 (Suc n) w * qty (Suc n) w" using assms by (simp add:mult_comp_val_process) also have "... = cls_val_process Mkt pf1 (Suc n) w * qty (Suc n) w" using assms unfolding self_financing_def by simp also have "... = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" using assms by (auto simp add: mult_comp_cls_val_process_Suc) finally show "val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w = cls_val_process Mkt (qty_mult_comp pf1 qty) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed lemma replace_comp_self_financing: assumes "\n w. prices Mkt x n w = cls_val_process Mkt pf2 n w" and "portfolio pf1" and "portfolio pf2" and "self_financing Mkt pf1" and "self_financing Mkt pf2" shows "self_financing Mkt (qty_replace_comp pf1 x pf2)" proof - have sfeq: "\n w. prices Mkt x n w = val_process Mkt pf2 n w" using assms by (simp add: self_financingE) have "\ n w. cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" proof (intro allI) fix n w have "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = cls_val_process Mkt pf1 (Suc n) w" using assms by (simp add:replace_comp_cls_val_process) also have "... = val_process Mkt pf1 (Suc n) w" using assms unfolding self_financing_def by simp also have "... = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" using assms sfeq by (simp add: replace_comp_val_process self_financing_def) finally show "cls_val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w = val_process Mkt (qty_replace_comp pf1 x pf2) (Suc n) w" . qed thus ?thesis unfolding self_financing_def by auto qed paragraph \Make a portfolio self-financing\ fun remaining_qty where init: "remaining_qty Mkt v pf asset 0 = (\w. 0)" | first: "remaining_qty Mkt v pf asset (Suc 0) = (\w. (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))" | step: "remaining_qty Mkt v pf asset (Suc (Suc n)) = (\w. (remaining_qty Mkt v pf asset (Suc n) w) + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" lemma (in disc_equity_market) remaining_qty_predict': assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F n)" proof (induct n) case 0 have "(\w. (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w))\ borel_measurable (F 0)" proof (rule borel_measurable_divide) have "val_process Mkt pf 0 \ borel_measurable (F 0)" using assms ats_val_process_adapted by (simp add:adapt_stoch_proc_def) thus "(\x. v - val_process Mkt pf 0 x) \ borel_measurable (F 0)" by simp show "prices Mkt asset 0 \ borel_measurable (F 0)" using assms unfolding adapt_stoch_proc_def by simp qed thus ?case by simp next case (Suc n) have "(\w. (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/ (prices Mkt asset (Suc n) w)) \ borel_measurable (F (Suc n))" proof (rule borel_measurable_divide) show "(\w. (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" proof (rule borel_measurable_diff) show "(\w. (cls_val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" using assms cls_val_process_adapted unfolding adapt_stoch_proc_def by auto show "(\w. (val_process Mkt pf (Suc n) w)) \ borel_measurable (F (Suc n))" using assms ats_val_process_adapted by (simp add:adapt_stoch_proc_def) qed show "prices Mkt asset (Suc n) \ borel_measurable (F (Suc n))" using assms unfolding adapt_stoch_proc_def by simp qed moreover have "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F (Suc n))" using Suc Suc_n_not_le_n increasing_measurable_info nat_le_linear by blast ultimately show ?case using Suc remaining_qty.simps(3)[of Mkt v pf asset n] by simp qed lemma (in disc_equity_market) remaining_qty_predict: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "borel_predict_stoch_proc F (remaining_qty Mkt v pf asset)" unfolding predict_stoch_proc_def proof (intro conjI allI) show "remaining_qty Mkt v pf asset 0 \ borel_measurable (F 0)" using init by simp fix n show "remaining_qty Mkt v pf asset (Suc n) \ borel_measurable (F n)" using assms by (simp add: remaining_qty_predict') qed lemma (in disc_equity_market) remaining_qty_adapt: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "remaining_qty Mkt v pf asset n \ borel_measurable (F n)" using adapt_stoch_proc_def assms(1) assms(2) predict_imp_adapt remaining_qty_predict by (metis assms(3)) lemma (in disc_equity_market) remaining_qty_adapted: assumes "borel_adapt_stoch_proc F (prices Mkt asset)" and "trading_strategy pf" and "support_adapt Mkt pf" shows "borel_adapt_stoch_proc F (remaining_qty Mkt v pf asset)" using assms unfolding adapt_stoch_proc_def using assms remaining_qty_adapt by blast definition self_finance where "self_finance Mkt v pf (asset::'a) = qty_sum pf (qty_single asset (remaining_qty Mkt v pf asset))" lemma self_finance_portfolio: assumes "portfolio pf" shows "portfolio (self_finance Mkt v pf asset)" unfolding self_finance_def by (simp add: assms single_comp_portfolio sum_portfolio) lemma self_finance_init: assumes "\w. prices Mkt asset 0 w \ 0" and "portfolio pf" shows "val_process Mkt (self_finance Mkt v pf asset) 0 w = v" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "val_process Mkt (self_finance Mkt v pf asset) 0 w = val_process Mkt pf 0 w + val_process Mkt scp 0 w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = val_process Mkt pf 0 w + (sum (\x. ((prices Mkt) x 0 w) * (scp x (Suc 0) w)) {asset})" using subset_val_process'[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * scp asset (Suc 0) w" by auto also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * (remaining_qty Mkt v pf asset) (Suc 0) w" unfolding scp_def qty_single_def by simp also have "... = val_process Mkt pf 0 w + prices Mkt asset 0 w * (v - val_process Mkt pf 0 w)/(prices Mkt asset 0 w)" by simp also have "... = val_process Mkt pf 0 w + (v - val_process Mkt pf 0 w)" using assms by simp also have "... = v" by simp finally show ?thesis . qed lemma self_finance_succ: assumes "prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + cls_val_process Mkt pf (Suc n) w" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = val_process Mkt pf (Suc n) w + val_process Mkt scp (Suc n) w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = val_process Mkt pf (Suc n) w + (sum (\x. ((prices Mkt) x (Suc n) w) * (scp x (Suc (Suc n)) w)) {asset})" using subset_val_process'[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * scp asset (Suc (Suc n)) w" by auto also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc (Suc n)) w" unfolding scp_def qty_single_def by simp also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset (Suc n) w + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w))" by simp also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + prices Mkt asset (Suc n) w * (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)/(prices Mkt asset (Suc n) w)" by (simp add: distrib_left) also have "... = val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + (cls_val_process Mkt pf (Suc n) w - val_process Mkt pf (Suc n) w)" using assms by simp also have "... = prices Mkt asset (Suc n) w * remaining_qty Mkt v pf asset (Suc n) w + cls_val_process Mkt pf (Suc n) w" by simp finally show ?thesis . qed lemma self_finance_updated: assumes "prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc n) w" proof - define scp where "scp = qty_single asset (remaining_qty Mkt v pf asset)" have "cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt pf (Suc n) w + cls_val_process Mkt scp (Suc n) w" unfolding scp_def using assms single_comp_portfolio[of asset] sum_cls_val_process[of pf "qty_single asset (remaining_qty Mkt v pf asset)" Mkt] by (simp add: \\qty. portfolio (qty_single asset qty)\ self_finance_def) also have "... = cls_val_process Mkt pf (Suc n) w + (sum (\x. ((prices Mkt) x (Suc n) w) * (scp x (Suc n) w)) {asset})" using subset_cls_val_process[of "{asset}" scp] unfolding scp_def by (auto simp add: single_comp_support) also have "... = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * scp asset (Suc n) w" by auto also have "... = cls_val_process Mkt pf (Suc n) w + prices Mkt asset (Suc n) w * (remaining_qty Mkt v pf asset) (Suc n) w" unfolding scp_def qty_single_def by simp finally show ?thesis . qed lemma self_finance_charact: assumes "\ n w. prices Mkt asset (Suc n) w \ 0" and "portfolio pf" shows "self_financing Mkt (self_finance Mkt v pf asset)" proof- have "\n w. val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w" proof (intro allI) fix n w show "val_process Mkt (self_finance Mkt v pf asset) (Suc n) w = cls_val_process Mkt (self_finance Mkt v pf asset) (Suc n) w" using assms self_finance_succ[of Mkt asset] by (simp add: self_finance_updated) qed thus ?thesis unfolding self_financing_def by auto qed subsubsection \Replicating portfolios\ definition (in disc_filtr_prob_space) price_structure::"('a \ real) \ nat \ real \ (nat \ 'a \ real) \ bool" where "price_structure pyf T \ pr \ ((\ w\ space M. pr 0 w = \) \ (AE w in M. pr T w = pyf w) \ (pr T \ borel_measurable (F T)))" lemma (in disc_filtr_prob_space) price_structure_init: assumes "price_structure pyf T \ pr" shows "\ w\ space M. pr 0 w = \" using assms unfolding price_structure_def by simp lemma (in disc_filtr_prob_space) price_structure_borel_measurable: assumes "price_structure pyf T \ pr" shows "pr T \ borel_measurable (F T)" using assms unfolding price_structure_def by simp lemma (in disc_filtr_prob_space) price_structure_maturity: assumes "price_structure pyf T \ pr" shows "AE w in M. pr T w = pyf w" using assms unfolding price_structure_def by simp definition (in disc_equity_market) replicating_portfolio where "replicating_portfolio pf der matur \ (stock_portfolio Mkt pf) \ (trading_strategy pf) \ (self_financing Mkt pf) \ (AE w in M. cls_val_process Mkt pf matur w = der w)" definition (in disc_equity_market) is_attainable where "is_attainable der matur \ (\ pf. replicating_portfolio pf der matur)" lemma (in disc_equity_market) replicating_price_process: assumes "replicating_portfolio pf der matur" and "support_adapt Mkt pf" shows "price_structure der matur (initial_value pf) (cls_val_process Mkt pf)" unfolding price_structure_def proof (intro conjI) show "AE w in M. cls_val_process Mkt pf matur w = der w" using assms unfolding replicating_portfolio_def by simp show "\w\space M. cls_val_process Mkt pf 0 w = initial_value pf" proof fix w assume "w\ space M" thus "cls_val_process Mkt pf 0 w = initial_value pf" unfolding initial_value_def using constant_imageI[of "cls_val_process Mkt pf 0"] trading_strategy_init[of pf] assms replicating_portfolio_def [of pf der matur] by (simp add: stock_portfolio_def cls_val_process_def) qed show "cls_val_process Mkt pf matur \ borel_measurable (F matur)" using assms unfolding replicating_portfolio_def using ats_val_process_adapted[of pf] cls_val_process_adapted by (simp add:adapt_stoch_proc_def) qed subsubsection \Arbitrages\ definition (in disc_filtr_prob_space) arbitrage_process where "arbitrage_process Mkt p \ (\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" lemma (in disc_filtr_prob_space) arbitrage_processE: assumes "arbitrage_process Mkt p" shows "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" using assms disc_filtr_prob_space.arbitrage_process_def disc_filtr_prob_space_axioms self_financingE by fastforce lemma (in disc_filtr_prob_space) arbitrage_processI: assumes "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w \ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" shows "arbitrage_process Mkt p" unfolding arbitrage_process_def using assms by (simp add: self_financingE cls_val_process_def) definition (in disc_filtr_prob_space) viable_market where "viable_market Mkt \ (\p. stock_portfolio Mkt p \ \ arbitrage_process Mkt p)" lemma (in disc_filtr_prob_space) arbitrage_val_process: assumes "arbitrage_process Mkt pf1" and "self_financing Mkt pf2" and "trading_strategy pf2" and "\ n w. cls_val_process Mkt pf1 n w = cls_val_process Mkt pf2 n w" shows "arbitrage_process Mkt pf2" proof - have "(\ m. (self_financing Mkt pf1) \ (trading_strategy pf1) \ (\w \ space M. cls_val_process Mkt pf1 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt pf1 m w) \ 0 < \

(w in M. cls_val_process Mkt pf1 m w > 0))" using assms arbitrage_processE[of Mkt pf1] by simp from this obtain m where "(self_financing Mkt pf1)" and "(trading_strategy pf1)" and "(\w \ space M. cls_val_process Mkt pf1 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt pf1 m w)" "0 < \

(w in M. cls_val_process Mkt pf1 m w > 0)" by auto have ae_eq:"\w\ space M. (cls_val_process Mkt pf1 0 w = 0) = (cls_val_process Mkt pf2 0 w = 0)" proof fix w assume "w\ space M" show "(cls_val_process Mkt pf1 0 w = 0) = (cls_val_process Mkt pf2 0 w = 0) " using assms by simp qed have ae_geq:"almost_everywhere M (\w. cls_val_process Mkt pf1 m w \ 0) = almost_everywhere M (\w. cls_val_process Mkt pf2 m w \ 0)" proof (rule AE_cong) fix w assume "w\ space M" show "(cls_val_process Mkt pf1 m w \ 0) = (cls_val_process Mkt pf2 m w \ 0) " using assms by simp qed have "self_financing Mkt pf2" using assms by simp moreover have "trading_strategy pf2" using assms by simp moreover have "(\w \ space M. cls_val_process Mkt pf2 0 w = 0)" using \(\w \ space M. cls_val_process Mkt pf1 0 w = 0)\ ae_eq by simp moreover have "AE w in M. 0 \ cls_val_process Mkt pf2 m w" using \AE w in M. 0 \ cls_val_process Mkt pf1 m w\ ae_geq by simp moreover have "0 < prob {w \ space M. 0 < cls_val_process Mkt pf2 m w}" proof - have "{w \ space M. 0 < cls_val_process Mkt pf2 m w} = {w \ space M. 0 < cls_val_process Mkt pf1 m w}" by (simp add: assms(4)) thus ?thesis by (simp add: \0 < prob {w \ space M. 0 < cls_val_process Mkt pf1 m w}\) qed ultimately show ?thesis using arbitrage_processI by blast qed definition coincides_on where "coincides_on Mkt Mkt2 A \ (stocks Mkt = stocks Mkt2 \ (\x. x\ A \ prices Mkt x = prices Mkt2 x))" lemma coincides_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. val_process Mkt pf n w = val_process Mkt2 pf n w" proof (intro allI) fix n w show "val_process Mkt pf n w = val_process Mkt2 pf n w" proof (cases "portfolio pf") case False thus ?thesis unfolding val_process_def by simp next case True hence "val_process Mkt pf n w = (\x\ support_set pf. prices Mkt x n w * pf x (Suc n) w)" using assms unfolding val_process_def by simp also have "... = (\x\ support_set pf. prices Mkt2 x n w * pf x (Suc n) w)" proof (rule sum.cong, simp) fix y assume "y\ support_set pf" hence "y\ A" using assms unfolding stock_portfolio_def by auto hence "prices Mkt y n w = prices Mkt2 y n w" using assms unfolding coincides_on_def by auto thus "prices Mkt y n w * pf y (Suc n) w = prices Mkt2 y n w * pf y (Suc n) w" by simp qed also have "... = val_process Mkt2 pf n w" by (metis (mono_tags, lifting) calculation val_process_def) finally show "val_process Mkt pf n w = val_process Mkt2 pf n w" . qed qed lemma coincides_cls_val_process': assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" proof (intro allI) fix n w show "cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" proof (cases "portfolio pf") case False thus ?thesis unfolding cls_val_process_def by simp next case True hence "cls_val_process Mkt pf (Suc n) w = (\x\ support_set pf. prices Mkt x (Suc n) w * pf x (Suc n) w)" using assms unfolding cls_val_process_def by simp also have "... = (\x\ support_set pf. prices Mkt2 x (Suc n) w * pf x (Suc n) w)" proof (rule sum.cong, simp) fix y assume "y\ support_set pf" hence "y\ A" using assms unfolding stock_portfolio_def by auto hence "prices Mkt y (Suc n) w = prices Mkt2 y (Suc n) w" using assms unfolding coincides_on_def by auto thus "prices Mkt y (Suc n) w * pf y (Suc n) w = prices Mkt2 y (Suc n) w * pf y (Suc n) w" by simp qed also have "... = cls_val_process Mkt2 pf (Suc n) w" by (metis (mono_tags, lifting) True up_cl_proc.simps(2) cls_val_process_def) finally show "cls_val_process Mkt pf (Suc n) w = cls_val_process Mkt2 pf (Suc n) w" . qed qed lemma coincides_cls_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set pf \ A" shows "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (intro allI) fix n w show "cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (cases "portfolio pf") case False thus ?thesis unfolding cls_val_process_def by simp next case True show "cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" proof (induct n) case 0 with assms show ?case by (simp add: cls_val_process_def coincides_val_process) next case Suc thus ?case using coincides_cls_val_process' assms by blast qed qed qed lemma (in disc_filtr_prob_space) coincides_on_self_financing: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "self_financing Mkt p" shows "self_financing Mkt2 p" proof - have "\ n w. val_process Mkt2 p (Suc n) w = cls_val_process Mkt2 p (Suc n) w" proof (intro allI) fix n w have "val_process Mkt2 p (Suc n) w = val_process Mkt p (Suc n) w" using assms(1) assms(2) coincides_val_process stock_portfolio_def by fastforce also have "... = cls_val_process Mkt p (Suc n) w" by (metis \self_financing Mkt p\ self_financing_def) also have "... = cls_val_process Mkt2 p (Suc n) w" using assms(1) assms(2) coincides_cls_val_process stock_portfolio_def by blast finally show "val_process Mkt2 p (Suc n) w = cls_val_process Mkt2 p (Suc n) w" . qed thus "self_financing Mkt2 p" unfolding self_financing_def by auto qed lemma (in disc_filtr_prob_space) coincides_on_arbitrage: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "arbitrage_process Mkt p" shows "arbitrage_process Mkt2 p" proof - have "(\ m. (self_financing Mkt p) \ (trading_strategy p) \ (\w\ space M. cls_val_process Mkt p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt p m w) \ 0 < \

(w in M. cls_val_process Mkt p m w > 0))" using assms using arbitrage_processE by simp from this obtain m where "(self_financing Mkt p)" and "(trading_strategy p)" and "(\w\ space M. cls_val_process Mkt p 0 w = 0)" and "(AE w in M. 0 \ cls_val_process Mkt p m w)" "0 < \

(w in M. cls_val_process Mkt p m w > 0)" by auto have ae_eq:"\w\ space M. (cls_val_process Mkt2 p 0 w = 0) = (cls_val_process Mkt p 0 w = 0)" proof fix w assume "w\ space M" show "(cls_val_process Mkt2 p 0 w = 0) = (cls_val_process Mkt p 0 w = 0) " using assms coincides_cls_val_process by (metis) qed have ae_geq:"almost_everywhere M (\w. cls_val_process Mkt2 p m w \ 0) = almost_everywhere M (\w. cls_val_process Mkt p m w \ 0)" proof (rule AE_cong) fix w assume "w\ space M" show "(cls_val_process Mkt2 p m w \ 0) = (cls_val_process Mkt p m w \ 0) " using assms coincides_cls_val_process by (metis) qed have "self_financing Mkt2 p" using assms coincides_on_self_financing using \self_financing Mkt p\ by blast moreover have "trading_strategy p" using \trading_strategy p\ . moreover have "(\w\ space M. cls_val_process Mkt2 p 0 w = 0)" using \(\w\ space M. cls_val_process Mkt p 0 w = 0)\ ae_eq by simp moreover have "AE w in M. 0 \ cls_val_process Mkt2 p m w" using \AE w in M. 0 \ cls_val_process Mkt p m w\ ae_geq by simp moreover have "0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" proof - have "{w \ space M. 0 < cls_val_process Mkt2 p m w} = {w \ space M. 0 < cls_val_process Mkt p m w}" by (metis (no_types, lifting) assms(1) assms(2) coincides_cls_val_process) thus ?thesis by (simp add: \0 < prob {w \ space M. 0 < cls_val_process Mkt p m w}\) qed ultimately show ?thesis using arbitrage_processI by blast qed lemma (in disc_filtr_prob_space) coincides_on_stocks_viable: assumes "coincides_on Mkt Mkt2 (stocks Mkt)" and "viable_market Mkt" shows "viable_market Mkt2" using coincides_on_arbitrage by (metis (mono_tags, opaque_lifting) assms(1) assms(2) coincides_on_def stock_portfolio_def viable_market_def) lemma coincides_stocks_val_process: assumes "stock_portfolio Mkt pf" and "coincides_on Mkt Mkt2 (stocks Mkt)" shows "\n w. val_process Mkt pf n w = val_process Mkt2 pf n w" using assms unfolding stock_portfolio_def by (simp add: coincides_val_process) lemma coincides_stocks_cls_val_process: assumes "stock_portfolio Mkt pf" and "coincides_on Mkt Mkt2 (stocks Mkt)" shows "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using assms unfolding stock_portfolio_def by (simp add: coincides_cls_val_process) lemma (in disc_filtr_prob_space) coincides_on_adapted_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "borel_adapt_stoch_proc F (val_process Mkt p)" shows "borel_adapt_stoch_proc F (val_process Mkt2 p)" unfolding adapt_stoch_proc_def proof fix n have "val_process Mkt p n \ borel_measurable (F n)" using assms unfolding adapt_stoch_proc_def by simp moreover have "\w. val_process Mkt p n w = val_process Mkt2 p n w" using assms coincides_val_process[of Mkt Mkt2 A] by auto thus "val_process Mkt2 p n \ borel_measurable (F n)" using calculation by presburger qed lemma (in disc_filtr_prob_space) coincides_on_adapted_cls_val_process: assumes "coincides_on Mkt Mkt2 A" and "support_set p \ A" and "borel_adapt_stoch_proc F (cls_val_process Mkt p)" shows "borel_adapt_stoch_proc F (cls_val_process Mkt2 p)" unfolding adapt_stoch_proc_def proof fix n have "cls_val_process Mkt p n \ borel_measurable (F n)" using assms unfolding adapt_stoch_proc_def by simp moreover have "\w. cls_val_process Mkt p n w = cls_val_process Mkt2 p n w" using assms coincides_cls_val_process[of Mkt Mkt2 A] by auto thus "cls_val_process Mkt2 p n \ borel_measurable (F n)" using calculation by presburger qed subsubsection \Fair prices\ definition (in disc_filtr_prob_space) fair_price where "fair_price Mkt \ pyf matur \ (\ pr. price_structure pyf matur \ pr \ (\ x Mkt2 p. (x\ stocks Mkt \ ((coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" lemma (in disc_filtr_prob_space) fair_priceI: assumes "fair_price Mkt \ pyf matur" shows "(\ pr. price_structure pyf matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using assms unfolding fair_price_def by simp paragraph \Existence when replicating portfolio\ lemma (in disc_equity_market) replicating_fair_price: assumes "viable_market Mkt" and "replicating_portfolio pf der matur" and "support_adapt Mkt pf" shows "fair_price Mkt (initial_value pf) der matur" proof (rule ccontr) define \ where "\ = (initial_value pf)" assume "\ fair_price Mkt \ der matur" hence imps: "\pr. (price_structure der matur \ pr) \ (\ x Mkt2 p. (x\ stocks Mkt \ (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p))" unfolding fair_price_def by simp have "(price_structure der matur \ (cls_val_process Mkt pf))" unfolding \_def using replicating_price_process assms by simp hence "\ x Mkt2 p. (x\ stocks Mkt \ (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = (cls_val_process Mkt pf)) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p)" using imps by simp from this obtain x Mkt2 p where "x\ stocks Mkt" and "coincides_on Mkt Mkt2 (stocks Mkt)" and "prices Mkt2 x = cls_val_process Mkt pf" and "portfolio p" and "support_set p\ stocks Mkt \ {x}" and "arbitrage_process Mkt2 p" by auto have "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using coincides_stocks_cls_val_process[of Mkt pf Mkt2] assms \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by simp have "\m. self_financing Mkt2 p \ trading_strategy p \ (AE w in M. cls_val_process Mkt2 p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt2 p m w) \ 0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" using \arbitrage_process Mkt2 p\ using arbitrage_processE[of Mkt2] by simp from this obtain m where "self_financing Mkt2 p" "trading_strategy p \ (AE w in M. cls_val_process Mkt2 p 0 w = 0) \ (AE w in M. 0 \ cls_val_process Mkt2 p m w) \ 0 < prob {w \ space M. 0 < cls_val_process Mkt2 p m w}" by auto note mprop = this define eq_stock where "eq_stock = qty_replace_comp p x pf" have "\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w" using assms unfolding replicating_portfolio_def using coincides_cls_val_process using \\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w\ by blast hence prx: "\n w. prices Mkt2 x n w = cls_val_process Mkt2 pf n w" using \prices Mkt2 x = cls_val_process Mkt pf\ by simp have "stock_portfolio Mkt2 eq_stock" by (metis (no_types, lifting) \coincides_on Mkt Mkt2 (stocks Mkt)\ \portfolio p\ \support_set p \ stocks Mkt \ {x}\ assms(2) coincides_on_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms eq_stock_def replace_comp_portfolio replace_comp_stocks stock_portfolio_def) moreover have "arbitrage_process Mkt2 eq_stock" proof (rule arbitrage_val_process) show "arbitrage_process Mkt2 p" using \arbitrage_process Mkt2 p\ . show vp: "\n w. cls_val_process Mkt2 p n w = cls_val_process Mkt2 eq_stock n w" using replace_comp_cls_val_process \prices Mkt2 x = cls_val_process Mkt pf\ unfolding eq_stock_def by (metis (no_types, lifting) \\n w. cls_val_process Mkt pf n w = cls_val_process Mkt2 pf n w\ \portfolio p\ assms(2) replicating_portfolio_def stock_portfolio_def) show "trading_strategy eq_stock" by (metis \arbitrage_process Mkt2 p\ arbitrage_process_def assms(2) eq_stock_def replace_comp_trading_strat replicating_portfolio_def) show "self_financing Mkt2 eq_stock" unfolding eq_stock_def proof (rule replace_comp_self_financing) show "portfolio pf" using assms unfolding replicating_portfolio_def stock_portfolio_def by simp show "portfolio p" using \portfolio p\ . show "\n w. prices Mkt2 x n w = cls_val_process Mkt2 pf n w" using prx . show "self_financing Mkt2 p" using \self_financing Mkt2 p\ . show "self_financing Mkt2 pf" using coincides_on_self_financing[of Mkt Mkt2 "stocks Mkt" pf] \coincides_on Mkt Mkt2 (stocks Mkt)\ assms(2) (*disc_equity_market.replicating_portfolio_def disc_equity_market_axioms*) unfolding stock_portfolio_def replicating_portfolio_def by auto qed qed moreover have "viable_market Mkt2" using assms coincides_on_stocks_viable[of Mkt Mkt2] by (simp add: \coincides_on Mkt Mkt2 (stocks Mkt)\) ultimately show False unfolding viable_market_def by simp qed paragraph \Uniqueness when replicating portfolio\ text \The proof of uniqueness requires the existence of a stock that always takes strictly positive values.\ locale disc_market_pos_stock = disc_equity_market + fixes pos_stock assumes in_stock: "pos_stock \ stocks Mkt" and positive: "\ n w. prices Mkt pos_stock n w > 0" and readable: "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" lemma (in disc_market_pos_stock) pos_stock_borel_adapted: shows "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using assets_def readable in_stock by auto definition static_quantities where "static_quantities p \ (\asset \ support_set p. \c::real. p asset = (\ n w. c))" lemma (in disc_filtr_prob_space) static_quantities_trading_strat: assumes "static_quantities p" and "finite (support_set p)" shows "trading_strategy p" unfolding trading_strategy_def proof (intro conjI ballI) show "portfolio p" using assms unfolding portfolio_def by simp fix asset assume "asset \ support_set p" hence "\c. p asset = (\ n w. c)" using assms unfolding static_quantities_def by simp then obtain c where "p asset = (\ n w. c)" by auto show "borel_predict_stoch_proc F (p asset)" unfolding predict_stoch_proc_def proof (intro conjI) show "p asset 0 \ borel_measurable (F 0)" using \p asset = (\ n w. c)\ by simp show "\n. p asset (Suc n) \ borel_measurable (F n)" proof fix n have "p asset (Suc n) = (\ w. c)" using \p asset = (\ n w. c)\ by simp thus "p asset (Suc n) \ borel_measurable (F n)" by simp qed qed qed lemma two_component_support_set: assumes "\ n w. a n w \ 0" and "\ n w. b n w\ 0" and "x \ y" shows "support_set ((\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b)) = {x,y}" proof let ?arb_pf = "(\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b)" have "\ n w. ?arb_pf x n w \ 0" using assms by simp moreover have "\n w. ?arb_pf y n w \ 0" using assms by simp ultimately show "{x, y} \ support_set ?arb_pf" unfolding support_set_def by simp show "support_set ?arb_pf \ {x, y}" proof (rule ccontr) assume "\ support_set ?arb_pf \ {x, y}" hence "\z. z\ support_set ?arb_pf \ z\ {x, y}" by auto from this obtain z where "z\ support_set ?arb_pf" and "z\ {x, y}" by auto have "\n w. ?arb_pf z n w \ 0" using \z\ support_set ?arb_pf\ unfolding support_set_def by simp from this obtain n w where "?arb_pf z n w \ 0" by auto have "?arb_pf z n w = 0" using \z\ {x, y}\ by simp thus False using \?arb_pf z n w \ 0\ by simp qed qed lemma two_component_val_process: assumes "arb_pf = ((\ (x::'b) (n::nat) (w::'a). 0::real)(x:= a, y:= b))" and "portfolio arb_pf" and "x \ y" and "\ n w. a n w \ 0" and "\ n w. b n w\ 0" shows "val_process Mkt arb_pf n w = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" proof - have "support_set arb_pf = {x,y}" using assms by (simp add:two_component_support_set) have "val_process Mkt arb_pf n w = (\x\support_set arb_pf. prices Mkt x n w * arb_pf x (Suc n) w)" unfolding val_process_def using \portfolio arb_pf\ by simp also have "... = (\x\{x, y}. prices Mkt x n w * arb_pf x (Suc n) w)" using \support_set arb_pf = {x, y}\ by simp also have "... = (\x\{y}. prices Mkt x n w * arb_pf x (Suc n) w) + prices Mkt x n w * arb_pf x (Suc n) w" using sum.insert[of "{y}" x "\x. prices Mkt x n w * arb_pf x n w"] assms(3) by auto also have "... = prices Mkt y n w * arb_pf y (Suc n) w + prices Mkt x n w * arb_pf x (Suc n) w" by simp also have "... = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" using assms by auto finally show "val_process Mkt arb_pf n w = prices Mkt y n w * b (Suc n) w + prices Mkt x n w * a (Suc n) w" . qed lemma quantity_update_support_set: assumes "\n w. pr n w \ 0" and "x\ support_set p" shows "support_set (p(x:=pr)) = support_set p \ {x}" proof show "support_set (p(x := pr)) \ support_set p \ {x}" proof fix y assume "y\ support_set (p(x := pr))" show "y \ support_set p \ {x}" proof (rule ccontr) assume "\y \ support_set p \ {x}" hence "y \ x" by simp have "\n w. (p(x := pr)) y n w \ 0" using \y\ support_set (p(x := pr))\ unfolding support_set_def by simp then obtain n w where nwprop: "(p(x := pr)) y n w \ 0" by auto have "y\ support_set p" using \\y \ support_set p \ {x}\ by simp hence "y = x" using nwprop using support_set_def by force thus False using \y\ x\ by simp qed qed show "support_set p \ {x} \ support_set (p(x := pr))" proof fix y assume "y \ support_set p \ {x}" show "y\ support_set (p(x := pr))" proof (cases "y\ support_set p") case True thus ?thesis proof - have f1: "y \ {b. \n a. p b n a \ 0}" by (metis True support_set_def) then have "y \ x" using assms(2) support_set_def by force then show ?thesis using f1 by (simp add: support_set_def) qed next case False hence "y = x" using \y \ support_set p \ {x}\ by auto thus ?thesis using assms by (simp add: support_set_def) qed qed qed lemma fix_asset_price: shows "\x Mkt2. x \ stocks Mkt \ coincides_on Mkt Mkt2 (stocks Mkt) \ prices Mkt2 x = pr" proof - have "\x. x\ stocks Mkt" by (metis UNIV_eq_I stk_strict_subs_def mkt_stocks_assets) from this obtain x where "x\ stocks Mkt" by auto let ?res = "discrete_market_of (stocks Mkt) ((prices Mkt)(x:=pr))" have "coincides_on Mkt ?res (stocks Mkt)" proof - have "stocks Mkt = stocks (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr)))" by (metis (no_types) stk_strict_subs_def mkt_stocks_assets stocks_of) then show ?thesis by (simp add: \x \ stocks Mkt\ coincides_on_def prices_of) qed have "prices ?res x = pr" by (simp add: prices_of) show ?thesis using \coincides_on Mkt (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr))) (stocks Mkt)\ \prices (discrete_market_of (stocks Mkt) ((prices Mkt)(x := pr))) x = pr\ \x \ stocks Mkt\ by blast qed lemma (in disc_market_pos_stock) arbitrage_portfolio_properties: assumes "price_structure der matur \ pr" and "replicating_portfolio pf der matur" and "(coincides_on Mkt Mkt2 (stocks Mkt))" and "(prices Mkt2 x = pr)" and "x\ stocks Mkt" and "diff_inv = (\ - initial_value pf) / constant_image (prices Mkt pos_stock 0)" and "diff_inv \ 0" and "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" and "contr_pf = qty_sum arb_pf pf" shows "self_financing Mkt2 contr_pf" and "trading_strategy contr_pf" and "\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0" and "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" and "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" and "support_set arb_pf = {x, pos_stock}" and "portfolio contr_pf" proof - have "0 < constant_image (prices Mkt pos_stock 0)" using trading_strategy_init proof - have "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using pos_stock_borel_adapted by simp hence "\c. \w\space M. prices Mkt pos_stock 0 w = c" using adapted_init[of "prices Mkt pos_stock"] by simp moreover have "\w\ space M. 0 < prices Mkt pos_stock 0 w" using positive by simp ultimately show ?thesis using constant_image_pos by simp qed show "support_set arb_pf = {x, pos_stock}" proof - have "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" using \arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))\ . moreover have "\n w. diff_inv \ 0" using assms by simp moreover have "x\ pos_stock" using \x \ stocks Mkt\ in_stock by auto ultimately show ?thesis by (simp add:two_component_support_set) qed hence "portfolio arb_pf" unfolding portfolio_def by simp have arb_vp:"\n w. val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" proof (intro allI) fix n w have "val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * (\ n w. diff_inv) n w + prices Mkt2 x n w * (\ n w. -1) n w" proof (rule two_component_val_process) show "x\ pos_stock" using \x \ stocks Mkt\ in_stock by auto show "arb_pf = (\x n w. 0)(x := \a b. - 1, pos_stock := \a b. diff_inv)" using assms by simp show "portfolio arb_pf" using \portfolio arb_pf\ by simp show "\n w. - (1::real) \ 0" by simp show "\n w. diff_inv \ 0" using assms by auto qed also have "... = prices Mkt2 pos_stock n w * diff_inv - pr n w" using \prices Mkt2 x = pr\ by simp finally show "val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" . qed have "static_quantities arb_pf" unfolding static_quantities_def proof fix asset assume "asset \ support_set arb_pf" thus "\c. arb_pf asset = (\n w. c)" proof (cases "asset = x") case True thus ?thesis using assms by auto next case False hence "asset = pos_stock" using \support_set arb_pf = {x, pos_stock}\ using \asset \ support_set arb_pf\ by blast thus ?thesis using assms by auto qed qed hence "trading_strategy arb_pf" using \portfolio arb_pf\ portfolio_def static_quantities_trading_strat by blast have "self_financing Mkt2 arb_pf" by (simp add: static_portfolio_self_financing \arb_pf = (\x n w. 0) (x := \n w. -1, pos_stock := \n w. diff_inv)\) hence arb_uvp: "\n w. cls_val_process Mkt2 arb_pf n w = prices Mkt2 pos_stock n w * diff_inv - pr n w" using assms arb_vp by (simp add:self_financingE) show "portfolio contr_pf" using assms by (metis \support_set arb_pf = {x, pos_stock}\ replicating_portfolio_def finite.emptyI finite.insertI portfolio_def stock_portfolio_def sum_portfolio) have "support_set contr_pf \ stocks Mkt \ {x}" proof - have "support_set contr_pf \ support_set arb_pf \ support_set pf" using assms by (simp add:sum_support_set) moreover have "support_set arb_pf \ stocks Mkt \ {x}" using \support_set arb_pf = {x, pos_stock}\ in_stock by simp moreover have "support_set pf \ stocks Mkt \ {x}" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto ultimately show ?thesis by auto qed show "self_financing Mkt2 contr_pf" proof - have "self_financing Mkt2 (qty_sum arb_pf pf)" proof (rule sum_self_financing) show "portfolio arb_pf" using \support_set arb_pf = {x, pos_stock}\ unfolding portfolio_def by auto show "portfolio pf" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto show "self_financing Mkt2 pf" using coincides_on_self_financing \(coincides_on Mkt Mkt2 (stocks Mkt))\ \(prices Mkt2 x = pr)\ assms(2) unfolding replicating_portfolio_def stock_portfolio_def by blast show "self_financing Mkt2 arb_pf" by (simp add: static_portfolio_self_financing \arb_pf = (\x n w. 0) (x := \n w. -1, pos_stock := \n w. diff_inv)\) qed thus ?thesis using assms by simp qed show "trading_strategy contr_pf" proof - have "trading_strategy (qty_sum arb_pf pf)" proof (rule sum_trading_strat) show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "trading_strategy arb_pf" using \trading_strategy arb_pf\ . qed thus ?thesis using assms by simp qed show "\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0" proof fix w assume "w\ space M" have "cls_val_process Mkt2 contr_pf 0 w = cls_val_process Mkt2 arb_pf 0 w + cls_val_process Mkt2 pf 0 w" using sum_cls_val_process0[of arb_pf pf Mkt2] using \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock 0 w * diff_inv - pr 0 w + cls_val_process Mkt2 pf 0 w" using arb_uvp by simp also have "... = constant_image (prices Mkt pos_stock 0) * diff_inv - pr 0 w + cls_val_process Mkt2 pf 0 w" proof - have f1: "prices Mkt pos_stock = prices Mkt2 pos_stock" using \coincides_on Mkt Mkt2 (stocks Mkt)\ in_stock unfolding coincides_on_def by blast have "prices Mkt pos_stock 0 w = constant_image (prices Mkt pos_stock 0)" using \w \ space M\ adapted_init constant_imageI pos_stock_borel_adapted by blast then show ?thesis using f1 by simp qed also have "... = (\ - initial_value pf) - pr 0 w + cls_val_process Mkt2 pf 0 w" using \0 < constant_image (prices Mkt pos_stock 0)\ assms by simp also have "... = (\ - initial_value pf) - \ + cls_val_process Mkt2 pf 0 w" using \price_structure der matur \ pr\ price_structure_init[of der matur \ pr] by (simp add: \w \ space M\) also have "... = (\ - initial_value pf) - \ + (initial_value pf)" using initial_valueI assms unfolding replicating_portfolio_def using \w \ space M\ coincides_stocks_cls_val_process self_financingE readable by (metis (no_types, opaque_lifting) support_adapt_def stock_portfolio_def subsetCE) also have "... = 0" by simp finally show "cls_val_process Mkt2 contr_pf 0 w = 0" . qed show "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" proof assume "0 < diff_inv" show "AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w" proof (rule AE_mp) have "AE w in M. prices Mkt2 x matur w = der w" using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ unfolding price_structure_def by auto moreover have "AE w in M. cls_val_process Mkt2 pf matur w = der w" using assms coincides_stocks_cls_val_process[of Mkt pf Mkt2] \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by auto ultimately show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" by auto show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w \ 0 < cls_val_process Mkt2 contr_pf matur w" proof (rule AE_I2, rule impI) fix w assume "w\ space M" and "prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" have "cls_val_process Mkt2 contr_pf matur w = cls_val_process Mkt2 arb_pf matur w + cls_val_process Mkt2 pf matur w" using sum_cls_val_process[of arb_pf pf Mkt2] \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock matur w * diff_inv - pr matur w + cls_val_process Mkt2 pf matur w" using arb_uvp by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv - prices Mkt2 x matur w + cls_val_process Mkt2 pf matur w" using \prices Mkt2 x = pr\ by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv" using \prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w\ by simp also have "... > 0" using positive \0 < diff_inv\ by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ coincides_on_def in_stock mult_pos_pos) finally have "cls_val_process Mkt2 contr_pf matur w > 0". thus "0 < cls_val_process Mkt2 contr_pf matur w" by simp qed qed qed show "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" proof assume "diff_inv < 0" show "AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w" proof (rule AE_mp) have "AE w in M. prices Mkt2 x matur w = der w" using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ unfolding price_structure_def by auto moreover have "AE w in M. cls_val_process Mkt2 pf matur w = der w" using assms coincides_stocks_cls_val_process[of Mkt pf Mkt2] \coincides_on Mkt Mkt2 (stocks Mkt)\ unfolding replicating_portfolio_def by auto ultimately show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" by auto show "AE w in M. prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w \ 0 > cls_val_process Mkt2 contr_pf matur w" proof (rule AE_I2, rule impI) fix w assume "w\ space M" and "prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w" have "cls_val_process Mkt2 contr_pf matur w = cls_val_process Mkt2 arb_pf matur w + cls_val_process Mkt2 pf matur w" using sum_cls_val_process[of arb_pf pf Mkt2] \portfolio arb_pf\ assms replicating_portfolio_def stock_portfolio_def by blast also have "... = prices Mkt2 pos_stock matur w * diff_inv - pr matur w + cls_val_process Mkt2 pf matur w" using arb_uvp by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv - prices Mkt2 x matur w + cls_val_process Mkt2 pf matur w" using \prices Mkt2 x = pr\ by simp also have "... = prices Mkt2 pos_stock matur w * diff_inv" using \prices Mkt2 x matur w = cls_val_process Mkt2 pf matur w\ by simp also have "... < 0" using positive \diff_inv < 0\ by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ coincides_on_def in_stock mult_pos_neg) finally have "cls_val_process Mkt2 contr_pf matur w < 0". thus "0 > cls_val_process Mkt2 contr_pf matur w" by simp qed qed qed qed lemma (in disc_equity_market) mult_comp_cls_val_process_measurable': assumes "cls_val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "qty n \ borel_measurable (F n)" and "0 \ n" shows "cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof - have "\m. n = Suc m" using assms by presburger from this obtain m where "n = Suc m" by auto hence "cls_val_process Mkt2 (qty_mult_comp pf qty) (Suc m) \ borel_measurable (F (Suc m))" using mult_comp_cls_val_process_Suc[of pf Mkt2 qty] borel_measurable_times[of "cls_val_process Mkt2 pf (Suc m)" "F (Suc m)" "qty (Suc m)"] assms \n= Suc m\ by presburger thus ?thesis using \n = Suc m\ by simp qed lemma (in disc_equity_market) mult_comp_cls_val_process_measurable: assumes "\n. cls_val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "\n. qty (Suc n) \ borel_measurable (F n)" shows "\n. cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof fix n show "cls_val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" proof (cases "n=0") case False hence "\m. n = Suc m" by presburger from this obtain m where "n = Suc m" by auto have "qty n \ borel_measurable (F n)" using Suc_n_not_le_n \n = Suc m\ assms(3) increasing_measurable_info nat_le_linear by blast hence "qty (Suc m) \ borel_measurable (F (Suc m))" using \n = Suc m\ by simp hence "cls_val_process Mkt2 (qty_mult_comp pf qty) (Suc m) \ borel_measurable (F (Suc m))" using mult_comp_cls_val_process_Suc[of pf Mkt2 qty] borel_measurable_times[of "cls_val_process Mkt2 pf (Suc m)" "F (Suc m)" "qty (Suc m)"] assms \n= Suc m\ by presburger thus ?thesis using \n = Suc m\ by simp next case True have "qty (Suc 0) \ borel_measurable (F 0)" using assms by simp moreover have "cls_val_process Mkt2 pf 0 \ borel_measurable (F 0)" using assms by simp ultimately have "(\w. cls_val_process Mkt2 pf 0 w * qty (Suc 0) w) \ borel_measurable (F 0)" by simp thus ?thesis using assms(2) True mult_comp_cls_val_process0 by (simp add: \(\w. cls_val_process Mkt2 pf 0 w * qty (Suc 0) w) \ borel_measurable (F 0)\ mult_comp_cls_val_process0 measurable_cong) qed qed lemma (in disc_equity_market) mult_comp_val_process_measurable: assumes "val_process Mkt2 pf n \borel_measurable (F n)" and "portfolio pf" and "qty (Suc n) \ borel_measurable (F n)" shows "val_process Mkt2 (qty_mult_comp pf qty) n \ borel_measurable (F n)" using mult_comp_val_process[of pf Mkt2 qty] borel_measurable_times[of "val_process Mkt2 pf n" "F n" "qty (Suc n)"] assms by presburger lemma (in disc_market_pos_stock) repl_fair_price_unique: assumes "replicating_portfolio pf der matur" and "fair_price Mkt \ der matur" shows "\ = initial_value pf" proof - have expr: "(\ pr. price_structure der matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using assms fair_priceI by simp then obtain pr where "price_structure der matur \ pr" and xasset: "(\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p)))" by auto define diff_inv where "diff_inv = (\ - initial_value pf) / constant_image (prices Mkt pos_stock 0)" { fix x assume "x\ stocks Mkt" hence mkprop: "(\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p)" using xasset by simp fix Mkt2 assume "(coincides_on Mkt Mkt2 (stocks Mkt))" and "(prices Mkt2 x = pr)" have "0 < constant_image (prices Mkt pos_stock 0)" using trading_strategy_init proof - have "borel_adapt_stoch_proc F (prices Mkt pos_stock)" using pos_stock_borel_adapted by simp hence "\c. \w\space M. prices Mkt pos_stock 0 w = c" using adapted_init[of "prices Mkt pos_stock"] by simp moreover have "\w\ space M. 0 < prices Mkt pos_stock 0 w" using positive by simp ultimately show ?thesis using constant_image_pos by simp qed define arb_pf where "arb_pf = (\ (x::'b) (n::nat) (w::'a). 0::real)(x:= (\ n w. -1), pos_stock := (\ n w. diff_inv))" define contr_pf where "contr_pf = qty_sum arb_pf pf" have 1:"0 \ diff_inv \ self_financing Mkt2 contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 2:"0 \ diff_inv \ trading_strategy contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 3:"0 \ diff_inv \ (\w\ space M. cls_val_process Mkt2 contr_pf 0 w = 0)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 4: "0 < diff_inv \ (AE w in M. 0 < cls_val_process Mkt2 contr_pf matur w)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 5: "diff_inv < 0 \ (AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w)" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 6: "0 \ diff_inv \ support_set arb_pf = {x, pos_stock}" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 7: "0 \ diff_inv \support_set contr_pf \ stocks Mkt \ {x}" proof - have "0 \ diff_inv \ support_set contr_pf \ support_set arb_pf \ support_set pf" unfolding contr_pf_def by (simp add:sum_support_set) moreover have "0 \ diff_inv \support_set arb_pf \ stocks Mkt \ {x}" using \0 \ diff_inv \ support_set arb_pf = {x, pos_stock}\ in_stock by simp moreover have "0 \ diff_inv \support_set pf \ stocks Mkt \ {x}" using assms unfolding replicating_portfolio_def stock_portfolio_def by auto ultimately show ?thesis by auto qed have 8:"0 \ diff_inv \portfolio contr_pf" using arbitrage_portfolio_properties[of der matur \ pr pf Mkt2 x diff_inv arb_pf contr_pf] using \coincides_on Mkt Mkt2 (stocks Mkt)\ \price_structure der matur \ pr\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ arb_pf_def assms(1) contr_pf_def diff_inv_def by blast have 9: "0 \ diff_inv \ cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" proof assume "0 \ diff_inv" have 10:"\ asset \ support_set arb_pf \ support_set pf. prices Mkt2 asset matur \ borel_measurable (F matur)" proof fix asset assume "asset \ support_set arb_pf \ support_set pf" show "prices Mkt2 asset matur \ borel_measurable (F matur)" proof (cases "asset \ support_set pf") case True thus ?thesis using assms readable by (metis (no_types, lifting) \coincides_on Mkt Mkt2 (stocks Mkt)\ adapt_stoch_proc_def coincides_on_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms stock_portfolio_def subsetCE) next case False hence "asset\ support_set arb_pf" using \asset \ support_set arb_pf \ support_set pf\ by auto show ?thesis proof (cases "asset = x") case True thus ?thesis using \price_structure der matur \ pr\ \prices Mkt2 x = pr\ price_structure_borel_measurable by blast next case False hence "asset = pos_stock" using \asset\ support_set arb_pf\ \0 \ diff_inv \ support_set arb_pf = {x, pos_stock}\ \0 \ diff_inv\ by auto thus ?thesis by (metis \coincides_on Mkt Mkt2 (stocks Mkt)\ adapt_stoch_proc_def coincides_on_def in_stock pos_stock_borel_adapted) qed qed qed moreover have "\asset\support_set contr_pf. contr_pf asset matur \ borel_measurable (F matur)" using \0 \ diff_inv \trading_strategy contr_pf\ \0 \ diff_inv\ by (metis adapt_stoch_proc_def disc_filtr_prob_space.predict_imp_adapt disc_filtr_prob_space_axioms trading_strategy_def) ultimately show "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" proof- have "\asset\support_set contr_pf. contr_pf asset (Suc matur) \ borel_measurable (F matur)" using \0 \ diff_inv \trading_strategy contr_pf\ \0 \ diff_inv\ by (simp add: predict_stoch_proc_def trading_strategy_def) moreover have "\asset\support_set contr_pf. prices Mkt2 asset matur \ borel_measurable (F matur)" using 10 unfolding contr_pf_def using sum_support_set[of arb_pf pf] by auto ultimately show ?thesis by (metis (no_types, lifting) "1" \0 \ diff_inv\ quantity_adapted self_financingE) qed qed { assume "0 > diff_inv" define opp_pf where "opp_pf = qty_mult_comp contr_pf (\ n w. -1)" have "arbitrage_process Mkt2 opp_pf" proof (rule arbitrage_processI, rule exI, intro conjI) show "self_financing Mkt2 opp_pf" using 1 \0 > diff_inv\ mult_time_constant_self_financing[of contr_pf] 8 unfolding opp_pf_def by auto show "trading_strategy opp_pf" unfolding opp_pf_def proof (rule mult_comp_trading_strat) show "trading_strategy contr_pf" using 2 \0 > diff_inv\ by auto show "borel_predict_stoch_proc F (\n w. - 1)" by (simp add: constant_process_borel_predictable) qed show "\w\space M. cls_val_process Mkt2 opp_pf 0 w = 0" proof fix w assume "w\ space M" show "cls_val_process Mkt2 opp_pf 0 w = 0" using 3 8 \0 > diff_inv\ using \w \ space M\ mult_comp_cls_val_process0 opp_pf_def by fastforce qed have "AE w in M. 0 < cls_val_process Mkt2 opp_pf matur w" proof (rule AE_mp) show "AE w in M. 0 > cls_val_process Mkt2 contr_pf matur w" using 5 \0 > diff_inv\ by auto show "AE w in M. cls_val_process Mkt2 contr_pf matur w < 0 \ 0 < cls_val_process Mkt2 opp_pf matur w" proof fix w assume "w\ space M" show "cls_val_process Mkt2 contr_pf matur w < 0 \ 0 < cls_val_process Mkt2 opp_pf matur w" proof assume "cls_val_process Mkt2 contr_pf matur w < 0" show "0 < cls_val_process Mkt2 opp_pf matur w" proof (cases "matur = 0") case False hence "\m. Suc m = matur" by presburger from this obtain m where "Suc m = matur" by auto hence "0 < cls_val_process Mkt2 opp_pf (Suc m) w" using 3 8 \0 > diff_inv\ \w \ space M\ mult_comp_cls_val_process_Suc opp_pf_def using \cls_val_process Mkt2 contr_pf matur w < 0\ by fastforce thus ?thesis using \Suc m = matur\ by simp next case True thus ?thesis using 3 8 \0 > diff_inv\ \w \ space M\ mult_comp_cls_val_process0 opp_pf_def using \cls_val_process Mkt2 contr_pf matur w < 0\ by auto qed qed qed qed thus "AE w in M. 0 \ cls_val_process Mkt2 opp_pf matur w" by auto show "0 < prob {w \ space M. 0 < cls_val_process Mkt2 opp_pf matur w}" proof - let ?P = "{w\ space M. 0 < cls_val_process Mkt2 opp_pf matur w}" have "cls_val_process Mkt2 opp_pf matur \ borel_measurable (F matur)" (*unfolding opp_pf_def *) proof - have "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" using 9 \0 > diff_inv\ by simp moreover have "portfolio contr_pf" using 8 \0 > diff_inv\ by simp moreover have "(\w. - 1) \ borel_measurable (F matur)" by (simp add:constant_process_borel_adapted) ultimately show ?thesis using mult_comp_cls_val_process_measurable proof - have "diff_inv \ 0" using \diff_inv < 0\ by blast then have "self_financing Mkt2 contr_pf" by (metis "1") then show ?thesis by (metis (no_types) \(\w. - 1) \ borel_measurable (F matur)\ \portfolio contr_pf\ \self_financing Mkt2 opp_pf\ \cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)\ mult_comp_val_process_measurable opp_pf_def self_financingE) qed qed moreover have "space M = space (F matur)" using filtration by (simp add: filtration_def subalgebra_def) ultimately have "?P \ sets (F matur)" using borel_measurable_iff_greater[of "val_process Mkt2 contr_pf matur" "F matur"] by auto hence "?P \ sets M" by (meson filtration filtration_def subalgebra_def subsetCE) hence "measure M ?P = 1" using prob_Collect_eq_1[of "\x. 0 < cls_val_process Mkt2 opp_pf matur x"] \AE w in M. 0 < cls_val_process Mkt2 opp_pf matur w\ \0 > diff_inv\ by blast thus ?thesis by simp qed qed have "\ p. portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p" proof(intro exI conjI) show "arbitrage_process Mkt2 opp_pf" using \arbitrage_process Mkt2 opp_pf\ . show "portfolio opp_pf" unfolding opp_pf_def using 8 \0 > diff_inv\ by (auto simp add: mult_comp_portfolio) show "support_set opp_pf \ stocks Mkt \ {x}" unfolding opp_pf_def using 7 \0 > diff_inv\ using mult_comp_support_set by fastforce qed } note negp = this { assume "0 < diff_inv" have "arbitrage_process Mkt2 contr_pf" proof (rule arbitrage_processI, rule exI, intro conjI) show "self_financing Mkt2 contr_pf" using 1 \0 < diff_inv\ by auto show "trading_strategy contr_pf" using 2 \0 < diff_inv\ by auto show "\w\space M. cls_val_process Mkt2 contr_pf 0 w = 0" using 3 \0 < diff_inv\ by auto show "AE w in M. 0 \ cls_val_process Mkt2 contr_pf matur w" using 4 \0 < diff_inv\ by auto show "0 < prob {w \ space M. 0 < cls_val_process Mkt2 contr_pf matur w}" proof - let ?P = "{w\ space M. 0 < cls_val_process Mkt2 contr_pf matur w}" have "cls_val_process Mkt2 contr_pf matur \ borel_measurable (F matur)" using 9 \0 < diff_inv\ by auto moreover have "space M = space (F matur)" using filtration by (simp add: filtration_def subalgebra_def) ultimately have "?P \ sets (F matur)" using borel_measurable_iff_greater[of "val_process Mkt2 contr_pf matur" "F matur"] by auto hence "?P \ sets M" by (meson filtration filtration_def subalgebra_def subsetCE) hence "measure M ?P = 1" using prob_Collect_eq_1[of "\x. 0 < cls_val_process Mkt2 contr_pf matur x"] 4 \0 < diff_inv\ by blast thus ?thesis by simp qed qed have "\ p. portfolio p \ support_set p \ stocks Mkt \ {x} \ arbitrage_process Mkt2 p" proof(intro exI conjI) show "arbitrage_process Mkt2 contr_pf" using \arbitrage_process Mkt2 contr_pf\ . show "portfolio contr_pf" using 8 \0 < diff_inv\ by auto show "support_set contr_pf \ stocks Mkt \ {x}" using 7 \0 < diff_inv\ by auto qed } note posp = this have "diff_inv \ 0 \ \(\ pr. price_structure der matur \ pr \ (\ x. (x\ stocks Mkt \ (\ Mkt2 p. (coincides_on Mkt Mkt2 (stocks Mkt)) \ (prices Mkt2 x = pr) \ portfolio p \ support_set p \ stocks Mkt \ {x} \ \ arbitrage_process Mkt2 p))))" using \coincides_on Mkt Mkt2 (stocks Mkt)\ \prices Mkt2 x = pr\ \x \ stocks Mkt\ xasset posp negp by force } hence "diff_inv = 0" using fix_asset_price expr by metis moreover have "constant_image (prices Mkt pos_stock 0) > 0" by (simp add: adapted_init constant_image_pos pos_stock_borel_adapted positive) ultimately show ?thesis unfolding diff_inv_def by auto qed subsection \Risk-neutral probability space\ subsubsection \risk-free rate and discount factor processes\ fun disc_rfr_proc:: "real \ nat \ 'a \ real" where rfr_base: "(disc_rfr_proc r) 0 w = 1"| rfr_step: "(disc_rfr_proc r) (Suc n) w = (1+r) * (disc_rfr_proc r) n w" lemma disc_rfr_proc_borel_measurable: shows "(disc_rfr_proc r) n \ borel_measurable M" proof (induct n) case (Suc n) thus ?case by (simp add:borel_measurable_times) qed auto lemma disc_rfr_proc_nonrandom: fixes r::real shows "\n. disc_rfr_proc r n \ borel_measurable (F 0)" using disc_rfr_proc_borel_measurable by auto lemma (in disc_equity_market) disc_rfr_constant_time: shows "\c. \w \ space (F 0). (disc_rfr_proc r n) w = c" proof (rule triv_measurable_cst) show "space (F 0) = space M" using filtration by (simp add:filtration_def subalgebra_def) show "sets (F 0) = {{}, space M}" using info_disc_filtr by (simp add: bot_nat_def init_triv_filt_def) show "(disc_rfr_proc r n) \ borel_measurable (F 0)" using disc_rfr_proc_nonrandom by blast show "space M \ {}" by (simp add:not_empty) qed lemma (in disc_filtr_prob_space) disc_rfr_proc_borel_adapted: shows "borel_adapt_stoch_proc F (disc_rfr_proc r)" unfolding adapt_stoch_proc_def using disc_rfr_proc_nonrandom filtration unfolding filtration_def by (meson increasing_measurable_info le0) lemma disc_rfr_proc_positive: assumes "-1 < r" shows "\n w . 0 < disc_rfr_proc r n w" proof - fix n fix w::'a show "0 < disc_rfr_proc r n w" proof (induct n) case 0 thus ?case using assms "disc_rfr_proc.simps" by simp next case (Suc n) thus ?case using assms "disc_rfr_proc.simps" by simp qed qed lemma (in prob_space) disc_rfr_constant_time_pos: assumes "-1 < r" shows "\c > 0. \w \ space M. (disc_rfr_proc r n) w = c" proof - let ?F = "sigma (space M) {{}, space M}" have ex: "\c. \w \ space ?F. (disc_rfr_proc r n) w = c" proof (rule triv_measurable_cst) show "space ?F = space M" by simp show "sets ?F = {{}, space M}" by (meson sigma_algebra.sets_measure_of_eq sigma_algebra_trivial) show "(disc_rfr_proc r n) \ borel_measurable ?F" using disc_rfr_proc_borel_measurable by blast show "space M \ {}" by (simp add:not_empty) qed from this obtain c where "\w \ space ?F. (disc_rfr_proc r n) w = c" by auto note cprops = this have "c>0" proof - have "\ w. w\ space M" using subprob_not_empty by blast from this obtain w where "w\ space M" by auto hence "c = disc_rfr_proc r n w" using cprops by simp also have "... > 0" using disc_rfr_proc_positive[of r n] assms by simp finally show ?thesis . qed moreover have "space M = space ?F" by simp ultimately show ?thesis using ex using cprops by blast qed lemma disc_rfr_proc_Suc_div: assumes "-1 < r" shows "\w. disc_rfr_proc r (Suc n) w/disc_rfr_proc r n w = 1+r" proof - fix w::'a show "disc_rfr_proc r (Suc n) w/disc_rfr_proc r n w = 1+r" using disc_rfr_proc_positive assms by (metis rfr_step less_irrefl nonzero_eq_divide_eq) qed definition discount_factor where "discount_factor r n = (\w. inverse (disc_rfr_proc r n w))" lemma discount_factor_times_rfr: assumes "-1 < r" shows "(1+r) * discount_factor r (Suc n) w = discount_factor r n w" unfolding discount_factor_def using assms by simp lemma discount_factor_borel_measurable: shows "discount_factor r n \ borel_measurable M" unfolding discount_factor_def proof (rule borel_measurable_inverse) show "disc_rfr_proc r n \ borel_measurable M" by (simp add:disc_rfr_proc_borel_measurable) qed lemma discount_factor_init: shows "discount_factor r 0 = (\w. 1)" unfolding discount_factor_def by simp lemma discount_factor_nonrandom: shows "discount_factor r n \ borel_measurable M" unfolding discount_factor_def proof (rule borel_measurable_inverse) show "disc_rfr_proc r n \ borel_measurable M" by (simp add:disc_rfr_proc_borel_measurable) qed lemma discount_factor_positive: assumes "-1 < r" shows "\n w . 0 < discount_factor r n w" using assms disc_rfr_proc_positive unfolding discount_factor_def by auto lemma (in prob_space) discount_factor_constant_time_pos: assumes "-1 < r" shows "\c > 0. \w \ space M. (discount_factor r n) w = c" using disc_rfr_constant_time_pos unfolding discount_factor_def by (metis assms inverse_positive_iff_positive) locale rsk_free_asset = fixes Mkt r risk_free_asset assumes acceptable_rate: "-1 < r" and rf_price: "prices Mkt risk_free_asset = disc_rfr_proc r" and rf_stock: "risk_free_asset \ stocks Mkt" locale rfr_disc_equity_market = disc_equity_market + rsk_free_asset + assumes rd: "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" sublocale rfr_disc_equity_market \ disc_market_pos_stock _ _ _ "risk_free_asset" by (unfold_locales, (auto simp add: rf_stock rd disc_rfr_proc_positive rf_price acceptable_rate)) subsubsection \Discounted value of a stochastic process\ definition discounted_value where "discounted_value r X = (\ n w. discount_factor r n w * X n w)" lemma (in rfr_disc_equity_market) discounted_rfr: shows "discounted_value r (prices Mkt risk_free_asset) n w = 1" unfolding discounted_value_def discount_factor_def using rf_price by (metis less_irrefl mult.commute positive right_inverse) lemma discounted_init: shows "\w. discounted_value r X 0 w = X 0 w" unfolding discounted_value_def by (simp add: discount_factor_init) lemma discounted_mult: shows "\n w. discounted_value r (\m x. X m x * Y m x) n w = X n w * (discounted_value r Y) n w" by (simp add: discounted_value_def) lemma discounted_mult': shows "discounted_value r (\m x. X m x * Y m x) n w = X n w * (discounted_value r Y) n w" by (simp add: discounted_value_def) lemma discounted_mult_times_rfr: assumes "-1 < r" shows "discounted_value r (\m w. (1+r) * X w) (Suc n) w = discounted_value r (\m w. X w) n w" unfolding discounted_value_def using assms discount_factor_times_rfr discounted_mult by (simp add: discount_factor_times_rfr mult.commute) lemma discounted_cong: assumes "\n w. X n w = Y n w" shows "\ n w. discounted_value r X n w = discounted_value r Y n w" by (simp add: assms discounted_value_def) lemma discounted_cong': assumes "X n w = Y n w" shows "discounted_value r X n w = discounted_value r Y n w" by (simp add: assms discounted_value_def) lemma discounted_AE_cong: assumes "AE w in N. X n w = Y n w" shows "AE w in N. discounted_value r X n w = discounted_value r Y n w" proof (rule AE_mp) show "AE w in N. X n w = Y n w" using assms by simp show "AE w in N. X n w = Y n w \ discounted_value r X n w = discounted_value r Y n w" proof fix w assume "w\ space N" thus "X n w = Y n w \ discounted_value r X n w = discounted_value r Y n w " by (simp add:discounted_value_def) qed qed lemma discounted_sum: assumes "finite I" shows "\n w. (\ i\ I. (discounted_value r (\m x. f i m x)) n w) = (discounted_value r (\m x. (\i\ I. f i m x)) n w)" using assms(1) subset_refl[of I] proof (induct rule: finite_subset_induct) case empty then show ?case by (simp add: discounted_value_def) next case (insert a F) show ?case proof (intro allI) fix n w have "(\i\insert a F. discounted_value r (f i) n w) = discounted_value r (f a) n w + (\i\F. discounted_value r (f i) n w)" by (simp add: insert.hyps(1) insert.hyps(3)) also have "... = discounted_value r (f a) n w + discounted_value r (\m x. \i\F. f i m x) n w" using insert.hyps(4) by simp also have "... = discounted_value r (\m x. \i\insert a F. f i m x) n w" by (simp add: discounted_value_def insert.hyps(1) insert.hyps(3) ring_class.ring_distribs(1)) finally show "(\i\insert a F. discounted_value r (f i) n w) = discounted_value r (\m x. \i\insert a F. f i m x) n w" . qed qed lemma discounted_adapted: assumes "borel_adapt_stoch_proc F X" shows "borel_adapt_stoch_proc F (discounted_value r X)" unfolding adapt_stoch_proc_def proof fix t show "discounted_value r X t \ borel_measurable (F t)" unfolding discounted_value_def proof (rule borel_measurable_times) show "X t \ borel_measurable (F t)" using assms unfolding adapt_stoch_proc_def by simp show "discount_factor r t \ borel_measurable (F t)" using discount_factor_borel_measurable by auto qed qed lemma discounted_measurable: assumes "X\ borel_measurable N" shows "discounted_value r (\m. X) m \ borel_measurable N" unfolding discounted_value_def proof (rule borel_measurable_times) show "X\ borel_measurable N" using assms by simp show "discount_factor r m \ borel_measurable N" using discount_factor_borel_measurable by auto qed lemma (in prob_space) discounted_integrable: assumes "integrable N (X n)" and "-1 < r" and "space N = space M" shows "integrable N (discounted_value r X n)" unfolding discounted_value_def proof - have "\c> 0. \w \ space M. (discount_factor r n) w = c" using discount_factor_constant_time_pos assms by simp from this obtain c where "c > 0" and "\w \ space M. (discount_factor r n) w = c" by auto note cprops = this hence "\w \ space M. discount_factor r n w = c" using cprops by simp hence "\w \ space N. discount_factor r n w = c" using assms by simp thus "integrable N (\w. discount_factor r n w * X n w)" using \\w \ space N. discount_factor r n w = c\ assms - integrable_cong[of N N "(\w. discount_factor r n w * X n w)" "(\w. c * X n w)"] by simp + Bochner_Integration.integrable_cong[of N N "(\w. discount_factor r n w * X n w)" "(\w. c * X n w)"] by simp qed subsubsection \Results on risk-neutral probability spaces\ definition (in rfr_disc_equity_market) risk_neutral_prob where "risk_neutral_prob N \ (prob_space N) \ (\ asset \ stocks Mkt. martingale N F (discounted_value r (prices Mkt asset)))" lemma integrable_val_process: assumes "\ asset \ support_set pf. integrable M (\w. prices Mkt asset n w * pf asset (Suc n) w)" shows "integrable M (val_process Mkt pf n)" proof (cases "portfolio pf") case False thus ?thesis unfolding val_process_def by simp next case True hence "val_process Mkt pf n = (\w. \x\support_set pf. prices Mkt x n w * pf x (Suc n) w)" unfolding val_process_def by simp moreover have "integrable M (\w. \x\support_set pf. prices Mkt x n w * pf x (Suc n) w)" using assms by simp ultimately show ?thesis by simp qed lemma integrable_self_fin_uvp: assumes "\ asset \ support_set pf. integrable M (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "self_financing Mkt pf" shows "integrable M (cls_val_process Mkt pf n)" proof - have "val_process Mkt pf n = cls_val_process Mkt pf n" using assms by (simp add:self_financingE) moreover have "integrable M (val_process Mkt pf n)" using assms by (simp add:integrable_val_process) ultimately show ?thesis by simp qed lemma (in rfr_disc_equity_market) stocks_portfolio_risk_neutral: assumes "risk_neutral_prob N" and "trading_strategy pf" and "subalgebra N M" and "support_set pf \ stocks Mkt" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" shows "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" proof have nsigfin: "\n. sigma_finite_subalgebra N (F n)" using assms unfolding risk_neutral_prob_def martingale_def subalgebra_def using filtration filtration_def risk_neutral_prob_def prob_space.subalgebra_sigma_finite in_stock by metis have "disc_filtr_prob_space N F" proof - have "prob_space N" using assms unfolding risk_neutral_prob_def by simp moreover have "disc_filtr N F" using assms subalgebra_filtration by (metis (no_types, lifting) filtration disc_filtr_def filtration_def) ultimately show ?thesis by (simp add: disc_filtr_prob_space_axioms_def disc_filtr_prob_space_def) qed fix asset assume "asset \ support_set pf" hence "asset \ stocks Mkt" using assms by auto have "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable M" using assms readable by (meson \asset \ stocks Mkt\ borel_adapt_stoch_proc_borel_measurable discounted_adapted rfr_disc_equity_market.risk_neutral_prob_def rfr_disc_equity_market_axioms) hence b: "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable N" using assms Conditional_Expectation.measurable_from_subalg[of N M _ borel] by auto show "AEeq N (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) (discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n)" proof - have "AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) w = (real_cond_exp N (F n) (\z. pf asset (Suc n) z * discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "AE w in N. discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n) w = pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) (Suc n) w" by (simp add: discounted_value_def) show "discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n) \ borel_measurable N" proof - have "(\y. prices Mkt asset (Suc n) y * pf asset (Suc n) y) \ borel_measurable N" using assms \asset\ support_set pf\ by (simp add:borel_measurable_integrable) thus ?thesis unfolding discounted_value_def using discount_factor_borel_measurable[of r "Suc n" N] by simp qed show "(\z. pf asset (Suc n) z * discounted_value r (prices Mkt asset) (Suc n) z) \ borel_measurable N" proof - have "pf asset (Suc n) \ borel_measurable M" using assms \asset\ support_set pf\ unfolding trading_strategy_def using borel_predict_stoch_proc_borel_measurable[of "pf asset"] by auto hence a: "pf asset (Suc n) \ borel_measurable N" using assms Conditional_Expectation.measurable_from_subalg[of N M _ borel] by blast show ?thesis using a b by simp qed qed also have "AE w in N. (real_cond_exp N (F n) (\z. pf asset (Suc n) z * discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w = pf asset (Suc n) w * (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w" proof (rule sigma_finite_subalgebra.real_cond_exp_mult) show "discounted_value r (prices Mkt asset) (Suc n) \ borel_measurable N" using b by simp show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "pf asset (Suc n) \ borel_measurable (F n)" using assms \asset\ support_set pf\ unfolding trading_strategy_def predict_stoch_proc_def by auto show "integrable N (\z. pf asset (Suc n) z * discounted_value r (prices Mkt asset) (Suc n) z)" proof - have "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" using assms \asset \ support_set pf\ by auto hence "integrable N (discounted_value r (\m w. prices Mkt asset m w * pf asset m w) (Suc n))" using assms unfolding risk_neutral_prob_def using acceptable_rate by (auto simp add:discounted_integrable subalgebra_def) - thus ?thesis using discounted_mult - integrable_cong[of N N "discounted_value r (\m w. prices Mkt asset m w * pf asset m w) (Suc n)" "(\z. pf asset (Suc n) z * discounted_value r (prices Mkt asset) (Suc n) z)"] - by (simp add: discounted_value_def) + thus ?thesis + by (smt (verit, ccfv_SIG) Bochner_Integration.integrable_cong discounted_value_def mult.assoc mult.commute) qed qed also have "AE w in N. pf asset (Suc n) w * (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) w = pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) n w" proof - have "AEeq N (real_cond_exp N (F n) (\z. discounted_value r (\m y. prices Mkt asset m y) (Suc n) z)) (\z. discounted_value r (\m y. prices Mkt asset m y) n z)" proof - have "martingale N F (discounted_value r (prices Mkt asset))" using assms \asset \ stocks Mkt\ unfolding risk_neutral_prob_def by simp moreover have "filtrated_prob_space N F" using \disc_filtr_prob_space N F\ using assms(2) disc_filtr_prob_space.axioms(1) filtrated_prob_space.intro filtrated_prob_space_axioms.intro filtration prob_space_axioms by (metis assms(3) subalgebra_filtration) ultimately show ?thesis using martingaleAE[of N F "discounted_value r (prices Mkt asset)" n "Suc n"] assms by simp qed thus ?thesis by auto qed also have "AE w in N. pf asset (Suc n) w * discounted_value r (\m y. prices Mkt asset m y) n w = discounted_value r (\m y. pf asset (Suc m) y * prices Mkt asset m y) n w" by (simp add: discounted_value_def) also have "AE w in N. discounted_value r (\m y. pf asset (Suc m) y * prices Mkt asset m y) n w = discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n w" by (simp add: discounted_value_def) finally show "AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))) w = (\x. discounted_value r (\m y. prices Mkt asset m y * pf asset (Suc m) y) n x) w" . qed qed lemma (in rfr_disc_equity_market) self_fin_trad_strat_mart: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "trading_strategy pf" and "self_financing Mkt pf" and "stock_portfolio Mkt pf" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" shows "martingale N F (discounted_value r (cls_val_process Mkt pf))" (*unfolding martingale_def*) proof (rule disc_martingale_charact) show nsigfin: "\n. sigma_finite_subalgebra N (F n)" using filt_equiv_prob_space_subalgebra assms using filtration filtration_def risk_neutral_prob_def subalgebra_sigma_finite by fastforce show "filtration N F" using assms by (simp add:filt_equiv_filtration) have "borel_adapt_stoch_proc F (discounted_value r (cls_val_process Mkt pf))" using assms discounted_adapted cls_val_process_adapted[of pf] stock_portfolio_def by (metis (mono_tags, opaque_lifting) support_adapt_def readable subsetCE) thus "\m. discounted_value r (cls_val_process Mkt pf) m \ borel_measurable (F m)" unfolding adapt_stoch_proc_def by simp show "\t. integrable N (discounted_value r (cls_val_process Mkt pf) t)" proof fix t have "integrable N (cls_val_process Mkt pf t)" using assms by (simp add: integrable_self_fin_uvp) thus "integrable N (discounted_value r (cls_val_process Mkt pf) t)" using assms discounted_integrable acceptable_rate by (metis filt_equiv_space) qed show "\n. AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" proof fix n show "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" proof - { fix w assume "w\ space M" have "discounted_value r (cls_val_process Mkt pf) (Suc n) w = discount_factor r (Suc n) w * (\x\support_set pf. prices Mkt x (Suc n) w * pf x (Suc n) w)" unfolding discounted_value_def cls_val_process_def using assms unfolding trading_strategy_def by simp also have "... = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" by (metis (no_types, lifting) mult.assoc sum.cong sum_distrib_left) finally have "discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" . } hence space: "\w\ space M. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" by simp hence nspace: "\w\ space N. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\x\support_set pf. discount_factor r (Suc n) w * prices Mkt x (Suc n) w * pf x (Suc n) w)" using assms by (simp add:filt_equiv_space) have sup_disc: "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" using assms by (simp add:stocks_portfolio_risk_neutral filt_equiv_imp_subalgebra stock_portfolio_def) have "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = real_cond_exp N (F n) (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w" proof (rule sigma_finite_subalgebra.real_cond_exp_cong') show "sigma_finite_subalgebra N (F n)" using nsigfin .. show "\w\space N. discounted_value r (cls_val_process Mkt pf) (Suc n) w = (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w" using nspace by (metis (no_types, lifting) discounted_value_def mult.assoc sum.cong) show "(discounted_value r (cls_val_process Mkt pf) (Suc n)) \ borel_measurable N" using assms using \\t. integrable N (discounted_value r (cls_val_process Mkt pf) t)\ by blast qed also have "AE w in N. real_cond_exp N (F n) (\y. \x\support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n) y) w = (\x\ support_set pf. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w)" proof (rule sigma_finite_subalgebra.real_cond_exp_bsum) show "sigma_finite_subalgebra N (F n)" using filt_equiv_prob_space_subalgebra assms using filtration filtration_def risk_neutral_prob_def subalgebra_sigma_finite by fastforce fix asset assume "asset \ support_set pf" show "integrable N (discounted_value r (\m y. prices Mkt asset m y * pf asset m y) (Suc n))" proof (rule discounted_integrable) show "integrable N (\y. prices Mkt asset (Suc n) y * pf asset (Suc n) y)" using assms \asset\ support_set pf\ by simp show "space N = space M" using assms by (metis filt_equiv_space) show "-1 < r" using acceptable_rate by simp qed qed also have "AE w in N. (\x\ support_set pf. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w) = (\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w)" proof (rule AE_sum) show "finite (support_set pf)" using assms(3) portfolio_def trading_strategy_def by auto show "\x \ support_set pf. AE w in N. (real_cond_exp N (F n) (discounted_value r (\m y. prices Mkt x m y * pf x m y) (Suc n))) w = discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w" using sup_disc by simp qed also have "AE w in N. (\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (cls_val_process Mkt pf) n w" proof fix w assume "w\ space N" have "(\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (\m y. (\x\ support_set pf. prices Mkt x m y * pf x (Suc m) y)) n w" using discounted_sum assms(3) portfolio_def trading_strategy_def by (simp add: discounted_value_def sum_distrib_left) also have "... = discounted_value r (val_process Mkt pf) n w" unfolding val_process_def by (simp add: portfolio_def) also have "... = discounted_value r (cls_val_process Mkt pf) n w" using assms by (simp add:self_financingE discounted_cong) finally show "(\x\ support_set pf. discounted_value r (\m y. prices Mkt x m y * pf x (Suc m) y) n w) = discounted_value r (cls_val_process Mkt pf) n w" . qed finally show "AE w in N. real_cond_exp N (F n) (discounted_value r (cls_val_process Mkt pf) (Suc n)) w = discounted_value r (cls_val_process Mkt pf) n w" . qed qed qed lemma (in disc_filtr_prob_space) finite_integrable_vp: assumes "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "prob_space N" and "filt_equiv F M N" and "trading_strategy pf" and "\n. \ asset \ support_set pf. prices Mkt asset n \ borel_measurable M" shows "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms by simp have "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset n `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(2) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms by blast moreover have "(\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset n ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset n \ borel_measurable M" using assms \asset \ support_set pf\ by simp moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed lemma (in disc_filtr_prob_space) finite_integrable_uvp: assumes "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "prob_space N" and "filt_equiv F M N" and "trading_strategy pf" and "\n. \ asset \ support_set pf. prices Mkt asset n \ borel_measurable M" shows "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms by simp have "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset (Suc n) `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(2) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms by blast moreover have "(\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset (Suc n) ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset (Suc n) \ borel_measurable M" using assms using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed lemma (in rfr_disc_equity_market) self_fin_trad_strat_mart_finite: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "trading_strategy pf" and "self_financing Mkt pf" and "support_set pf \ stocks Mkt" and "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "\ asset\ stocks Mkt. borel_adapt_stoch_proc F (prices Mkt asset)" shows "martingale N F (discounted_value r (cls_val_process Mkt pf))" proof (rule self_fin_trad_strat_mart, (simp add:assms)+) show "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms unfolding risk_neutral_prob_def by auto have "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset n `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(7) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms(6) by blast moreover have "(\ y\ prices Mkt asset n `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset n `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset n ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset n w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset n w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset n \ borel_measurable M" using assms readable using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset n w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed show "\n. \asset\support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (intro allI ballI) fix n fix asset assume "asset\support_set pf" show "integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule prob_space.finite_borel_measurable_integrable) show "prob_space N" using assms unfolding risk_neutral_prob_def by auto have "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M)" proof - have "\y\ prices Mkt asset (Suc n) `(space M). finite ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)" by (metis \asset \ support_set pf\ assms(7) finite_imageI image_image) hence "finite (\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y))" using \asset \ support_set pf\ assms(6) by blast moreover have "(\ y\ prices Mkt asset (Suc n) `(space M). ((\ z. (\w. z * pf asset (Suc n) w) ` space M) y)) = (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" by simp moreover have "((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M) \ (\ y\ prices Mkt asset (Suc n) `(space M). (\w. y * pf asset (Suc n) w) ` space M)" proof fix x assume "x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M" show "x \ (\y\prices Mkt asset (Suc n) ` space M. (\w. y * pf asset (Suc n) w) ` space M)" using \x \ (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space M\ by auto qed ultimately show ?thesis by (simp add:finite_subset) qed thus "finite ((\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) ` space N)" using assms by (simp add:filt_equiv_space) have "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable M" proof - have "prices Mkt asset (Suc n) \ borel_measurable M" using assms readable using \asset \ support_set pf\ borel_adapt_stoch_proc_borel_measurable by blast moreover have "pf asset (Suc n) \ borel_measurable M" using assms unfolding trading_strategy_def using \asset \ support_set pf\ borel_predict_stoch_proc_borel_measurable by blast ultimately show ?thesis by simp qed thus "(\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w) \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) qed qed show "stock_portfolio Mkt pf" using assms stock_portfolio_def by (simp add: stock_portfolio_def trading_strategy_def) qed lemma (in rfr_disc_equity_market) replicating_expectation: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "replicating_portfolio pf pyf matur" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" and "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" and "viable_market Mkt" and "sets (F 0) = {{}, space M}" and "pyf \ borel_measurable (F matur)" shows "fair_price Mkt (prob_space.expectation N (discounted_value r (\m. pyf) matur)) pyf matur" proof - have fn: "filtrated_prob_space N F" using assms by (simp add: \pyf \ borel_measurable (F matur)\ filtrated_prob_space_axioms.intro filtrated_prob_space_def risk_neutral_prob_def filt_equiv_filtration) have "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using assms(3) disc_equity_market.replicating_portfolio_def disc_equity_market_axioms discounted_adapted filtrated_prob_space.borel_adapt_stoch_proc_borel_measurable fn cls_val_process_adapted by (metis (no_types, opaque_lifting) support_adapt_def readable stock_portfolio_def subsetCE) have "discounted_value r (\m. pyf) matur \ borel_measurable N" proof - have "(\m. pyf) matur \ borel_measurable (F matur)" using assms by simp hence "(\m. pyf) matur \ borel_measurable M" using filtration filtrationE1 measurable_from_subalg by blast hence "(\m. pyf) matur \ borel_measurable N" using assms by (simp add:filt_equiv_measurable) thus ?thesis by (simp add:discounted_measurable) qed have mpyf: "AE w in M. cls_val_process Mkt pf matur w = pyf w" using assms unfolding replicating_portfolio_def by simp have "AE w in N. cls_val_process Mkt pf matur w = pyf w" proof (rule filt_equiv_borel_AE_eq) show "filt_equiv F M N" using assms by simp show "pyf \ borel_measurable (F matur)" using assms by simp show "AE w in M. cls_val_process Mkt pf matur w = pyf w" using mpyf by simp show "cls_val_process Mkt pf matur \ borel_measurable (F matur)" using assms(3) price_structure_def replicating_price_process by (meson support_adapt_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms readable stock_portfolio_def subsetCE) qed hence disc:"AE w in N. discounted_value r (cls_val_process Mkt pf) matur w = discounted_value r (\m. pyf) matur w" by (simp add:discounted_AE_cong) have "AEeq N (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) (real_cond_exp N (F 0) (discounted_value r (\m. pyf) matur))" proof (rule sigma_finite_subalgebra.real_cond_exp_cong) show "sigma_finite_subalgebra N (F 0)" using filtrated_prob_space.axioms(1) filtrated_prob_space.filtration fn filtrationE1 prob_space.subalgebra_sigma_finite by blast show "AEeq N (discounted_value r (cls_val_process Mkt pf) matur) (discounted_value r (\m. pyf) matur)" using disc by simp show "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using \discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N\ . show "discounted_value r (\m. pyf) matur \ borel_measurable N" using \discounted_value r (\m. pyf) matur \ borel_measurable N\ . qed have "martingale N F (discounted_value r (cls_val_process Mkt pf))" using assms unfolding replicating_portfolio_def using self_fin_trad_strat_mart[of N pf] by (simp add: stock_portfolio_def) hence "AEeq N (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) (discounted_value r (cls_val_process Mkt pf) 0)" using martingaleAE[of N F "discounted_value r (cls_val_process Mkt pf)" 0 matur] fn by simp also have "AE w in N. (discounted_value r (cls_val_process Mkt pf) 0 w) = initial_value pf" proof fix w assume "w\ space N" have "discounted_value r (cls_val_process Mkt pf) 0 w = cls_val_process Mkt pf 0 w" by (simp add:discounted_init) also have "... = val_process Mkt pf 0 w" unfolding cls_val_process_def using assms unfolding replicating_portfolio_def stock_portfolio_def by simp also have "... = initial_value pf" using assms unfolding replicating_portfolio_def using \w\ space N\ by (metis (no_types, lifting) support_adapt_def filt_equiv_space initial_valueI readable stock_portfolio_def subsetCE) finally show "discounted_value r (cls_val_process Mkt pf) 0 w = initial_value pf" . qed finally have "AE w in N. (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) w = initial_value pf" . moreover have "\w\ space N. (real_cond_exp N (F 0) (discounted_value r (cls_val_process Mkt pf) matur)) w = prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur)" proof (rule prob_space.trivial_subalg_cond_expect_eq) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "subalgebra N (F 0)" using \prob_space N\ filtrated_prob_space.filtration fn filtrationE1 by blast show "sets (F 0) = {{}, space N}" using assms by (simp add:filt_equiv_space) show "integrable N (discounted_value r (cls_val_process Mkt pf) matur)" proof (rule discounted_integrable) show "space N = space M" using assms by (simp add:filt_equiv_space) show "integrable N (cls_val_process Mkt pf matur)" using assms unfolding replicating_portfolio_def by (simp add: integrable_self_fin_uvp) show "-1 < r" using acceptable_rate by simp qed qed ultimately have "AE w in N. prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = initial_value pf" by simp hence "prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = initial_value pf" using assms unfolding risk_neutral_prob_def using prob_space.emeasure_space_1[of N] AE_eq_cst[of _ _ N] by simp moreover have "prob_space.expectation N (discounted_value r (cls_val_process Mkt pf) matur) = prob_space.expectation N (discounted_value r (\m. pyf) matur)" proof (rule integral_cong_AE) show "AEeq N (discounted_value r (cls_val_process Mkt pf) matur) (discounted_value r (\m. pyf) matur)" using disc by simp show "discounted_value r (\m. pyf) matur \ borel_measurable N" using \discounted_value r (\m. pyf) matur \ borel_measurable N\ . show "discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N" using \discounted_value r (cls_val_process Mkt pf) matur \ borel_measurable N\ . qed ultimately have "prob_space.expectation N (discounted_value r (\m. pyf) matur) = initial_value pf" by simp thus ?thesis using assms by (metis (full_types) support_adapt_def disc_equity_market.replicating_portfolio_def disc_equity_market_axioms readable replicating_fair_price stock_portfolio_def subsetCE) qed lemma (in rfr_disc_equity_market) replicating_expectation_finite: assumes "risk_neutral_prob N" and "filt_equiv F M N" and "replicating_portfolio pf pyf matur" and "\n. \ asset \ support_set pf. finite (prices Mkt asset n `(space M))" and "\n. \ asset \ support_set pf. finite (pf asset n `(space M))" and "viable_market Mkt" and "sets (F 0) = {{}, space M}" and "pyf \ borel_measurable (F matur)" shows "fair_price Mkt (prob_space.expectation N (discounted_value r (\m. pyf) matur)) pyf matur" proof - have "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset n w * pf asset (Suc n) w)" proof (rule finite_integrable_vp, (auto simp add:assms)) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "\n asset. asset \ support_set pf \ random_variable borel (prices Mkt asset n)" proof- fix n fix asset assume "asset \ support_set pf" show "random_variable borel (prices Mkt asset n)" using assms unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def using readable by (meson \asset \ support_set pf\ adapt_stoch_proc_borel_measurable subsetCE) qed qed moreover have "\n. \ asset \ support_set pf. integrable N (\w. prices Mkt asset (Suc n) w * pf asset (Suc n) w)" proof (rule finite_integrable_uvp, (auto simp add:assms)) show "prob_space N" using assms unfolding risk_neutral_prob_def by simp show "trading_strategy pf" using assms unfolding replicating_portfolio_def by simp show "\n asset. asset \ support_set pf \ random_variable borel (prices Mkt asset n)" proof- fix n fix asset assume "asset \ support_set pf" show "random_variable borel (prices Mkt asset n)" using assms unfolding replicating_portfolio_def stock_portfolio_def adapt_stoch_proc_def using readable by (meson \asset \ support_set pf\ adapt_stoch_proc_borel_measurable subsetCE) qed qed ultimately show ?thesis using assms replicating_expectation by simp qed end \ No newline at end of file diff --git a/thys/Ergodic_Theory/Kingman.thy b/thys/Ergodic_Theory/Kingman.thy --- a/thys/Ergodic_Theory/Kingman.thy +++ b/thys/Ergodic_Theory/Kingman.thy @@ -1,2121 +1,2121 @@ (* Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr License: BSD *) section \Subcocycles, subadditive ergodic theory\ theory Kingman imports Ergodicity Fekete begin text \Subadditive ergodic theory is the branch of ergodic theory devoted to the study of subadditive cocycles (named subcocycles in what follows), i.e., functions such that $u(n+m, x) \leq u(n, x) + u(m, T^n x)$ for all $x$ and $m,n$. For instance, Birkhoff sums are examples of such subadditive cocycles (in fact, they are additive), but more interesting examples are genuinely subadditive. The main result of the theory is Kingman's theorem, asserting the almost sure convergence of $u_n / n$ (this is a generalization of Birkhoff theorem). If the asymptotic average $\lim \int u_n / n$ (which exists by subadditivity and Fekete lemma) is not $-\infty$, then the convergence takes also place in $L^1$. We prove all this below. \ context mpt begin subsection \Definition and basic properties\ definition subcocycle::"(nat \ 'a \ real) \ bool" where "subcocycle u = ((\n. integrable M (u n)) \ (\n m x. u (n+m) x \ u n x + u m ((T^^n) x)))" lemma subcocycle_ineq: assumes "subcocycle u" shows "u (n+m) x \ u n x + u m ((T^^n) x)" using assms unfolding subcocycle_def by blast lemma subcocycle_0_nonneg: assumes "subcocycle u" shows "u 0 x \ 0" proof - have "u (0+0) x \ u 0 x + u 0 ((T^^0) x)" using assms unfolding subcocycle_def by blast then show ?thesis by auto qed lemma subcocycle_integrable: assumes "subcocycle u" shows "integrable M (u n)" "u n \ borel_measurable M" using assms unfolding subcocycle_def by auto lemma subcocycle_birkhoff: assumes "integrable M f" shows "subcocycle (birkhoff_sum f)" unfolding subcocycle_def by (auto simp add: assms birkhoff_sum_integral(1) birkhoff_sum_cocycle) text \The set of subcocycles is stable under addition, multiplication by positive numbers, and $\max$.\ lemma subcocycle_add: assumes "subcocycle u" "subcocycle v" shows "subcocycle (\n x. u n x + v n x)" unfolding subcocycle_def proof (auto) fix n show "integrable M (\x. u n x + v n x)" using assms unfolding subcocycle_def by simp next fix n m x have "u (n+m) x \ u n x + u m ((T ^^ n) x)" using assms(1) subcocycle_def by simp moreover have "v (n+m) x \ v n x + v m ((T ^^ n) x)" using assms(2) subcocycle_def by simp ultimately show "u (n + m) x + v (n + m) x \ u n x + v n x + (u m ((T ^^ n) x) + v m ((T ^^ n) x))" by simp qed lemma subcocycle_cmult: assumes "subcocycle u" "c \ 0" shows "subcocycle (\n x. c * u n x)" using assms unfolding subcocycle_def by (auto, metis distrib_left mult_left_mono) lemma subcocycle_max: assumes "subcocycle u" "subcocycle v" shows "subcocycle (\n x. max (u n x) (v n x))" unfolding subcocycle_def proof (auto) fix n show "integrable M (\x. max (u n x) (v n x))" using assms unfolding subcocycle_def by auto next fix n m x have "u (n+m) x \ u n x + u m ((T^^n) x)" using assms(1) unfolding subcocycle_def by auto then show "u (n + m) x \ max (u n x) (v n x) + max (u m ((T ^^ n) x)) (v m ((T ^^ n) x))" by simp next fix n m x have "v (n+m) x \ v n x + v m ((T^^n) x)" using assms(2) unfolding subcocycle_def by auto then show "v (n + m) x \ max (u n x) (v n x) + max (u m ((T ^^ n) x)) (v m ((T ^^ n) x))" by simp qed text \Applying inductively the subcocycle equation, it follows that a subcocycle is bounded by the Birkhoff sum of the subcocycle at time $1$.\ lemma subcocycle_bounded_by_birkhoff1: assumes "subcocycle u" "n > 0" shows "u n x \ birkhoff_sum (u 1) n x" using \n > 0\ proof (induction rule: ind_from_1) case 1 show ?case by auto next case (Suc p) have "u (Suc p) x \ u p x + u 1 ((T^^p)x)" using assms(1) subcocycle_def by (metis Suc_eq_plus1) then show ?case using Suc birkhoff_sum_cocycle[where ?n = p and ?m = 1] \ p>0 \ by (simp add: birkhoff_sum_def) qed text \It is often important to bound a cocycle $u_n(x)$ by the Birkhoff sums of $u_N/N$. Compared to the trivial upper bound for $u_1$, there are additional boundary errors that make the estimate more cumbersome (but these terms only come from a $N$-neighborhood of $0$ and $n$, so they are negligible if $N$ is fixed and $n$ tends to infinity.\ lemma subcocycle_bounded_by_birkhoffN: assumes "subcocycle u" "n > 2*N" "N>0" shows "u n x \ birkhoff_sum (\x. u N x / real N) (n - 2 * N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\)" proof - have Iar: "u (a*N+r) x \ u r x + (\i u(a*N+r) x + u N ((T^^(a*N+r))x)" using assms(1) unfolding subcocycle_def by auto also have "... \ u r x + (\ii (\i0" for a using that proof (induction rule: ind_from_1) case 1 show ?case by auto next case (Suc a) have "u ((a+1)*N) x = u((a*N) + N) x" by (simp add: semiring_normalization_rules(2) semiring_normalization_rules(23)) also have "... \ u(a*N) x + u N ((T^^(a*N))x)" using assms(1) unfolding subcocycle_def by auto also have "... \ (\iiii<2*N. abs(u 1 ((T^^(n-(2*N-i))) x)))" have "E2 \ 0" unfolding E2_def by auto obtain a0 s0 where 0: "s0 < N" "n = a0 * N + s0" using \0 < N\ mod_div_decomp mod_less_divisor by blast then have "a0 \ 1" using \n > 2 * N\ \N>0\ by (metis Nat.add_0_right add.commute add_lessD1 add_mult_distrib comm_monoid_mult_class.mult_1 eq_imp_le less_imp_add_positive less_imp_le_nat less_one linorder_neqE_nat mult.left_neutral mult_not_zero not_add_less1 one_add_one) define a s where "a = a0-1" and "s = s0+N" then have as: "n = a * N + s" unfolding a_def s_def using \a0 \ 1\ 0 by (simp add: mult_eq_if) have s: "s \ N" "s < 2*N" using 0 unfolding s_def by auto have a: "a*N > n - 2*N" "a*N \ n - N" using as s \n > 2*N\ by auto then have "(a*N - (n-2*N)) \ N" using \n > 2*N\ by auto have "a*N \ n - 2*N" using a by simp { fix r::nat assume "r < N" have "a*N+r > n - 2*N" using \n>2*N\ as s by auto define tr where "tr = n-(a*N+r)" have "tr > 0" unfolding tr_def using as s \r by auto then have *: "n = (a*N+r) + tr" unfolding tr_def by auto have "birkhoff_sum (u 1) tr ((T^^(a*N+r))x) = (\ii\{a*N+r.. (\i\{a*N+r.. (\i\{n-2*N..r tr_def by auto also have "... = E2" unfolding E2_def apply (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - (n-2*N)"]) using \n > 2*N\ by auto finally have A: "birkhoff_sum (u 1) tr ((T^^(a*N+r))x) \ E2" by simp have "u n x \ u (a*N+r) x + u tr ((T^^(a*N+r))x)" using assms(1) * unfolding subcocycle_def by auto also have "... \ u (a*N+r) x + birkhoff_sum (u 1) tr ((T^^(a*N+r))x)" using subcocycle_bounded_by_birkhoff1[OF assms(1)] \tr > 0\ by auto finally have B: "u n x \ u (a*N+r) x + E2" using A by auto have "u (a*N+r) x \ (\ii0" using \a*N+r > n - 2*N\ not_less by fastforce have "u(a*N+r) x \ (\ia>0\] True by auto moreover have "0 \ (\i u r x + (\i (\i (\i (\ir by auto finally show ?thesis using I by auto qed then have "u n x \ E1 + (\i E2" if "j\{n-2*N.. (\iN>0\] unfolding birkhoff_sum_def by auto also have "... = (\i (\ik\{j.. (\i\{n-2*N..j\{n-2*N.. \a*N \ n - N\ by auto also have "... = E2" unfolding E2_def apply (rule sum.reindex_bij_betw[symmetric], rule bij_betw_byWitness[where ?f' = "\i. i - (n-2*N)"]) using \n > 2*N\ by auto finally show ?thesis by auto qed have "(\jjj\{n-2*N..a*N \ n - 2*N\, of 0 "\j. u N ((T^^j) x)", symmetric] atLeast0LessThan by simp also have "... \ (\j\{n-2*N.. N * E2" using \(a*N - (n-2*N)) \ N\ \E2 \ 0\ by (simp add: mult_right_mono) finally have J: "(\j (\jr (\rirrij N *(E1+E2) + (\jjN>0\ by (simp add: sum_distrib_left) also have "... = N*(E1 + E2 + (\j E1 + 2*E2 + birkhoff_sum (\x. u N x / N) (n-2*N) x" unfolding birkhoff_sum_def using \N>0\ by auto then show ?thesis unfolding E1_def E2_def by auto qed text \Many natural cocycles are only defined almost everywhere, and then the subadditivity property only makes sense almost everywhere. We will now show that such an a.e.-subcocycle coincides almost everywhere with a genuine subcocycle in the above sense. Then, all the results for subcocycles will apply to such a.e.-subcocycles. (Usually, in ergodic theory, subcocycles only satisfy the subadditivity property almost everywhere, but we have requested it everywhere for simplicity in the proofs.) The subcocycle will be defined in a recursive way. This means that is can not be defined in a proof (since complicated function definitions are not available inside proofs). Since it is defined in terms of $u$, then $u$ has to be available at the top level, which is most conveniently done using a context. \ context fixes u::"nat \ 'a \ real" assumes H: "\m n. AE x in M. u (n+m) x \ u n x + u m ((T^^n) x)" "\n. integrable M (u n)" begin private fun v :: "nat \ 'a \ real" where "v n x = ( if n = 0 then max (u 0 x) 0 else if n = 1 then u 1 x else min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..v 0 x = max (u 0 x) 0\ by simp private lemma v1 [simp]: \v (Suc 0) x = u 1 x\ by simp private lemma v2 [simp]: \v n x = min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<.. if \n \ 2\ using that by (subst v.simps) (simp del: v.simps) declare v.simps [simp del] private lemma integrable_v: "integrable M (v n)" for n proof (induction n rule: nat_less_induct) case (1 n) consider "n = 0" | "n = 1" | "n>1" by linarith then show ?case proof (cases) assume "n = 0" have "v 0 x = max (u 0 x) 0" for x by simp then show ?thesis using integrable_max[OF H(2)[of 0]] \n = 0\ by auto next assume "n = 1" have "v 1 x = u 1 x" for x by simp then show ?thesis using H(2)[of 1] \n = 1\ by auto next assume "n > 1" hence "v n x = min (u n x) (MIN k \ {0<..x. min (u n x) (MIN k \ {0<..n >1\ apply auto[1] apply (rule Bochner_Integration.integrable_add) using "1.IH" apply auto[1] apply (rule Tn_integral_preserving(1)) using "1.IH" by (metis \1 < n\ diff_less greaterThanLessThan_iff max_0_1(2) max_less_iff_conj) ultimately show ?case by auto qed qed private lemma u_eq_v_AE: "AE x in M. v n x = u n x" for n proof (induction n rule: nat_less_induct) case (1 n) consider "n = 0" | "n = 1" | "n>1" by linarith then show ?case proof (cases) assume "n = 0" have "AE x in M. u 0 x \ u 0 x + u 0 x" using H(1)[of 0 0] by auto then have "AE x in M. u 0 x \ 0" by auto moreover have "v 0 x = max (u 0 x) 0" for x by simp ultimately show ?thesis using \n = 0\ by auto next assume "n = 1" have "v 1 x = u 1 x" for x by simp then show ?thesis using \n = 1\ by simp next assume "n > 1" { fix k assume "ks. v k ((T^^s) x) = u k ((T^^s) x)" by simp } note * = this have "AE x in M. \k \ {..s. v k ((T^^s) x) = u k ((T^^s) x)" apply (rule AE_finite_allI) using * by simp_all moreover have "AE x in M. \i j. u (i+j) x \ u i x + u j ((T^^i) x)" apply (subst AE_all_countable, intro allI)+ using H(1) by simp moreover { fix x assume "\k \ {..s. v k ((T^^s) x) = u k ((T^^s) x)" "\i j. u (i+j) x \ u i x + u j ((T^^i) x)" then have Hx: "\k s. k < n \ v k ((T^^s) x) = u k ((T^^s) x)" "\i j. u (i+j) x \ u i x + u j ((T^^i) x)" by auto { fix k assume "k \ {0<.. u k x + u (n-k) ((T^^k) x)" using Hx(2) K by (metis le_add_diff_inverse less_imp_le_nat) also have "... = v k x + v (n-k) ((T^^k)x)" using Hx(1)[OF \k , of 0] Hx(1)[OF \n-k , of k] by auto finally have "u n x \ v k x + v (n-k) ((T^^k)x)" by simp } then have *: "\z. z \ (\k. v k x + v (n-k) ((T^^k) x))`{0<.. u n x \ z" by blast have "u n x \ Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..n>1\ * by auto moreover have "v n x = min (u n x) (Min ((\k. v k x + v (n-k) ((T^^k) x))`{0<..1 by auto ultimately have "v n x = u n x" by auto } ultimately show ?thesis by auto qed qed private lemma subcocycle_v: "v (n+m) x \ v n x + v m ((T^^n) x)" proof - consider "n = 0" | "m = 0" | "n>0 \ m >0" by auto then show ?thesis proof (cases) case 1 then have "v n x \ 0" by simp then show ?thesis using \n = 0\ by auto next case 2 then have "v m x \ 0" by simp then show ?thesis using \m = 0\ by auto next case 3 then have "n+m > 1" by simp then have "v (n+m) x = min (u(n+m) x) (Min ((\k. v k x + v ((n+m)-k) ((T^^k) x))`{0<.. Min ((\k. v k x + v ((n+m)-k) ((T^^k) x))`{0<.. v n x + v ((n+m)-n) ((T^^n) x)" apply (rule Min_le, simp) by (metis (lifting) \0 < n \ 0 < m\ add.commute greaterThanLessThan_iff image_iff less_add_same_cancel2) finally show ?thesis by simp qed qed lemma subcocycle_AE_in_context: "\w. subcocycle w \ (AE x in M. \n. w n x = u n x)" proof - have "subcocycle v" using subcocycle_v integrable_v unfolding subcocycle_def by auto moreover have "AE x in M. \n. v n x = u n x" by (subst AE_all_countable, intro allI, rule u_eq_v_AE) ultimately show ?thesis by blast qed end lemma subcocycle_AE: fixes u::"nat \ 'a \ real" assumes "\m n. AE x in M. u (n+m) x \ u n x + u m ((T^^n) x)" "\n. integrable M (u n)" shows "\w. subcocycle w \ (AE x in M. \n. w n x = u n x)" using subcocycle_AE_in_context assms by blast subsection \The asymptotic average\ text \In this subsection, we define the asymptotic average of a subcocycle $u$, i.e., the limit of $\int u_n(x)/n$ (the convergence follows from subadditivity of $\int u_n$) and study its basic properties, especially in terms of operations on subcocycles. In general, it can be $-\infty$, so we define it in the extended reals.\ definition subcocycle_avg_ereal::"(nat \ 'a \ real) \ ereal" where "subcocycle_avg_ereal u = Inf {ereal((\x. u n x \M) / n) |n. n > 0}" lemma subcocycle_avg_finite: "subcocycle_avg_ereal u < \" unfolding subcocycle_avg_ereal_def using Inf_less_iff less_ereal.simps(4) by blast lemma subcocycle_avg_subadditive: assumes "subcocycle u" shows "subadditive (\n. (\x. u n x \M))" unfolding subadditive_def proof (intro allI) have int_u [measurable]: "\n. integrable M (u n)" using assms unfolding subcocycle_def by auto fix m n have "(\x. u (n+m) x \M) \ (\x. u n x + u m ((T^^n) x) \M)" apply (rule integral_mono) using int_u apply (auto simp add: Tn_integral_preserving(1)) using assms unfolding subcocycle_def by auto also have "... \ (\x. u n x \M) + (\x. u m ((T^^n) x) \M)" using int_u by (auto simp add: Tn_integral_preserving(1)) also have "... = (\x. u n x \M) + (\x. u m x \M)" using int_u by (auto simp add: Tn_integral_preserving(2)) finally show "(\x. u (n+m) x \M) \ (\x. u n x \M) + (\x. u m x \M)" by simp qed lemma subcocycle_int_tendsto_avg_ereal: assumes "subcocycle u" shows "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" unfolding subcocycle_avg_ereal_def using subadditive_converges_ereal[OF subcocycle_avg_subadditive[OF assms]] by auto text \The average behaves well under addition, scalar multiplication and max, trivially.\ lemma subcocycle_avg_ereal_add: assumes "subcocycle u" "subcocycle v" shows "subcocycle_avg_ereal (\n x. u n x + v n x) = subcocycle_avg_ereal u + subcocycle_avg_ereal v" proof - have int [simp]: "\n. integrable M (u n)" "\n. integrable M (v n)" using assms unfolding subcocycle_def by auto { fix n have "(\x. u n x / n \M) + (\x. v n x / n \M) = (\x. u n x / n + v n x / n \M)" by (rule Bochner_Integration.integral_add[symmetric], auto) also have "... = (\x. (u n x + v n x) / n \M)" by (rule Bochner_Integration.integral_cong, auto simp add: add_divide_distrib) finally have "ereal (\x. u n x / n \M) + (\x. v n x / n \M) = (\x. (u n x + v n x) / n \M)" by auto } moreover have "(\n. ereal (\x. u n x / n \M) + (\x. v n x / n \M)) \ subcocycle_avg_ereal u + subcocycle_avg_ereal v" apply (intro tendsto_intros subcocycle_int_tendsto_avg_ereal[OF assms(1)] subcocycle_int_tendsto_avg_ereal[OF assms(2)]) using subcocycle_avg_finite by auto ultimately have "(\n. (\x. (u n x + v n x) / n \M)) \ subcocycle_avg_ereal u + subcocycle_avg_ereal v" by auto moreover have "(\n. (\x. (u n x + v n x) / n \M)) \ subcocycle_avg_ereal (\n x. u n x + v n x)" by (rule subcocycle_int_tendsto_avg_ereal[OF subcocycle_add[OF assms]]) ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_avg_ereal_cmult: assumes "subcocycle u" "c \ (0::real)" shows "subcocycle_avg_ereal (\n x. c * u n x) = c * subcocycle_avg_ereal u" proof (cases "c = 0") case True have *: "ereal (\x. (c * u n x) / n \M) = 0" if "n>0" for n by (subst True, auto) have "(\n. ereal (\x. (c * u n x) / n \M)) \ 0" by (subst lim_explicit, metis * less_le_trans zero_less_one) moreover have "(\n. ereal (\x. (c * u n x) / n \M)) \ subcocycle_avg_ereal (\n x. c * u n x)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_cmult[OF assms]] by auto ultimately have "subcocycle_avg_ereal (\n x. c * u n x) = 0" using LIMSEQ_unique by blast then show ?thesis using True by auto next case False have int: "\n. integrable M (u n)" using assms unfolding subcocycle_def by auto have "ereal (\x. c * u n x / n \M) = c * ereal (\x. u n x / n \M)" for n by auto then have "(\n. c * ereal (\x. u n x / n \M)) \ subcocycle_avg_ereal (\n x. c * u n x)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_cmult[OF assms]] by auto moreover have "(\n. c * ereal (\x. u n x / n \M)) \ c * subcocycle_avg_ereal u" apply (rule tendsto_mult_ereal) using False subcocycle_int_tendsto_avg_ereal[OF assms(1)] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_avg_ereal_max: assumes "subcocycle u" "subcocycle v" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) \ max (subcocycle_avg_ereal u) (subcocycle_avg_ereal v)" proof (auto) have int: "integrable M (u n)" "integrable M (v n)" for n using assms unfolding subcocycle_def by auto have int2: "integrable M (\x. max (u n x) (v n x))" for n using integrable_max int by auto have "(\x. u n x / n \M) \ (\x. max (u n x) (v n x) / n \M)" for n apply (rule integral_mono) using int int2 by (auto simp add: divide_simps) then show "subcocycle_avg_ereal u \ subcocycle_avg_ereal (\n x. max (u n x) (v n x))" using LIMSEQ_le[OF subcocycle_int_tendsto_avg_ereal[OF assms(1)] subcocycle_int_tendsto_avg_ereal[OF subcocycle_max[OF assms]]] by auto have "(\x. v n x / n \M) \ (\x. max (u n x) (v n x) / n \M)" for n apply (rule integral_mono) using int int2 by (auto simp add: divide_simps) then show "subcocycle_avg_ereal v \ subcocycle_avg_ereal (\n x. max (u n x) (v n x))" using LIMSEQ_le[OF subcocycle_int_tendsto_avg_ereal[OF assms(2)] subcocycle_int_tendsto_avg_ereal[OF subcocycle_max[OF assms]]] by auto qed text \For a Birkhoff sum, the average at each time is the same, equal to the average of the function, so the asymptotic average is also equal to this common value.\ lemma subcocycle_avg_ereal_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) = (\x. u x \M)" proof - have *: "ereal (\x. (birkhoff_sum u n x) / n \M) = (\x. u x \M)" if "n>0" for n using birkhoff_sum_integral(2)[OF assms] that by auto have "(\n. ereal (\x. (birkhoff_sum u n x) / n \M)) \ (\x. u x \M)" by (subst lim_explicit, metis * less_le_trans zero_less_one) moreover have "(\n. ereal (\x. (birkhoff_sum u n x) / n \M)) \ subcocycle_avg_ereal (birkhoff_sum u)" using subcocycle_int_tendsto_avg_ereal[OF subcocycle_birkhoff[OF assms]] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed text \In nice situations, where one can avoid the use of ereal, the following definition is more convenient. The kind of statements we are after is as follows: if the ereal average is finite, then something holds, likely involving the real average. In particular, we show in this setting what we have proved above under this new assumption: convergence (in real numbers) of the average to the asymptotic average, as well as good behavior under sum, scalar multiplication by positive numbers, max, formula for Birkhoff sums.\ definition subcocycle_avg::"(nat \ 'a \ real) \ real" where "subcocycle_avg u = real_of_ereal(subcocycle_avg_ereal u)" lemma subcocycle_avg_real_ereal: assumes "subcocycle_avg_ereal u > - \" shows "subcocycle_avg_ereal u = ereal(subcocycle_avg u)" unfolding subcocycle_avg_def using assms subcocycle_avg_finite[of u] ereal_real by auto lemma subcocycle_int_tendsto_avg: assumes "subcocycle u" "subcocycle_avg_ereal u > - \" shows "(\n. (\x. u n x / n \M)) \ subcocycle_avg u" using subcocycle_avg_real_ereal[OF assms(2)] subcocycle_int_tendsto_avg_ereal[OF assms(1)] by auto lemma subcocycle_avg_add: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > - \" "subcocycle_avg_ereal v > - \" shows "subcocycle_avg_ereal (\n x. u n x + v n x) > -\" "subcocycle_avg (\n x. u n x + v n x) = subcocycle_avg u + subcocycle_avg v" using assms subcocycle_avg_finite real_of_ereal_add unfolding subcocycle_avg_def subcocycle_avg_ereal_add[OF assms(1) assms(2)] by auto lemma subcocycle_avg_cmult: assumes "subcocycle u" "c \ (0::real)" "subcocycle_avg_ereal u > - \" shows "subcocycle_avg_ereal (\n x. c * u n x) > - \" "subcocycle_avg (\n x. c * u n x) = c * subcocycle_avg u" using assms subcocycle_avg_finite unfolding subcocycle_avg_def subcocycle_avg_ereal_cmult[OF assms(1) assms(2)] by auto lemma subcocycle_avg_max: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > - \" "subcocycle_avg_ereal v > - \" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > -\" "subcocycle_avg (\n x. max (u n x) (v n x)) \ max (subcocycle_avg u) (subcocycle_avg v)" proof - show *: "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" using assms(3) subcocycle_avg_ereal_max[OF assms(1) assms(2)] by auto have "ereal (subcocycle_avg (\n x. max (u n x) (v n x))) \ max (ereal(subcocycle_avg u)) (ereal(subcocycle_avg v))" using subcocycle_avg_real_ereal[OF assms(3)] subcocycle_avg_real_ereal[OF assms(4)] subcocycle_avg_real_ereal[OF *] subcocycle_avg_ereal_max[OF assms(1) assms(2)] by auto then show "subcocycle_avg (\n x. max (u n x) (v n x)) \ max (subcocycle_avg u) (subcocycle_avg v)" by auto qed lemma subcocycle_avg_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) > - \" "subcocycle_avg (birkhoff_sum u) = (\x. u x \M)" unfolding subcocycle_avg_def subcocycle_avg_ereal_birkhoff[OF assms(1)] by auto end subsection \Almost sure convergence of subcocycles\ text \In this paragraph, we prove Kingman's theorem, i.e., the almost sure convergence of subcocycles. Their limit is almost surely invariant. There is no really easy proof. The one we use below is arguably the simplest known one, due to Steele (1989). The idea is to show that the limsup of the subcocycle is bounded by the liminf (which is almost surely constant along trajectories), by using subadditivity along time intervals where the liminf is almost reached, of length at most $N$. For some points, the liminf takes a large time $>N$ to be reached. We neglect those times, introducing an additional error that gets smaller with $N$, thanks to Birkhoff ergodic theorem applied to the set of bad points. The error is most easily managed if the subcocycle is assumed to be nonpositive, which one can assume in a first step. The general case is reduced to this one by replacing $u_n$ with $u_n - S_n u_1 \leq 0$, and using Birkhoff theorem to control $S_n u_1$.\ context fmpt begin text \First, as explained above, we prove the theorem for nonpositive subcocycles.\ lemma kingman_theorem_AE_aux1: assumes "subcocycle u" "\x. u 1 x \ 0" shows "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" proof - define l where "l = (\x. liminf (\n. u n x / n))" have u_meas [measurable]: "\n. u n \ borel_measurable M" using assms(1) unfolding subcocycle_def by auto have l_meas [measurable]: "l \ borel_measurable M" unfolding l_def by auto { fix x assume *: "(\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" then have "(\n. birkhoff_sum (u 1) n x / n) \ ereal(real_cond_exp M Invariants (u 1) x)" by auto then have a: "liminf (\n. birkhoff_sum (u 1) n x / n) = ereal(real_cond_exp M Invariants (u 1) x)" using lim_imp_Liminf by force have "ereal(u n x / n) \ ereal(birkhoff_sum (u 1) n x / n)" if "n>0" for n using subcocycle_bounded_by_birkhoff1[OF assms(1) that, of x] that by (simp add: divide_right_mono) with eventually_mono[OF eventually_gt_at_top[of 0] this] have "eventually (\n. ereal(u n x / n) \ ereal(birkhoff_sum (u 1) n x / n)) sequentially" by auto then have "liminf (\n. u n x / n) \ liminf (\n. birkhoff_sum (u 1) n x / n)" by (simp add: Liminf_mono) then have "l x < \" unfolding l_def using a by auto } then have "AE x in M. l x < \" using birkhoff_theorem_AE_nonergodic[of "u 1"] subcocycle_def assms(1) by auto have l_dec: "l x \ l (T x)" for x proof - have "l x = liminf (\n. ereal ((u (n+1) x)/(n+1)))" unfolding l_def by (rule liminf_shift[of "\n. ereal (u n x / real n)", symmetric]) also have "... \ liminf (\n. ereal((u 1 x)/(n+1)) + ereal((u n (T x))/(n+1)))" proof (rule Liminf_mono[OF eventuallyI]) fix n have "u (1+n) x \ u 1 x + u n ((T^^1) x)" using assms(1) unfolding subcocycle_def by blast then have "u (n+1) x \ u 1 x + u n (T x)" by auto then have "(u (n+1) x)/(n+1) \ (u 1 x)/(n+1) + (u n (T x))/(n+1)" by (metis add_divide_distrib divide_right_mono of_nat_0_le_iff) then show "ereal ((u (n+1) x)/(n+1)) \ ereal((u 1 x)/(n+1)) + ereal((u n (T x))/(n+1))" by auto qed also have "... = 0 + liminf(\n. ereal((u n (T x))/(n+1)))" proof (rule ereal_liminf_lim_add[of "\n. ereal((u 1 x)/real(n+1))" 0 "\n. ereal((u n (T x))/(n+1))"]) have "(\n. ereal((u 1 x)*(1/real(n+1)))) \ ereal((u 1 x) * 0)" by (intro tendsto_intros LIMSEQ_ignore_initial_segment) then show "(\n. ereal((u 1 x)/real(n+1))) \ 0" by (simp add: zero_ereal_def) qed (simp) also have "... = 1 * liminf(\n. ereal((u n (T x))/(n+1)))" by simp also have "... = liminf(\n. (n+1)/n * ereal((u n (T x))/(n+1)))" proof (rule ereal_liminf_lim_mult[symmetric]) have "real (n+1) / real n = 1 + 1/real n" if "n>0" for n by (simp add: divide_simps mult.commute that) with eventually_mono[OF eventually_gt_at_top[of "0::nat"] this] have "eventually (\n. real (n+1) / real n = 1 + 1/real n) sequentially" by simp moreover have "(\n. 1 + 1/real n) \ 1 + 0" by (intro tendsto_intros) ultimately have "(\n. real (n+1) / real n) \ 1" using Lim_transform_eventually by (simp add: filterlim_cong) then show "(\n. ereal(real (n+1) / real n)) \ 1" by (simp add: one_ereal_def) qed (auto) also have "... = l (T x)" unfolding l_def by auto finally show "l x \ l (T x)" by simp qed have "AE x in M. l (T x) = l x" apply (rule AE_increasing_then_invariant) using l_dec by auto then obtain g0 where g0: "g0 \ borel_measurable Invariants" "AE x in M. l x = g0 x" using Invariants_quasi_Invariants_functions[OF l_meas] by auto define g where "g = (\x. if g0 x = \ then 0 else g0 x)" have g: "g \ borel_measurable Invariants" "AE x in M. g x = l x" unfolding g_def using g0(1) \AE x in M. l x = g0 x\ \AE x in M. l x < \\ by auto have [measurable]: "g \ borel_measurable M" using g(1) Invariants_measurable_func by blast have "\x. g x < \" unfolding g_def by auto define A where "A = {x \ space M. l x < \ \ (\n. l ((T^^n) x) = g ((T^^n) x))}" have A_meas [measurable]: "A \ sets M" unfolding A_def by auto have "AE x in M. x \ A" unfolding A_def using T_AE_iterates[OF g(2)] \AE x in M. l x < \\ by auto then have "space M - A \ null_sets M" by (simp add: AE_iff_null set_diff_eq) have l_inv: "l((T^^n) x) = l x" if "x \ A" for x n proof - have "l((T^^n) x) = g((T^^n) x)" using \ x \ A \ unfolding A_def by blast also have "... = g x" using g(1) A_def Invariants_func_is_invariant_n that by blast also have "... = g((T^^0) x)" by auto also have "... = l((T^^0) x)" using \ x \ A \ unfolding A_def by (metis (mono_tags, lifting) mem_Collect_eq) finally show ?thesis by auto qed define F where "F = (\ K e x. real_of_ereal(max (l x) (-ereal K)) + e)" have F_meas [measurable]: "F K e \ borel_measurable M" for K e unfolding F_def by auto define B where "B = (\N K e. {x \ A. \n\{1..N}. u n x - n * F K e x < 0})" have B_meas [measurable]: "B N K e \ sets M" for N K e unfolding B_def by (measurable) define I where "I = (\N K e x. (indicator (- B N K e) x)::real)" have I_meas [measurable]: "I N K e \ borel_measurable M" for N K e unfolding I_def by auto have I_int: "integrable M (I N K e)" for N K e - unfolding I_def apply (subst integrable_cong[where ?g = "indicator (space M - B N K e)::_ \ real"], auto) + unfolding I_def apply (subst Bochner_Integration.integrable_cong[where ?g = "indicator (space M - B N K e)::_ \ real"], auto) by (auto split: split_indicator simp: less_top[symmetric]) have main: "AE x in M. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" if "N>(1::nat)" "K>(0::real)" "e>(0::real)" for N K e proof - let ?B = "B N K e" and ?I = "I N K e" and ?F = "F K e" define t where "t = (\x. if x \ ?B then Min {n\{1..N}. u n x - n * ?F x < 0} else 1)" have [measurable]: "t \ measurable M (count_space UNIV)" unfolding t_def by measurable have t1: "t x \ {1..N}" for x proof (cases "x \ ?B") case False then have "t x = 1" by (simp add: t_def) then show ?thesis using \N>1\by auto next case True let ?A = "{n\{1..N}. u n x - n * ?F x < 0}" have "t x = Min ?A" using True by (simp add: t_def) moreover have "Min ?A \ ?A" apply (rule Min_in, simp) using True B_def by blast ultimately show ?thesis by auto qed have bound1: "u (t x) x \ t x * ?F x + birkhoff_sum ?I (t x) x * abs(?F x)" for x proof (cases "x \ ?B") case True let ?A = "{n\{1..N}. u n x - n * F K e x < 0}" have "t x = Min ?A" using True by (simp add: t_def) moreover have "Min ?A \ ?A" apply (rule Min_in, simp) using True B_def by blast ultimately have "u (t x) x \ (t x) * ?F x" by auto moreover have "0 \ birkhoff_sum ?I (t x) x * abs(?F x)" unfolding birkhoff_sum_def I_def by (simp add: sum_nonneg) ultimately show ?thesis by auto next case False then have "0 \ ?F x + ?I x * abs(?F x)" unfolding I_def by auto then have "u 1 x \ ?F x + ?I x * abs(?F x)" using assms(2)[of x] by auto moreover have "t x = 1" unfolding t_def using False by auto ultimately show ?thesis by auto qed define TB where "TB = (\x. (T^^(t x)) x)" have [measurable]: "TB \ measurable M M" unfolding TB_def by auto define S where "S = (\n x. (\i measurable M (count_space UNIV)" for n unfolding S_def by measurable have TB_pow: "(TB^^n) x = (T^^(S n x)) x" for n x unfolding S_def TB_def by (induction n, auto, metis (mono_tags, lifting) add.commute funpow_add o_apply) have uS: "u (S n x) x \ (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x)" if "x \ A" "n>0" for x n using \n > 0\ proof (induction rule: ind_from_1) case 1 show ?case unfolding S_def using bound1 by auto next case (Suc n) have *: "?F((TB^^n) x) = ?F x" apply (subst TB_pow) unfolding F_def using l_inv[OF \x\A\] by auto have **: "S n x + t ((TB^^n) x) = S (Suc n) x" unfolding S_def by auto have "u (S (Suc n) x) x = u (S n x + t((TB^^n) x)) x" unfolding S_def by auto also have "... \ u (S n x) x + u (t((TB^^n) x)) ((T^^(S n x)) x)" using assms(1) unfolding subcocycle_def by auto also have "... \ u (S n x) x + u (t((TB^^n) x)) ((TB^^n) x)" using TB_pow by auto also have "... \ (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x) + t ((TB^^n) x) * ?F ((TB^^n) x) + birkhoff_sum ?I (t ((TB^^n) x)) ((TB^^n) x) * abs(?F ((TB^^n) x))" using Suc bound1[of "((TB^^n) x)"] by auto also have "... = (S n x) * ?F x + birkhoff_sum ?I (S n x) x * abs(?F x) + t ((TB^^n) x) * ?F x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x) * abs(?F x)" using * TB_pow by auto also have "... = (real(S n x) + t ((TB^^n) x)) * ?F x + (birkhoff_sum ?I (S n x) x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x)) * abs(?F x)" by (simp add: mult.commute ring_class.ring_distribs(1)) also have "... = (S n x + t ((TB^^n) x)) * ?F x + (birkhoff_sum ?I (S n x) x + birkhoff_sum ?I (t ((TB^^n) x)) ((T^^(S n x)) x)) * abs(?F x)" by simp also have "... = (S (Suc n) x) * ?F x + birkhoff_sum ?I (S (Suc n) x) x * abs(?F x)" by (subst birkhoff_sum_cocycle[symmetric], subst **, subst **, simp) finally show ?case by simp qed have un: "u n x \ n * ?F x + N * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" if "x \ A" "n>N" for x n proof - let ?A = "{i. S i x > n}" let ?iA = "Inf ?A" have "n < (\i S (n+1) x" unfolding S_def apply (rule sum_mono) using t1 by auto finally have "?A \ {}" by blast then have "?iA \ ?A" by (meson Inf_nat_def1) moreover have "0 \ ?A" unfolding S_def by auto ultimately have "?iA \ 0" by fastforce define j where "j = ?iA - 1" then have "j < ?iA" using \?iA \ 0\ by auto then have "j \ ?A" by (meson bdd_below_def cInf_lower le0 not_less) then have "S j x \ n" by auto define k where "k = n - S j x" have "n = S j x + k" unfolding k_def using \S j x \ n\ by auto have "n < S (j+1) x" unfolding j_def using \?iA \ 0\ \?iA \ ?A\ by auto also have "... = S j x + t((TB^^j) x)" unfolding S_def by auto also have "... \ S j x + N" using t1 by auto finally have "k \ N" unfolding k_def using \n > N\ by auto then have "S j x > 0" unfolding k_def using \n > N\ by auto then have "j > 0" unfolding S_def using not_gr0 by fastforce have "birkhoff_sum ?I (S j x) x \ birkhoff_sum ?I n x" unfolding birkhoff_sum_def I_def using \S j x \ n\ by (metis finite_Collect_less_nat indicator_pos_le lessThan_def lessThan_subset_iff sum_mono2) have "u n x \ u (S j x) x" proof (cases "k = 0") case True show ?thesis using True unfolding k_def using \S j x \ n\ by auto next case False then have "k > 0" by simp have "u k ((T^^(S j x)) x) \ birkhoff_sum (u 1) k ((T ^^ S j x) x)" using subcocycle_bounded_by_birkhoff1[OF assms(1) \k>0\, of "(T^^(S j x)) x"] by simp also have "... \ 0" unfolding birkhoff_sum_def using sum_mono assms(2) by (simp add: sum_nonpos) also have "u n x \ u (S j x) x + u k ((T^^(S j x)) x)" apply (subst \n = S j x + k\) using assms(1) subcocycle_def by auto ultimately show ?thesis by auto qed also have "... \ (S j x) * ?F x + birkhoff_sum ?I (S j x) x * abs(?F x)" using uS[OF \x \ A\ \j>0\] by simp also have "... \ (S j x) * ?F x + birkhoff_sum ?I n x * abs(?F x)" using \birkhoff_sum ?I (S j x) x \ birkhoff_sum ?I n x\ by (simp add: mult_right_mono) also have "... = n * ?F x - k * ?F x + birkhoff_sum ?I n x * abs(?F x)" by (metis \n = S j x + k\ add_diff_cancel_right' le_add2 left_diff_distrib' of_nat_diff) also have "... \ n * ?F x + k * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" by (auto, metis abs_ge_minus_self abs_mult abs_of_nat) also have "... \ n * ?F x + N * abs(?F x) + birkhoff_sum ?I n x * abs(?F x)" using \k \ N\ by (simp add: mult_right_mono) finally show ?thesis by simp qed have "limsup (\n. u n x / n) \ ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" if "x \ A" for x proof - have "(\n. ereal(?F x + N * abs(?F x) * (1 / n))) \ ereal(?F x + N * abs (?F x) * 0)" by (intro tendsto_intros) then have *: "limsup (\n. ?F x + N * abs(?F x)/n) = ?F x" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by force { fix n assume "n > N" have "u n x / real n \ ?F x + N * abs(?F x) / n + abs(?F x) * birkhoff_sum ?I n x / n" using un[OF \x \ A\ \n > N\] using \n>N\ by (auto simp add: divide_simps mult.commute) then have "ereal(u n x/n) \ ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n)" by auto } then have "eventually (\n. ereal(u n x / n) \ ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n)) sequentially" using eventually_mono[OF eventually_gt_at_top[of N]] by auto with Limsup_mono[OF this] have "limsup (\n. u n x / n) \ limsup (\n. ereal(?F x + N * abs(?F x) / n) + abs(?F x) * ereal(birkhoff_sum ?I n x / n))" by auto also have "... \ limsup (\n. ?F x + N * abs(?F x) / n) + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" by (rule ereal_limsup_add_mono) also have "... = ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" using * by auto finally show ?thesis by auto qed then have *: "AE x in M. limsup (\n. u n x / n) \ ?F x + limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n))" using \AE x in M. x \ A\ by auto { fix x assume H: "(\n. birkhoff_sum ?I n x / n) \ real_cond_exp M Invariants ?I x" have "(\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) \ abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" by (rule tendsto_mult_ereal, auto simp add: H) then have "limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) = abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast } moreover have "AE x in M. (\n. birkhoff_sum ?I n x / n) \ real_cond_exp M Invariants ?I x" by (rule birkhoff_theorem_AE_nonergodic[OF I_int]) ultimately have "AE x in M. limsup (\n. abs(?F x) * ereal(birkhoff_sum ?I n x / n)) = abs(?F x) * ereal(real_cond_exp M Invariants ?I x)" by auto then show ?thesis using * by auto qed have bound2: "AE x in M. limsup (\n. u n x / n) \ F K e x" if "K > 0" "e > 0" for K e proof - define C where "C = (\N. A - B N K e)" have C_meas [measurable]: "\N. C N \ sets M" unfolding C_def by auto { fix x assume "x \ A" have "F K e x > l x" using \x \ A\ \e > 0\ unfolding F_def A_def by (cases "l x", auto, metis add.commute ereal_max less_add_same_cancel2 max_less_iff_conj real_of_ereal.simps(1)) then have "\n>0. ereal(u n x / n) < F K e x" unfolding l_def using liminf_upper_bound by fastforce then obtain n where "n>0" "ereal(u n x / n) < F K e x" by auto then have "u n x - n * F K e x < 0" by (simp add: divide_less_eq mult.commute) then have "x \ C n" unfolding C_def B_def using \x \ A\ \n>0\ by auto then have "x \ (\n. C n)" by auto } then have "(\n. C n) = {}" unfolding C_def by auto then have *: "0 = measure M (\n. C n)" by auto have "(\n. measure M (C n)) \ 0" apply (subst *, rule finite_Lim_measure_decseq, auto) unfolding C_def B_def decseq_def by auto moreover have "measure M (C n) = (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)" for n proof - have *: "AE x in M. 0 \ real_cond_exp M Invariants (I n K e) x" apply (rule real_cond_exp_pos, auto) unfolding I_def by auto have "measure M (C n) = (\x. indicator (C n) x \M)" by auto also have "... = (\x. I n K e x \M)" apply (rule integral_cong_AE, auto) unfolding C_def I_def indicator_def using \AE x in M. x \ A\ by auto also have "... = (\x. real_cond_exp M Invariants (I n K e) x \M)" by (rule real_cond_exp_int(2)[symmetric, OF I_int]) also have "... = (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)" apply (rule integral_cong_AE, auto) using * by auto finally show ?thesis by auto qed ultimately have *: "(\n. (\x. norm(real_cond_exp M Invariants (I n K e) x) \M)) \ 0" by simp have "\r. strict_mono r \ (AE x in M. (\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0)" apply (rule tendsto_L1_AE_subseq) using * real_cond_exp_int[OF I_int] by auto then obtain r where "strict_mono r" "AE x in M. (\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0" by auto moreover have "AE x in M. \N \ {1<..}. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" apply (rule AE_ball_countable') using main[OF _ \K>0\ \e>0\] by auto moreover { fix x assume H: "(\n. real_cond_exp M Invariants (I (r n) K e) x) \ 0" "\N. N > 1 \ limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I N K e) x)" have 1: "eventually (\N. limsup (\n. u n x / n) \ F K e x + abs(F K e x) * ereal(real_cond_exp M Invariants (I (r N) K e) x)) sequentially" apply (rule eventually_mono[OF eventually_gt_at_top[of 1] H(2)]) using \strict_mono r\ less_le_trans seq_suble by blast have 2: "(\N. F K e x + (abs(F K e x) * ereal(real_cond_exp M Invariants (I (r N) K e) x))) \ ereal(F K e x) + (abs(F K e x) * ereal 0)" by (intro tendsto_intros) (auto simp add: H(1)) have "limsup (\n. u n x / n) \ F K e x" apply (rule LIMSEQ_le_const) using 1 2 by (auto simp add: eventually_at_top_linorder) } ultimately show "AE x in M. limsup (\n. u n x / n) \ F K e x" by auto qed have "AE x in M. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" if "K>(0::nat)" for K apply (rule AE_upper_bound_inf_ereal) using bound2 \K>0\ unfolding F_def by auto then have "AE x in M. \K\{(0::nat)<..}. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" by (rule AE_ball_countable', auto) moreover have "(\n. u n x / n) \ l x" if H: "\K\{(0::nat)<..}. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))" for x proof - have "limsup (\n. u n x / n) \ l x" proof (cases "l x = \") case False then have "(\K. real_of_ereal(max (l x) (-ereal K))) \ l x" using ereal_truncation_real_bottom by auto moreover have "eventually (\K. limsup (\n. u n x / n) \ real_of_ereal(max (l x) (-ereal K))) sequentially" using H by (metis (no_types, lifting) eventually_at_top_linorder eventually_gt_at_top greaterThan_iff) ultimately show "limsup (\n. u n x / n) \ l x" using Lim_bounded2 eventually_sequentially by auto qed (simp) then have "limsup (\n. ereal (u n x / real n)) = l x" using Liminf_le_Limsup l_def eq_iff sequentially_bot by blast then show "(\n. u n x / n) \ l x" by (simp add: l_def tendsto_iff_Liminf_eq_Limsup) qed ultimately have "AE x in M. (\n. u n x / n) \ l x" by auto then have "AE x in M. (\n. u n x / n) \ g x" using g(2) by auto then show "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" using g(1) \\x. g x < \\ by auto qed text \We deduce it for general subcocycles, by reducing to nonpositive subcocycles by subtracting the Birkhoff sum of $u_1$ (for which the convergence follows from Birkhoff theorem).\ theorem kingman_theorem_AE_aux2: assumes "subcocycle u" shows "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" proof - define v where "v = (\n x. u n x + birkhoff_sum (\x. - u 1 x) n x)" have "subcocycle v" unfolding v_def apply (rule subcocycle_add[OF assms], rule subcocycle_birkhoff) using assms unfolding subcocycle_def by auto have "\(gv::'a\ereal). (gv\borel_measurable Invariants \ (\x. gv x < \) \ (AE x in M. (\n. v n x / n) \ gv x))" apply (rule kingman_theorem_AE_aux1[OF \subcocycle v\]) unfolding v_def by auto then obtain gv where gv: "gv \ borel_measurable Invariants" "AE x in M. (\n. v n x / n) \ (gv x::ereal)" "\x. gv x < \" by blast define g where "g = (\x. gv x + ereal(real_cond_exp M Invariants (u 1) x))" have g_meas: "g \ borel_measurable Invariants" unfolding g_def using gv(1) by auto have g_fin: "\x. g x < \" unfolding g_def using gv(3) by auto have "AE x in M. (\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" apply (rule birkhoff_theorem_AE_nonergodic) using assms unfolding subcocycle_def by auto moreover { fix x assume H: "(\n. v n x / n) \ (gv x)" "(\n. birkhoff_sum (u 1) n x / n) \ real_cond_exp M Invariants (u 1) x" then have "(\n. ereal(birkhoff_sum (u 1) n x / n)) \ ereal(real_cond_exp M Invariants (u 1) x)" by auto { fix n have "u n x = v n x + birkhoff_sum (u 1) n x" unfolding v_def birkhoff_sum_def apply auto by (simp add: sum_negf) then have "u n x / n = v n x / n + birkhoff_sum (u 1) n x / n" by (simp add: add_divide_distrib) then have "ereal(u n x / n) = ereal(v n x / n) + ereal(birkhoff_sum (u 1) n x / n)" by auto } note * = this have "(\n. ereal(u n x / n)) \ g x" unfolding * g_def apply (intro tendsto_intros) using H by auto } ultimately have "AE x in M. (\n. ereal(u n x / n)) \ g x" using gv(2) by auto then show ?thesis using g_meas g_fin by blast qed text \For applications, it is convenient to have a limit which is really measurable with respect to the invariant sigma algebra and does not come from a hard to use abstract existence statement. Hence we introduce the following definition for the would-be limit -- Kingman's theorem shows that it is indeed a limit. We introduce the definition for any function, not only subcocycles, but it will only be usable for subcocycles. We introduce an if clause in the definition so that the limit is always measurable, even when $u$ is not a subcocycle and there is no convergence.\ definition subcocycle_lim_ereal::"(nat \ 'a \ real) \ ('a \ ereal)" where "subcocycle_lim_ereal u = ( if (\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))) then (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x)) else (\_. 0))" definition subcocycle_lim::"(nat \ 'a \ real) \ ('a \ real)" where "subcocycle_lim u = (\x. real_of_ereal(subcocycle_lim_ereal u x))" lemma subcocycle_lim_meas_Inv [measurable]: "subcocycle_lim_ereal u \ borel_measurable Invariants" "subcocycle_lim u \ borel_measurable Invariants" proof - show "subcocycle_lim_ereal u \ borel_measurable Invariants" proof (cases "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))") case True then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF True] by auto next case False then have "subcocycle_lim_ereal u = (\_. 0)" unfolding subcocycle_lim_ereal_def by auto then show ?thesis by auto qed then show "subcocycle_lim u \ borel_measurable Invariants" unfolding subcocycle_lim_def by auto qed lemma subcocycle_lim_meas [measurable]: "subcocycle_lim_ereal u \ borel_measurable M" "subcocycle_lim u \ borel_measurable M" using subcocycle_lim_meas_Inv Invariants_measurable_func apply blast using subcocycle_lim_meas_Inv Invariants_measurable_func by blast lemma subcocycle_lim_ereal_not_PInf: "subcocycle_lim_ereal u x < \" proof (cases "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))") case True then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF True] by auto next case False then have "subcocycle_lim_ereal u = (\_. 0)" unfolding subcocycle_lim_ereal_def by auto then show ?thesis by auto qed text \We reformulate the subadditive ergodic theorem of Kingman with this definition. From this point on, the technical definition of \verb+subcocycle_lim_ereal+ will never be used, only the following property will be relevant.\ theorem kingman_theorem_AE_nonergodic_ereal: assumes "subcocycle u" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" proof - have *: "\(g::'a\ereal). (g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" using kingman_theorem_AE_aux2[OF assms] by auto then have "subcocycle_lim_ereal u = (SOME (g::'a\ereal). g\borel_measurable Invariants \ (\x. g x < \) \ (AE x in M. (\n. u n x / n) \ g x))" unfolding subcocycle_lim_ereal_def by auto then show ?thesis using someI_ex[OF *] by auto qed text \The subcocycle limit behaves well under addition, multiplication by a positive scalar, max, and is simply the conditional expectation with respect to invariants for Birkhoff sums, thanks to Birkhoff theorem.\ lemma subcocycle_lim_ereal_add: assumes "subcocycle u" "subcocycle v" shows "AE x in M. subcocycle_lim_ereal (\n x. u n x + v n x) x = subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" proof - have "AE x in M. (\n. (u n x + v n x)/n) \ subcocycle_lim_ereal (\n x. u n x + v n x) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_add[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim_ereal v x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(2)]) moreover { fix x assume H: "(\n. (u n x + v n x)/n) \ subcocycle_lim_ereal (\n x. u n x + v n x) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. v n x / n) \ subcocycle_lim_ereal v x" have *: "(u n x + v n x)/n = ereal (u n x / n) + (v n x / n)" for n by (simp add: add_divide_distrib) have "(\n. (u n x + v n x)/n) \ subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" unfolding * apply (intro tendsto_intros H(2) H(3)) using subcocycle_lim_ereal_not_PInf by auto then have "subcocycle_lim_ereal (\n x. u n x + v n x) x = subcocycle_lim_ereal u x + subcocycle_lim_ereal v x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_cmult: assumes "subcocycle u" "c\(0::real)" shows "AE x in M. subcocycle_lim_ereal (\n x. c * u n x) x = c * subcocycle_lim_ereal u x" proof - have "AE x in M. (\n. (c * u n x)/n) \ subcocycle_lim_ereal (\n x. c * u n x) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_cmult[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover { fix x assume H: "(\n. (c * u n x)/n) \ subcocycle_lim_ereal (\n x. c * u n x) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" have "(\n. c * ereal (u n x / n)) \ c * subcocycle_lim_ereal u x" by (rule tendsto_cmult_ereal[OF _ H(2)], auto) then have "subcocycle_lim_ereal (\n x. c * u n x) x = c * subcocycle_lim_ereal u x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_max: assumes "subcocycle u" "subcocycle v" shows "AE x in M. subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" proof - have "AE x in M. (\n. max (u n x) (v n x) / n) \ subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_max[OF assms]]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim_ereal v x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(2)]) moreover { fix x assume H: "(\n. max (u n x) (v n x) / n) \ subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. v n x / n) \ subcocycle_lim_ereal v x" have "(\n. max (ereal(u n x / n)) (ereal(v n x / n))) \ max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" apply (rule tendsto_max) using H by auto moreover have "max (ereal(u n x / n)) (ereal(v n x / n)) = max (u n x) (v n x) / n" for n by (simp del: ereal_max add:ereal_max[symmetric] max_divide_distrib_right) ultimately have "(\n. max (u n x) (v n x) / n) \ max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" by auto then have "subcocycle_lim_ereal (\n x. max (u n x) (v n x)) x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" using H(1) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed lemma subcocycle_lim_ereal_birkhoff: assumes "integrable M u" shows "AE x in M. subcocycle_lim_ereal (birkhoff_sum u) x = ereal(real_cond_exp M Invariants u x)" proof - have "AE x in M. (\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" by (rule birkhoff_theorem_AE_nonergodic[OF assms]) moreover have "AE x in M. (\n. birkhoff_sum u n x / n) \ subcocycle_lim_ereal (birkhoff_sum u) x" by (rule kingman_theorem_AE_nonergodic_ereal[OF subcocycle_birkhoff[OF assms]]) moreover { fix x assume H: "(\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" "(\n. birkhoff_sum u n x / n) \ subcocycle_lim_ereal (birkhoff_sum u) x" have "(\n. birkhoff_sum u n x / n) \ ereal(real_cond_exp M Invariants u x)" using H(1) by auto then have "subcocycle_lim_ereal (birkhoff_sum u) x = ereal(real_cond_exp M Invariants u x)" using H(2) by (simp add: LIMSEQ_unique) } ultimately show ?thesis by auto qed subsection \$L^1$ and a.e.\ convergence of subcocycles with finite asymptotic average\ text \In this subsection, we show that the almost sure convergence in Kingman theorem also takes place in $L^1$ if the limit is integrable, i.e., if the asymptotic average of the subcocycle is $> -\infty$. To deduce it from the almost sure convergence, we only need to show that there is no loss of mass, i.e., that the integral of the limit is not strictly larger than the limit of the integrals (thanks to Scheffe criterion). This follows from comparison to Birkhoff sums, for which we know that the average of the limit is the same as the average of the function.\ text \First, we show that the subcocycle limit is bounded by the limit of the Birkhoff sums of $u_N$, i.e., its conditional expectation. This follows from the fact that $u_n$ is bounded by the Birkhoff sum of $u_N$ (up to negligible boundary terms).\ lemma subcocycle_lim_ereal_atmost_uN_invariants: assumes "subcocycle u" "N>(0::nat)" shows "AE x in M. subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" proof - have "AE x in M. (\n. u 1 ((T^^n) x) / n) \ 0" apply (rule limit_foTn_over_n') using assms(1) unfolding subcocycle_def by auto moreover have "AE x in M. (\n. birkhoff_sum (\x. u N x/N) n x / n) \ real_cond_exp M Invariants (\x. u N x / N) x" apply (rule birkhoff_theorem_AE_nonergodic) using assms(1) unfolding subcocycle_def by auto moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim_ereal u x" by (rule kingman_theorem_AE_nonergodic_ereal[OF assms(1)]) moreover { fix x assume H: "(\n. u 1 ((T^^n) x) / n) \ 0" "(\n. birkhoff_sum (\x. u N x/N) n x / n) \ real_cond_exp M Invariants (\x. u N x / N) x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" let ?f = "\n. birkhoff_sum (\x. u N x / real N) (n - 2 * N) x / n + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\ / n)" { fix n assume "n\2*N+1" then have "n > 2 * N" by simp have "u n x / n \ (birkhoff_sum (\x. u N x / real N) (n - 2 * N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. \u 1 ((T ^^ (n - (2 * N - i))) x)\)) / n" using subcocycle_bounded_by_birkhoffN[OF assms(1) \n>2*N\ \N>0\, of x] \n>2*N\ by (simp add: divide_right_mono) also have "... = ?f n" apply (subst add_divide_distrib)+ by (auto simp add: sum_divide_distrib[symmetric]) finally have "u n x / n \ ?f n" by simp then have "u n x / n \ ereal(?f n)" by simp } have "(\n. ?f n) \ real_cond_exp M Invariants (\x. u N x / N) x + (\iu 1 ((T ^^ i) x)\) + 2 * (\i<2*N. 0)" apply (intro tendsto_intros) using H(2) tendsto_norm[OF H(1)] by auto then have "(\n. ereal(?f n)) \ real_cond_exp M Invariants (\x. u N x / N) x" by auto with lim_mono[OF \\n. n \ 2*N+1 \ u n x / n \ ereal(?f n)\ H(3) this] have "subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by simp } ultimately show ?thesis by auto qed text \To apply Scheffe criterion, we need to deal with nonnegative functions, or equivalently with nonpositive functions after a change of sign. Hence, as in the proof of the almost sure version of Kingman theorem above, we first give the proof assuming that the subcocycle is nonpositive, and deduce the general statement by adding a suitable Birkhoff sum.\ lemma kingman_theorem_L1_aux: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" "\x. u 1 x \ 0" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "integrable M (subcocycle_lim u)" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof - have int_u [measurable]: "\n. integrable M (u n)" using assms(1) subcocycle_def by auto then have int_F [measurable]: "\n. integrable M (\x. - u n x/ n)" by auto have F_pos: "- u n x / n \ 0" for x n proof (cases "n > 0") case True have "u n x \ (\in>0\] unfolding birkhoff_sum_def by simp also have "... \ 0" using sum_mono[OF assms(3)] by auto finally have "u n x \ 0" by simp then have "-u n x \ 0" by simp with divide_nonneg_nonneg[OF this] show "- u n x / n \ 0" using \n>0\ by auto qed (auto) { fix x assume *: "(\n. u n x / n) \ subcocycle_lim_ereal u x" have H: "(\n. - u n x / n) \ - subcocycle_lim_ereal u x" using tendsto_cmult_ereal[OF _ *, of "-1"] by auto have "liminf (\n. -u n x / n) = - subcocycle_lim_ereal u x" "(\n. - u n x / n) \ - subcocycle_lim_ereal u x" "- subcocycle_lim_ereal u x \ 0" using H apply (simp add: tendsto_iff_Liminf_eq_Limsup, simp) apply (rule LIMSEQ_le_const[OF H]) using F_pos by auto } then have AE_1: "AE x in M. liminf (\n. -u n x / n) = - subcocycle_lim_ereal u x" "AE x in M. (\n. - u n x / n) \ - subcocycle_lim_ereal u x" "AE x in M. - subcocycle_lim_ereal u x \ 0" using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] by auto have "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) = (\\<^sup>+ x. liminf (\n. -u n x / n) \M)" apply (rule nn_integral_cong_AE) using AE_1(1) by auto also have "... \ liminf (\n. \\<^sup>+ x. -u n x / n \M)" apply (subst e2ennreal_Liminf) apply (simp_all add: e2ennreal_ereal) using F_pos by (intro nn_integral_liminf) (simp add: int_F) also have "... = - subcocycle_avg_ereal u" proof - have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" by (rule subcocycle_int_tendsto_avg_ereal[OF assms(1)]) with tendsto_cmult_ereal[OF _ this, of "-1"] have "(\n. (\x. - u n x / n \M)) \ - subcocycle_avg_ereal u" by simp then have "- subcocycle_avg_ereal u = liminf (\n. (\x. - u n x / n \M))" by (simp add: tendsto_iff_Liminf_eq_Limsup) moreover have "(\\<^sup>+ x. ennreal (-u n x / n) \M) = ennreal(\x. - u n x / n \M)" for n apply (rule nn_integral_eq_integral[OF int_F]) using F_pos by auto ultimately show ?thesis by (auto simp: e2ennreal_Liminf e2ennreal_ereal) qed finally have "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) \ - subcocycle_avg_ereal u" by simp also have "\ < \" using assms(2) by (cases "subcocycle_avg_ereal u") (auto simp: e2ennreal_ereal e2ennreal_neg) finally have *: "(\\<^sup>+ x. -subcocycle_lim_ereal u x \M) < \" . have "AE x in M. e2ennreal (- subcocycle_lim_ereal u x) \ \" apply (rule nn_integral_PInf_AE) using * by auto then have **: "AE x in M. - subcocycle_lim_ereal u x \ \" using AE_1(3) by eventually_elim simp { fix x assume H: "- subcocycle_lim_ereal u x \ \" "(\n. u n x / n) \ subcocycle_lim_ereal u x" "- subcocycle_lim_ereal u x \ 0" then have 1: "abs(subcocycle_lim_ereal u x) \ \" by auto then have 2: "(\n. u n x / n) \ subcocycle_lim u x" using H(2) unfolding subcocycle_lim_def by auto then have 3: "(\n. - (u n x / n)) \ - subcocycle_lim u x" using tendsto_mult[OF _ 2, of "\_. -1", of "-1"] by auto have 4: "-subcocycle_lim u x \ 0" using H(3) unfolding subcocycle_lim_def by auto have "abs(subcocycle_lim_ereal u x) \ \" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. - (u n x / n)) \ - subcocycle_lim u x" "-subcocycle_lim u x \ 0" using 1 2 3 4 by auto } then have AE_2: "AE x in M. abs(subcocycle_lim_ereal u x) \ \" "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "AE x in M. (\n. - (u n x / n)) \ - subcocycle_lim u x" "AE x in M. -subcocycle_lim u x \ 0" using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] ** AE_1(3) by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by simp have "(\\<^sup>+x. abs(subcocycle_lim u x) \M) = (\\<^sup>+x. -subcocycle_lim_ereal u x \M)" apply (rule nn_integral_cong_AE) using AE_2 unfolding subcocycle_lim_def abs_real_of_ereal apply eventually_elim by (auto simp: e2ennreal_ereal) then have A: "(\\<^sup>+x. abs(subcocycle_lim u x) \M) < \" using * by auto show int_Gr: "integrable M (subcocycle_lim u)" apply (rule integrableI_bounded) using A by auto have B: "(\n. (\\<^sup>+ x. norm((- u n x /n) - (-subcocycle_lim u x)) \M)) \ 0" proof (rule Scheffe_lemma1, auto simp add: int_Gr int_u AE_2(2) AE_2(3)) { fix n assume "n>(0::nat)" have *: "AE x in M. subcocycle_lim u x \ real_cond_exp M Invariants (\x. u n x / n) x" using subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1) \n>0\] AE_2(1) unfolding subcocycle_lim_def by auto have "(\x. subcocycle_lim u x \M) \ (\x. real_cond_exp M Invariants (\x. u n x / n) x \M)" apply (rule integral_mono_AE[OF int_Gr _ *], rule real_cond_exp_int(1)) using int_u by auto also have "... = (\x. u n x / n \M)" apply (rule real_cond_exp_int(2)) using int_u by auto finally have A: "(\x. subcocycle_lim u x \M) \ (\x. u n x / n \M)" by auto have "(\\<^sup>+x. abs(u n x) / n \M) = (\\<^sup>+x. - u n x / n \M)" apply (rule nn_integral_cong) using F_pos abs_of_nonneg by (intro arg_cong[where f = ennreal]) fastforce also have "... = (\x. - u n x / n \M)" apply (rule nn_integral_eq_integral) using F_pos int_F by auto also have "... \ (\x. - subcocycle_lim u x \M)" using A by (auto intro!: ennreal_leI) also have "... = (\\<^sup>+x. - subcocycle_lim u x \M)" apply (rule nn_integral_eq_integral[symmetric]) using int_Gr AE_2(4) by auto also have "... = (\\<^sup>+x. abs(subcocycle_lim u x) \M)" apply (rule nn_integral_cong_AE) using AE_2(4) by auto finally have "(\\<^sup>+x. abs(u n x) / n \M) \ (\\<^sup>+x. abs(subcocycle_lim u x) \M)" by simp } with eventually_mono[OF eventually_gt_at_top[of 0] this] have "eventually (\n. (\\<^sup>+x. abs(u n x) / n \M) \ (\\<^sup>+x. abs(subcocycle_lim u x) \M)) sequentially" by fastforce then show "limsup (\n. \\<^sup>+ x. abs(u n x) / n \M) \ \\<^sup>+ x. abs(subcocycle_lim u x) \M" using Limsup_bounded by fastforce qed moreover have "norm((- u n x /n) - (-subcocycle_lim u x)) = abs(u n x / n - subcocycle_lim u x)" for n x by auto ultimately show "(\n. \\<^sup>+ x. ennreal \u n x / real n - subcocycle_lim u x\ \M) \ 0" by auto qed text \We can then remove the nonpositivity assumption, by subtracting the Birkhoff sums of $u_1$ to a general subcocycle $u$.\ theorem kingman_theorem_nonergodic: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "integrable M (subcocycle_lim u)" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof - have [measurable]: "u n \ borel_measurable M" for n using assms(1) unfolding subcocycle_def by auto have int_u [measurable]: "integrable M (u 1)" using assms(1) subcocycle_def by auto define v where "v = (\n x. u n x + birkhoff_sum (\x. - u 1 x) n x)" have [measurable]: "v n \ borel_measurable M" for n unfolding v_def by auto define w where "w = birkhoff_sum (u 1)" have [measurable]: "w n \ borel_measurable M" for n unfolding w_def by auto have "subcocycle v" unfolding v_def apply (rule subcocycle_add[OF assms(1)], rule subcocycle_birkhoff) using assms unfolding subcocycle_def by auto have "subcocycle w" unfolding w_def by (rule subcocycle_birkhoff[OF int_u]) have uvw: "u n x = v n x + w n x" for n x unfolding v_def w_def birkhoff_sum_def by (auto simp add: sum_negf) then have "subcocycle_avg_ereal (\n x. u n x) = subcocycle_avg_ereal v + subcocycle_avg_ereal w" using subcocycle_avg_ereal_add[OF \subcocycle v\ \subcocycle w\] by auto then have "subcocycle_avg_ereal u = subcocycle_avg_ereal v + subcocycle_avg_ereal w" by auto then have "subcocycle_avg_ereal v > -\" unfolding w_def using subcocycle_avg_ereal_birkhoff[OF int_u] assms(2) by auto have "subcocycle_avg_ereal w > - \" unfolding w_def using subcocycle_avg_birkhoff[OF int_u] by auto have "\x. v 1 x \ 0" unfolding v_def by auto have v: "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" "integrable M (subcocycle_lim v)" "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M)) \ 0" using kingman_theorem_L1_aux[OF \subcocycle v\ \subcocycle_avg_ereal v > -\\ \\x. v 1 x \ 0\] by auto have w: "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" "integrable M (subcocycle_lim w)" "(\n. (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" proof - show "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] birkhoff_theorem_AE_nonergodic[OF int_u] by auto show "integrable M (subcocycle_lim w)" apply (subst integrable_cong_AE[where ?g = "\x. real_cond_exp M Invariants (u 1) x"]) unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] real_cond_exp_int(1)[OF int_u] by auto have "(\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M) = (\\<^sup>+x. abs(birkhoff_sum (u 1) n x / n - real_cond_exp M Invariants (u 1) x) \M)" for n apply (rule nn_integral_cong_AE) unfolding w_def subcocycle_lim_def using subcocycle_lim_ereal_birkhoff[OF int_u] by auto then show "(\n. (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" using birkhoff_theorem_L1_nonergodic[OF int_u] by auto qed { fix x assume H: "(\n. v n x / n) \ subcocycle_lim v x" "(\n. w n x / n) \ subcocycle_lim w x" "(\n. u n x / n) \ subcocycle_lim_ereal u x" then have "(\n. v n x / n + w n x / n) \ subcocycle_lim v x + subcocycle_lim w x" using tendsto_add[OF H(1) H(2)] by simp then have *: "(\n. ereal(u n x / n)) \ ereal(subcocycle_lim v x + subcocycle_lim w x)" unfolding uvw by (simp add: add_divide_distrib) then have "subcocycle_lim_ereal u x = ereal(subcocycle_lim v x + subcocycle_lim w x)" using H(3) LIMSEQ_unique by blast then have **: "subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x" using subcocycle_lim_def by auto have "u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x" for n apply (subst **, subst uvw) using add_divide_distrib add.commute by auto then have "(\n. u n x / n) \ subcocycle_lim u x \ subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x \ (\n. u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x)" using * ** by auto } then have AE: "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" "AE x in M. subcocycle_lim u x = subcocycle_lim v x + subcocycle_lim w x" "AE x in M. \n. u n x / n - subcocycle_lim u x = v n x / n - subcocycle_lim v x + w n x / n - subcocycle_lim w x" using v(1) w(1) kingman_theorem_AE_nonergodic_ereal[OF assms(1)] by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by simp show "integrable M (subcocycle_lim u)" apply (subst integrable_cong_AE[where ?g = "\x. subcocycle_lim v x + subcocycle_lim w x"]) by (auto simp add: AE(2) v(2) w(2)) show "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)) \ 0" proof (rule tendsto_sandwich[where ?f = "\_. 0" and ?h = "\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)"], auto) { fix n have "(\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) = (\\<^sup>+x. abs((v n x / n - subcocycle_lim v x) + (w n x / n - subcocycle_lim w x)) \M)" apply (rule nn_integral_cong_AE) using AE(3) by auto also have "... \ (\\<^sup>+x. ennreal(abs(v n x / n - subcocycle_lim v x)) + abs(w n x / n - subcocycle_lim w x) \M)" by (rule nn_integral_mono, auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus) also have "... = (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)" by (rule nn_integral_add, auto, measurable) finally have "(\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) \ (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)" using tendsto_sandwich by simp } then show "eventually (\n. (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M) \ (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) sequentially" by auto have "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0 + 0" by (rule tendsto_add[OF v(3) w(3)]) then show "(\n. (\\<^sup>+x. abs(v n x / n - subcocycle_lim v x) \M) + (\\<^sup>+x. abs(w n x / n - subcocycle_lim w x) \M)) \ 0" by simp qed qed text \From the almost sure convergence, we can prove the basic properties of the (real) subcocycle limit: relationship to the asymptotic average, behavior under sum, multiplication, max, behavior for Birkhoff sums.\ lemma subcocycle_lim_avg: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "(\x. subcocycle_lim u x \M) = subcocycle_avg u" proof - have H: "(\n. (\\<^sup>+x. norm(u n x / n - subcocycle_lim u x) \M)) \ 0" "integrable M (subcocycle_lim u)" using kingman_theorem_nonergodic[OF assms] by auto have "(\n. (\x. u n x / n \M)) \ (\x. subcocycle_lim u x \M)" apply (rule tendsto_L1_int[OF _ H(2) H(1)]) using subcocycle_integrable[OF assms(1)] by auto then have "(\n. (\x. u n x / n \M)) \ ereal (\x. subcocycle_lim u x \M)" by auto moreover have "(\n. (\x. u n x / n \M)) \ ereal (subcocycle_avg u)" using subcocycle_int_tendsto_avg[OF assms] by auto ultimately show ?thesis using LIMSEQ_unique by blast qed lemma subcocycle_lim_real_ereal: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" proof - { fix x assume H: "(\n. u n x / n) \ subcocycle_lim_ereal u x" "(\n. u n x / n) \ subcocycle_lim u x" then have "(\n. u n x / n) \ ereal(subcocycle_lim u x)" by auto then have "subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" using H(1) LIMSEQ_unique by blast } then show ?thesis using kingman_theorem_AE_nonergodic_ereal[OF assms(1)] kingman_theorem_nonergodic(1)[OF assms] by auto qed lemma subcocycle_lim_add: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > -\" "subcocycle_avg_ereal v > -\" shows "subcocycle_avg_ereal (\n x. u n x + v n x) > - \" "AE x in M. subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" proof - show *: "subcocycle_avg_ereal (\n x. u n x + v n x) > - \" using subcocycle_avg_add[OF assms(1) assms(2)] assms(3) assms(4) by auto have "AE x in M. (\n. (u n x + v n x)/n) \ subcocycle_lim (\n x. u n x + v n x) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_add[OF assms(1) assms(2)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic[OF assms(1) assms(3)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" by (rule kingman_theorem_nonergodic[OF assms(2) assms(4)]) moreover { fix x assume H: "(\n. (u n x + v n x)/n) \ subcocycle_lim (\n x. u n x + v n x) x" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. v n x / n) \ subcocycle_lim v x" have *: "(u n x + v n x)/n = (u n x / n) + (v n x / n)" for n by (simp add: add_divide_distrib) have "(\n. (u n x + v n x)/n) \ subcocycle_lim u x + subcocycle_lim v x" unfolding * by (intro tendsto_intros H) then have "subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. u n x + v n x) x = subcocycle_lim u x + subcocycle_lim v x" by auto qed lemma subcocycle_lim_cmult: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" "c\(0::real)" shows "subcocycle_avg_ereal (\n x. c * u n x) > - \" "AE x in M. subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" proof - show *: "subcocycle_avg_ereal (\n x. c * u n x) > - \" using subcocycle_avg_cmult[OF assms(1) assms(3)] assms(2) assms(3) by auto have "AE x in M. (\n. (c * u n x)/n) \ subcocycle_lim (\n x. c * u n x) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_cmult[OF assms(1) assms(3)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic(1)[OF assms(1) assms(2)]) moreover { fix x assume H: "(\n. (c * u n x)/n) \ subcocycle_lim (\n x. c * u n x) x" "(\n. u n x / n) \ subcocycle_lim u x" have "(\n. c * (u n x / n)) \ c * subcocycle_lim u x" by (rule tendsto_mult[OF _ H(2)], auto) then have "subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. c * u n x) x = c * subcocycle_lim u x" by auto qed lemma subcocycle_lim_max: assumes "subcocycle u" "subcocycle v" "subcocycle_avg_ereal u > -\" "subcocycle_avg_ereal v > -\" shows "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" "AE x in M. subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" proof - show *: "subcocycle_avg_ereal (\n x. max (u n x) (v n x)) > - \" using subcocycle_avg_max(1)[OF assms(1) assms(2)] assms(3) assms(4) by auto have "AE x in M. (\n. max (u n x) (v n x) / n) \ subcocycle_lim (\n x. max (u n x) (v n x)) x" by (rule kingman_theorem_nonergodic[OF subcocycle_max[OF assms(1) assms(2)] *]) moreover have "AE x in M. (\n. u n x / n) \ subcocycle_lim u x" by (rule kingman_theorem_nonergodic[OF assms(1) assms(3)]) moreover have "AE x in M. (\n. v n x / n) \ subcocycle_lim v x" by (rule kingman_theorem_nonergodic[OF assms(2) assms(4)]) moreover { fix x assume H: "(\n. max (u n x) (v n x) / n) \ subcocycle_lim (\n x. max (u n x) (v n x)) x" "(\n. u n x / n) \ subcocycle_lim u x" "(\n. v n x / n) \ subcocycle_lim v x" have "(\n. max (u n x / n) (v n x / n)) \ max (subcocycle_lim u x) (subcocycle_lim v x)" apply (rule tendsto_max) using H by auto moreover have "max (u n x / n) (v n x / n) = max (u n x) (v n x) / n" for n by (simp add: max_divide_distrib_right) ultimately have "(\n. max (u n x) (v n x) / n) \ max (subcocycle_lim u x) (subcocycle_lim v x)" by auto then have "subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" using H(1) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (\n x. max (u n x) (v n x)) x = max (subcocycle_lim u x) (subcocycle_lim v x)" by auto qed lemma subcocycle_lim_birkhoff: assumes "integrable M u" shows "subcocycle_avg_ereal (birkhoff_sum u) > -\" "AE x in M. subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" proof - show *: "subcocycle_avg_ereal (birkhoff_sum u) > -\" using subcocycle_avg_birkhoff[OF assms] by auto have "AE x in M. (\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" by (rule birkhoff_theorem_AE_nonergodic[OF assms]) moreover have "AE x in M. (\n. birkhoff_sum u n x / n) \ subcocycle_lim (birkhoff_sum u) x" by (rule kingman_theorem_nonergodic(1)[OF subcocycle_birkhoff[OF assms] *]) moreover { fix x assume H: "(\n. birkhoff_sum u n x / n) \ real_cond_exp M Invariants u x" "(\n. birkhoff_sum u n x / n) \ subcocycle_lim (birkhoff_sum u) x" then have "subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" using H(2) by (simp add: LIMSEQ_unique) } ultimately show "AE x in M. subcocycle_lim (birkhoff_sum u) x = real_cond_exp M Invariants u x" by auto qed subsection \Conditional expectations of subcocycles\ text \In this subsection, we show that the conditional expectations of a subcocycle (with respect to the invariant subalgebra) also converge, with the same limit as the cocycle. Note that the conditional expectation of a subcocycle $u$ is still a subcocycle, with the same average at each step so with the same asymptotic average. Kingman theorem can be applied to it, and what we have to show is that the limit of this subcocycle is the same as the limit of the original subcocycle. When the asymptotic average is $>-\infty$, both limits have the same integral, and moreover the domination of the subcocycle by the Birkhoff sums of $u_n$ for fixed $n$ (which converge to the conditional expectation of $u_n$) implies that one limit is smaller than the other. Hence, they coincide almost everywhere. The case when the asymptotic average is $-\infty$ is deduced from the previous one by truncation. \ text \First, we prove the result when the asymptotic average with finite.\ theorem kingman_theorem_nonergodic_invariant: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" "(\n. (\\<^sup>+x. abs(real_cond_exp M Invariants (u n) x / n - subcocycle_lim u x) \M)) \ 0" proof - have int [simp]: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by auto then have int2: "integrable M (real_cond_exp M Invariants (u n))" for n using real_cond_exp_int by auto { fix n m have "u (n+m) x \ u n x + u m ((T^^n) x)" for x using subcocycle_ineq[OF assms(1)] by auto have "AE x in M. real_cond_exp M Invariants (u (n+m)) x \ real_cond_exp M Invariants (\x. u n x + u m ((T^^n) x)) x" apply (rule real_cond_exp_mono) using subcocycle_ineq[OF assms(1)] apply auto by (rule Bochner_Integration.integrable_add, auto simp add: Tn_integral_preserving) moreover have "AE x in M. real_cond_exp M Invariants (\x. u n x + u m ((T^^n) x)) x = real_cond_exp M Invariants (u n) x + real_cond_exp M Invariants (\x. u m ((T^^n) x)) x" by (rule real_cond_exp_add, auto simp add: Tn_integral_preserving) moreover have "AE x in M. real_cond_exp M Invariants (u m \ ((T^^n))) x = real_cond_exp M Invariants (u m) x" by (rule Invariants_of_foTn, simp) moreover have "AE x in M. real_cond_exp M Invariants (u m) x = real_cond_exp M Invariants (u m) ((T^^n) x)" using Invariants_func_is_invariant_n[symmetric, of "real_cond_exp M Invariants (u m)"] by auto ultimately have "AE x in M. real_cond_exp M Invariants (u (n+m)) x \ real_cond_exp M Invariants (u n) x + real_cond_exp M Invariants (u m) ((T^^n) x)" unfolding o_def by auto } with subcocycle_AE[OF this int2] obtain w where w: "subcocycle w" "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" by blast have [measurable]: "integrable M (w n)" for n using subcocycle_integrable[OF w(1)] by simp { fix n::nat have "(\x. w n x / n \M) = (\x. real_cond_exp M Invariants (u n) x / n \M)" apply (rule integral_cong_AE) using w(2) by auto also have "... = (\x. real_cond_exp M Invariants (u n) x \M) / n" by (rule integral_divide_zero) also have "... = (\x. u n x \M) / n" by (simp add: divide_simps real_cond_exp_int(2)[OF int[of n]]) also have "... = (\x. u n x / n \M)" by (rule integral_divide_zero[symmetric]) finally have "ereal (\x. w n x / n \M) = ereal (\x. u n x / n \M)" by simp } note * = this have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal w" apply (rule Lim_transform_eventually[OF subcocycle_int_tendsto_avg_ereal[OF w(1)]]) using * by auto then have "subcocycle_avg_ereal u = subcocycle_avg_ereal w" using subcocycle_int_tendsto_avg_ereal[OF assms(1)] LIMSEQ_unique by auto then have "subcocycle_avg_ereal w > -\" using assms(2) by simp have "subcocycle_avg u = subcocycle_avg w" using \subcocycle_avg_ereal u = subcocycle_avg_ereal w\ unfolding subcocycle_avg_def by simp have *: "AE x in M. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" for N by (cases "N = 0", auto simp add: subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1)]) have "AE x in M. \N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by (subst AE_all_countable, intro allI, simp add: *) moreover have "AE x in M. subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" by (rule subcocycle_lim_real_ereal[OF assms]) moreover have "AE x in M. (\N. u N x / N) \ subcocycle_lim u x" using kingman_theorem_nonergodic[OF assms] by simp moreover have "AE x in M. (\N. w N x / N) \ subcocycle_lim w x" using kingman_theorem_nonergodic[OF w(1) \subcocycle_avg_ereal w > -\\ ] by simp moreover have "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" using w(2) by simp moreover have "AE x in M. \n. real_cond_exp M Invariants (u n) x / n = real_cond_exp M Invariants (\x. u n x / n) x" apply (subst AE_all_countable, intro allI) using AE_symmetric[OF real_cond_exp_cdiv[OF int]] by auto moreover { fix x assume x: "\N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" "subcocycle_lim_ereal u x = ereal(subcocycle_lim u x)" "(\N. u N x / N) \ subcocycle_lim u x" "(\N. w N x / N) \ subcocycle_lim w x" "\n. w n x = real_cond_exp M Invariants (u n) x" "\n. real_cond_exp M Invariants (u n) x / n = real_cond_exp M Invariants (\x. u n x / n) x" { fix N::nat assume "N\1" have "subcocycle_lim u x \ real_cond_exp M Invariants (\x. u N x / N) x" using x(1) x(2) \N\1\ by auto also have "... = real_cond_exp M Invariants (u N) x / N" using x(6) by simp also have "... = w N x / N" using x(5) by simp finally have "subcocycle_lim u x \ w N x / N" by simp } note * = this have "subcocycle_lim u x \ subcocycle_lim w x" apply (rule LIMSEQ_le_const[OF x(4)]) using * by auto } ultimately have *: "AE x in M. subcocycle_lim u x \ subcocycle_lim w x" by auto have **: "(\x. subcocycle_lim u x \M) = (\x. subcocycle_lim w x \M)" using subcocycle_lim_avg[OF assms] subcocycle_lim_avg[OF w(1) \subcocycle_avg_ereal w > -\\] \subcocycle_avg u = subcocycle_avg w\ by simp have AE_eq: "AE x in M. subcocycle_lim u x = subcocycle_lim w x" by (rule integral_ineq_eq_0_then_AE[OF * kingman_theorem_nonergodic(2)[OF assms] kingman_theorem_nonergodic(2)[OF w(1) \subcocycle_avg_ereal w > -\\] **]) moreover have "AE x in M. (\n. w n x / n) \ subcocycle_lim w x" by (rule kingman_theorem_nonergodic(1)[OF w(1) \subcocycle_avg_ereal w > -\\]) moreover have "AE x in M. \n. w n x = real_cond_exp M Invariants (u n) x" using w(2) by auto moreover { fix x assume H: "subcocycle_lim u x = subcocycle_lim w x" "(\n. w n x / n) \ subcocycle_lim w x" "\n. w n x = real_cond_exp M Invariants (u n) x" then have "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" by auto } ultimately show "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim u x" by auto { fix n::nat have "AE x in M. subcocycle_lim u x = subcocycle_lim w x" using AE_eq by simp moreover have "AE x in M. w n x = real_cond_exp M Invariants (u n) x" using w(2) by auto moreover { fix x assume H: "subcocycle_lim u x = subcocycle_lim w x" "w n x = real_cond_exp M Invariants (u n) x" then have "ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ = ennreal \w n x / real n - subcocycle_lim w x\" by auto } ultimately have "AE x in M. ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ = ennreal \w n x / real n - subcocycle_lim w x\" by auto then have "(\\<^sup>+ x. ennreal \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ \M) = (\\<^sup>+ x. ennreal \w n x / real n - subcocycle_lim w x\ \M)" by (rule nn_integral_cong_AE) } moreover have "(\n. (\\<^sup>+ x. \w n x / real n - subcocycle_lim w x\ \M)) \ 0" by (rule kingman_theorem_nonergodic(3)[OF w(1) \subcocycle_avg_ereal w > -\\]) ultimately show "(\n. (\\<^sup>+ x. \real_cond_exp M Invariants (u n) x / real n - subcocycle_lim u x\ \M)) \ 0" by auto qed text \Then, we extend it by truncation to the general case, i.e., to the asymptotic limit in extended reals.\ theorem kingman_theorem_AE_nonergodic_invariant_ereal: assumes "subcocycle u" shows "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" proof - have [simp]: "subcocycle u" using assms by simp have int [simp]: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by auto have limsup_ineq_K: "AE x in M. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" for K::nat proof - define v where "v = (\ (n::nat) (x::'a). (-n * real K))" have [simp]: "subcocycle v" unfolding v_def subcocycle_def by (auto simp add: algebra_simps) have "ereal (\x. v n x / n \M) = ereal(- real K * measure M (space M))" if "n\1" for n unfolding v_def using that by simp then have "(\n. ereal (\x. v n x / n \M)) \ ereal(- real K * measure M (space M))" using lim_explicit by force moreover have "(\n. ereal (\x. v n x / n \M)) \ subcocycle_avg_ereal v" using subcocycle_int_tendsto_avg_ereal[OF \subcocycle v\] by auto ultimately have "subcocycle_avg_ereal v = - real K * measure M (space M)" using LIMSEQ_unique by blast then have "subcocycle_avg_ereal v > -\" by auto { fix x assume H: "(\n. v n x / n) \ subcocycle_lim_ereal v x" have "ereal(v n x / n) = -real K" if "n\1" for n unfolding v_def using that by auto then have "(\n. ereal(v n x / n)) \ - real K" using lim_explicit by force then have "subcocycle_lim_ereal v x = -real K" using H LIMSEQ_unique by blast } then have "AE x in M. subcocycle_lim_ereal v x = -real K" using kingman_theorem_AE_nonergodic_ereal[OF \subcocycle v\] by auto define w where "w = (\n x. max (u n x) (v n x))" have [simp]: "subcocycle w" unfolding w_def by (rule subcocycle_max, auto) have "subcocycle_avg_ereal w \ subcocycle_avg_ereal v" unfolding w_def using subcocycle_avg_ereal_max by auto then have "subcocycle_avg_ereal w > -\" using \subcocycle_avg_ereal v > -\\ by auto have *: "AE x in M. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" for n apply (rule real_cond_exp_mono) using subcocycle_integrable[OF assms, of n] subcocycle_integrable[OF \subcocycle w\, of n] apply auto unfolding w_def by auto have "AE x in M. \n. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" apply (subst AE_all_countable) using * by auto moreover have "AE x in M. (\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim w x" apply (rule kingman_theorem_nonergodic_invariant(1)) using \subcocycle_avg_ereal w > -\\ by auto moreover have "AE x in M. subcocycle_lim_ereal w x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" unfolding w_def using subcocycle_lim_ereal_max by auto moreover { fix x assume H: "(\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim w x" "subcocycle_lim_ereal w x = max (subcocycle_lim_ereal u x) (subcocycle_lim_ereal v x)" "subcocycle_lim_ereal v x = - real K" "\n. real_cond_exp M Invariants (u n) x \ real_cond_exp M Invariants (w n) x" have "subcocycle_lim_ereal w x > -\" - using H(2) H(3) + using H(2) H(3) by auto (metis MInfty_neq_ereal(1) ereal_infty_less_eq2(2) max.cobounded2) then have "subcocycle_lim_ereal w x = ereal(subcocycle_lim w x)" unfolding subcocycle_lim_def using subcocycle_lim_ereal_not_PInf[of w x] ereal_real by force moreover have "(\n. real_cond_exp M Invariants (w n) x / n) \ ereal(subcocycle_lim w x)" using H(1) by auto ultimately have "(\n. real_cond_exp M Invariants (w n) x / n) \ subcocycle_lim_ereal w x" by auto then have *: "limsup (\n. real_cond_exp M Invariants (w n) x / n) = subcocycle_lim_ereal w x" using tendsto_iff_Liminf_eq_Limsup trivial_limit_at_top_linorder by blast have "ereal(real_cond_exp M Invariants (u n) x / n) \ real_cond_exp M Invariants (w n) x / n" for n using H(4) by (auto simp add: divide_simps) then have "eventually (\n. ereal(real_cond_exp M Invariants (u n) x / n) \ real_cond_exp M Invariants (w n) x / n) sequentially" by auto then have "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ limsup (\n. real_cond_exp M Invariants (w n) x / n)" using Limsup_mono[of _ _ sequentially] by force then have "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" using * H(2) H(3) by auto } ultimately show ?thesis using \AE x in M. subcocycle_lim_ereal v x = -real K\ by auto qed have "AE x in M. \K::nat. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" apply (subst AE_all_countable) using limsup_ineq_K by auto moreover have "AE x in M. liminf (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" proof - have *: "AE x in M. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" for N by (cases "N = 0", auto simp add: subcocycle_lim_ereal_atmost_uN_invariants[OF assms(1)]) have "AE x in M. \N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" by (subst AE_all_countable, intro allI, simp add: *) moreover have "AE x in M. \n. real_cond_exp M Invariants (\x. u n x / n) x = real_cond_exp M Invariants (u n) x / n" apply (subst AE_all_countable, intro allI) using real_cond_exp_cdiv by auto moreover { fix x assume x: "\N. N > 0 \ subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" "\n. real_cond_exp M Invariants (\x. u n x / n) x = real_cond_exp M Invariants (u n) x / n" then have *: "subcocycle_lim_ereal u x \ real_cond_exp M Invariants (u n) x / n" if "n \ 1" for n using that by auto have "subcocycle_lim_ereal u x \ liminf (\n. real_cond_exp M Invariants (u n) x / n)" apply (subst liminf_bounded_iff) using * less_le_trans by blast } ultimately show ?thesis by auto qed moreover { fix x assume H: "\K::nat. limsup (\n. real_cond_exp M Invariants (u n) x / n) \ max (subcocycle_lim_ereal u x) (-real K)" "liminf (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" have "(\K::nat. max (subcocycle_lim_ereal u x) (-real K)) \ subcocycle_lim_ereal u x" by (rule ereal_truncation_bottom) with LIMSEQ_le_const[OF this] have *: "limsup (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" using H(1) by auto have "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" apply (subst tendsto_iff_Liminf_eq_Limsup[OF trivial_limit_at_top_linorder]) using H(2) * Liminf_le_Limsup[OF trivial_limit_at_top_linorder, of "(\n. real_cond_exp M Invariants (u n) x / n)"] by auto } ultimately show ?thesis by auto qed end subsection \Subcocycles in the ergodic case\ text \In this subsection, we describe how all the previous results simplify in the ergodic case. Indeed, subcocycle limits are almost surely constant, given by the asymptotic average.\ context ergodic_pmpt begin lemma subcocycle_ergodic_lim_avg: assumes "subcocycle u" shows "AE x in M. subcocycle_lim_ereal u x = subcocycle_avg_ereal u" "AE x in M. subcocycle_lim u x = subcocycle_avg u" proof - have I: "integrable M (u N)" for N using subcocycle_integrable[OF assms]by auto obtain c::ereal where c: "AE x in M. subcocycle_lim_ereal u x = c" using Invariant_func_is_AE_constant[OF subcocycle_lim_meas_Inv(1)] by blast have "c = subcocycle_avg_ereal u" proof (cases "subcocycle_avg_ereal u = - \") case True { fix N assume "N > (0::nat)" have "AE x in M. real_cond_exp M Invariants (\x. u N x / N) x = (\ x. u N x / N \M)" apply (rule Invariants_cond_exp_is_integral) using I by auto moreover have "AE x in M. subcocycle_lim_ereal u x \ real_cond_exp M Invariants (\x. u N x / N) x" using subcocycle_lim_ereal_atmost_uN_invariants[OF assms \N>0\] by simp ultimately have "AE x in M. c \ (\x. u N x / N \M)" using c by force then have "c \ (\x. u N x / N \M)" by auto } then have "\N\1. c \ (\x. u N x / N \M)" by auto with Lim_bounded2[OF subcocycle_int_tendsto_avg_ereal[OF assms] this] have "c \ subcocycle_avg_ereal u" by simp then show ?thesis using True by auto next case False then have fin: "subcocycle_avg_ereal u > - \" by simp obtain cr::real where cr: "AE x in M. subcocycle_lim u x = cr" using Invariant_func_is_AE_constant[OF subcocycle_lim_meas_Inv(2)] by blast have "AE x in M. c = ereal cr" using c cr subcocycle_lim_real_ereal[OF assms fin] by force then have "c = ereal cr" by auto have "subcocycle_avg u = (\x. subcocycle_lim u x \M)" using subcocycle_lim_avg[OF assms fin] by auto also have "... = (\x. cr \M)" apply (rule integral_cong_AE) using cr by auto also have "... = cr" by (simp add: prob_space.prob_space prob_space_axioms) finally have "ereal(subcocycle_avg u) = ereal cr" by simp then show ?thesis using \ c = ereal cr \ subcocycle_avg_real_ereal[OF fin] by auto qed then show "AE x in M. subcocycle_lim_ereal u x = subcocycle_avg_ereal u" using c by auto then show "AE x in M. subcocycle_lim u x = subcocycle_avg u" unfolding subcocycle_lim_def subcocycle_avg_def by auto qed theorem kingman_theorem_AE_ereal: assumes "subcocycle u" shows "AE x in M. (\n. u n x / n) \ subcocycle_avg_ereal u" using kingman_theorem_AE_nonergodic_ereal[OF assms] subcocycle_ergodic_lim_avg(1)[OF assms] by auto theorem kingman_theorem: assumes "subcocycle u" "subcocycle_avg_ereal u > -\" shows "AE x in M. (\n. u n x / n) \ subcocycle_avg u" "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M)) \ 0" proof - have *: "AE x in M. subcocycle_lim u x = subcocycle_avg u" using subcocycle_ergodic_lim_avg(2)[OF assms(1)] by auto then show "AE x in M. (\n. u n x / n) \ subcocycle_avg u" using kingman_theorem_nonergodic(1)[OF assms] by auto have "(\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M) = (\\<^sup>+x. abs(u n x / n - subcocycle_lim u x) \M)" for n apply (rule nn_integral_cong_AE) using * by auto then show "(\n. (\\<^sup>+x. abs(u n x / n - subcocycle_avg u) \M)) \ 0" using kingman_theorem_nonergodic(3)[OF assms] by auto qed end subsection \Subocycles for invertible maps\ text \If $T$ is invertible, then a subcocycle $u_n$ for $T$ gives rise to another subcocycle for $T^{-1}$. Intuitively, if $u$ is subadditive along the time interval $[0,n)$, then it should also be subadditive along the time interval $[-n,0)$. This is true, and formalized with the following statement.\ proposition (in mpt) subcocycle_u_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "mpt.subcocycle M Tinv (\n x. u n (((Tinv)^^n) x))" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: mpt M Tinv using Tinv_mpt[OF assms(2)] by simp show "I.subcocycle (\n x. u n (((Tinv)^^n) x))" unfolding I.subcocycle_def proof(auto) show "integrable M (\x. u n ((Tinv ^^ n) x))" for n using I.Tn_integral_preserving(1)[OF int[of n]] by simp fix n m::nat and x::'a define y where "y = (Tinv^^(m+n)) x" have "(T^^m) y = (T^^m) ((Tinv^^m) ((Tinv^^n) x))" unfolding y_def by (simp add: funpow_add) then have *: "(T^^m) y = (Tinv^^n) x" using fn_o_inv_fn_is_id[OF bij, of m] by (metis Tinv_def comp_def) have "u (n + m) ((Tinv ^^ (n + m)) x) = u (m+n) y" unfolding y_def by (simp add: add.commute[of n m]) also have "... \ u m y + u n ((T^^m) y)" using subcocycle_ineq[OF \subcocycle u\, of m n y] by simp also have "... = u m ((Tinv^^(m+n)) x) + u n ((Tinv^^n) x)" using * y_def by auto finally show "u (n + m) ((Tinv ^^ (n + m)) x) \ u n ((Tinv ^^ n) x) + u m ((Tinv ^^ m) ((Tinv ^^ n) x))" by (simp add: funpow_add) qed qed text \The subcocycle averages for $T$ and $T^{-1}$ coincide.\ proposition (in mpt) subcocycle_avg_ereal_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "mpt.subcocycle_avg_ereal M (\n x. u n (((Tinv)^^n) x)) = subcocycle_avg_ereal u" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: mpt M Tinv using Tinv_mpt[OF assms(2)] by simp have "(\n. (\x. u n (((Tinv)^^n) x) / n \M)) \ I.subcocycle_avg_ereal (\n x. u n (((Tinv)^^n) x))" using I.subcocycle_int_tendsto_avg_ereal[OF subcocycle_u_Tinv[OF assms]] by simp moreover have "(\x. u n x / n \M) = ereal (\x. u n (((Tinv)^^n) x) / n \M)" for n apply (simp) apply (rule disjI2) apply (rule I.Tn_integral_preserving(2)[symmetric]) apply (simp add: int) done ultimately have "(\n. (\x. u n x / n \M)) \ I.subcocycle_avg_ereal (\n x. u n (((Tinv)^^n) x))" by presburger moreover have "(\n. (\x. u n x / n \M)) \ subcocycle_avg_ereal u" using subcocycle_int_tendsto_avg_ereal[OF \subcocycle u\] by simp ultimately show ?thesis using LIMSEQ_unique by simp qed text \The asymptotic limit of the subcocycle is the same for $T$ and $T^{-1}$. This is clear in the ergodic case, and follows from the ergodic decomposition in the general case (on a standard probability space). We give a direct proof below (on a general probability space) using the fact that the asymptotic limit is the same for the subcocycle conditioned by the invariant sigma-algebra, which is clearly the same for $T$ and $T^{-1}$ as it is constant along orbits.\ proposition (in fmpt) subcocycle_lim_ereal_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "AE x in M. fmpt.subcocycle_lim_ereal M Tinv (\n x. u n (((Tinv)^^n) x)) x = subcocycle_lim_ereal u x" proof - have bij: "bij T" using \invertible_qmpt\ unfolding invertible_qmpt_def by auto have int: "integrable M (u n)" for n using subcocycle_integrable[OF assms(1)] by simp interpret I: fmpt M Tinv using Tinv_fmpt[OF assms(2)] by simp have *: "AE x in M. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" for n using I.Invariants_of_foTn int unfolding o_def by simp have "AE x in M. \n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" apply (subst AE_all_countable) using * by simp moreover have "AE x in M. (\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" using kingman_theorem_AE_nonergodic_invariant_ereal[OF \subcocycle u\] by simp moreover have "AE x in M. (\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" using I.kingman_theorem_AE_nonergodic_invariant_ereal[OF subcocycle_u_Tinv[OF assms]] by simp moreover { fix x assume H: "\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x = real_cond_exp M I.Invariants (u n) x" "(\n. real_cond_exp M Invariants (u n) x / n) \ subcocycle_lim_ereal u x" "(\n. real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" have "ereal(real_cond_exp M Invariants (u n) x / n) = ereal(real_cond_exp M I.Invariants (\ x. u n (((Tinv)^^n) x)) x / n)" for n using H(1) Invariants_Tinv[OF \invertible_qmpt\] by auto then have "(\n. real_cond_exp M Invariants (u n) x / n) \ I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x" using H(3) by presburger then have "I.subcocycle_lim_ereal (\ n x. u n (((Tinv)^^n) x)) x = subcocycle_lim_ereal u x" using H(2) LIMSEQ_unique by auto } ultimately show ?thesis by auto qed proposition (in fmpt) subcocycle_lim_Tinv: assumes "subcocycle u" "invertible_qmpt" shows "AE x in M. fmpt.subcocycle_lim M Tinv (\n x. u n (((Tinv)^^n) x)) x = subcocycle_lim u x" proof - interpret I: fmpt M Tinv using Tinv_fmpt[OF assms(2)] by simp show ?thesis unfolding subcocycle_lim_def I.subcocycle_lim_def using subcocycle_lim_ereal_Tinv[OF assms] by auto qed end (*of Kingman.thy*) diff --git a/thys/MDP-Rewards/MDP_disc.thy b/thys/MDP-Rewards/MDP_disc.thy --- a/thys/MDP-Rewards/MDP_disc.thy +++ b/thys/MDP-Rewards/MDP_disc.thy @@ -1,925 +1,925 @@ (* Author: Maximilian Schäffeler *) theory MDP_disc - imports + imports MDP_cont "HOL-Library.Omega_Words_Fun" begin section \Markov Decision Processes with Discrete State Spaces\ (* counterpart to nn_integral_stream_space *) lemma (in prob_space) integral_stream_space: fixes f :: "'a stream \ ('b :: {banach, second_countable_topology,real_normed_vector})" assumes int_f: "integrable (stream_space M) f" assumes [measurable]: "f \ borel_measurable (stream_space M)" shows "(\X. f X \stream_space M) = (\x. (\X. f (x ## X) \stream_space M) \M)" proof - interpret S: sequence_space M .. interpret P: pair_sigma_finite M "\\<^sub>M i::nat\UNIV. M" .. interpret P': pair_sigma_finite "\\<^sub>M i::nat\UNIV. M" M .. obtain i where "has_bochner_integral (stream_space M) f i" using int_f using integrable.cases by blast have "integrable S.S (\X. f (to_stream X))" using int_f by (metis integrable_distr measurable_to_stream stream_space_eq_distr) - hence "integrable (distr (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (Pi\<^sub>M UNIV (\i. M)) + hence "integrable (distr (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (Pi\<^sub>M UNIV (\i. M)) (\(x, y). case_nat x y)) (\X. f (to_stream X))" by (auto simp: S.PiM_iter) - moreover have "integrable (distr (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (Pi\<^sub>M UNIV (\i. M)) (\(x, y). case_nat x y)) - (\X. f (to_stream X)) \ + moreover have "integrable (distr (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (Pi\<^sub>M UNIV (\i. M)) (\(x, y). case_nat x y)) + (\X. f (to_stream X)) \ integrable (M \\<^sub>M S.S) (\X. f (to_stream ((\(s, \). case_nat s \) X)))" by (auto simp: integrable_distr_eq) ultimately have "integrable (M \\<^sub>M S.S) (\X. f (to_stream ((\(s, \). case_nat s \) X)))" by auto - hence "integrable (M \\<^sub>M (Pi\<^sub>M UNIV (\i. M))) + hence "integrable (M \\<^sub>M (Pi\<^sub>M UNIV (\i. M))) (\X. f (to_stream ((\(s, \). case_nat s \) X)))" by auto - moreover have "integrable (M \\<^sub>M (Pi\<^sub>M UNIV (\i. M))) - (\X. f (to_stream ((\(s, \). case_nat s \) X))) = + moreover have "integrable (M \\<^sub>M (Pi\<^sub>M UNIV (\i. M))) + (\X. f (to_stream ((\(s, \). case_nat s \) X))) = integrable (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (\(x, X). f (to_stream (case_nat x X)))" - by (fastforce intro!: integrable_cong) + by (fastforce intro!: Bochner_Integration.integrable_cong) ultimately have *: "integrable (M \\<^sub>M Pi\<^sub>M UNIV (\i. M)) (\(x, X). f (to_stream (case_nat x X)))" by auto have "(\X. f X \stream_space M) = (\X. f (to_stream X) \S.S)" by (subst stream_space_eq_distr) (simp add: integral_distr) also have "\ = (\X. f (to_stream ((\(s, \). case_nat s \) X)) \(M \\<^sub>M S.S))" by (subst S.PiM_iter[symmetric]) (simp add: integral_distr) also have "\ = (\x. \X. f (to_stream ((\(s, \). case_nat s \) (x, X)))\S.S \M)" using * by (auto simp: pair_sigma_finite.integral_fst P.pair_sigma_finite_axioms case_prod_unfold) also have "\ = (\x. \X. f (x ## to_stream X) \S.S \M)" by (auto intro!: integral_cong simp: to_stream_nat_case) also have "\ = (\x. \X. f (x ## X) \distr (Pi\<^sub>M UNIV (\i. M)) (stream_space M) to_stream \M)" - by (subst Bochner_Integration.integral_cong[OF refl]) (auto simp: integral_distr) + by (subst Bochner_Integration.integral_cong[OF refl]) (auto simp: integral_distr) also have "\ = (\x. \X. f (x ## X) \stream_space M \M)" using stream_space_eq_distr by metis finally show ?thesis . qed -lemma prefix_cons: +lemma prefix_cons: "Omega_Words_Fun.prefix (Suc n) seq = seq 0#Omega_Words_Fun.prefix n (\n. seq (Suc n))" by (metis map_upt_Suc subsequence_def) lemma restrict_Suc: "restrict y {0..n. y (Suc n)) {0.. Pi\<^sub>M {0.. Pi\<^sub>M {0.._. count_space (UNIV :: ('s ::countable \ 'a::countable) set)) \\<^sub>M count_space UNIV" proof (induction i) case 0 then show ?case by simp next case (Suc i) - have aux: "(\w. (restrict w {0.. Pi\<^sub>M {0.._. count_space UNIV) \\<^sub>M + have aux: "(\w. (restrict w {0.. Pi\<^sub>M {0.._. count_space UNIV) \\<^sub>M Pi\<^sub>M {0.._. count_space UNIV) \\<^sub>M (count_space UNIV)" by auto - have aux': "(\(w,wi). Omega_Words_Fun.prefix i (restrict w {0.. Pi\<^sub>M {0..(w,wi). Omega_Words_Fun.prefix i (restrict w {0.. Pi\<^sub>M {0.._. count_space (UNIV :: ('s \ 'a) set)) \\<^sub>M (count_space UNIV) \\<^sub>M count_space UNIV" using Suc.IH by auto - have f_eq: "\w. Omega_Words_Fun.prefix i (restrict w {0..w. Omega_Words_Fun.prefix i (restrict w {0..(w,wi). Omega_Words_Fun.prefix i w @[wi]) ((restrict w {0..w:: nat \ 's \ 'a. Omega_Words_Fun.prefix i (restrict w {0.. Pi\<^sub>M {0.._. count_space UNIV) \\<^sub>M count_space UNIV" using aux aux'[unfolded prefix_restrict] by (subst f_eq) auto thus ?case unfolding prefix_restrict[of _ "i"] by auto qed no_notation Omega_Words_Fun.build (infixr \##\ 65) locale discrete_MDP = fixes A :: "'s::countable \ 'a::countable set" \ \enabled actions\ and K :: "'s \ 'a \ 's pmf" \ \MDP kernel, transition probabilities\ assumes A_ne: "\s. A s \ {}" \ \set of enabled actions is nonempty\ begin subsection \Policies\ text \Type synonym for decision rules.\ type_synonym ('c, 'd) dec = "'c \ 'd pmf" definition is_dec :: "('s, 'a) dec \ bool" where "is_dec d \ \s. d s \ A s" -lemma is_decI[intro]: +lemma is_decI[intro]: "(\s. set_pmf (d s) \ A s) \ is_dec d" unfolding is_dec_def by auto abbreviation "D\<^sub>R \ {d. is_dec d}" definition is_dec_det :: "('s \ 'a) \ bool" where "is_dec_det d \ \s. d s \ A s" abbreviation "D\<^sub>D \ {d. is_dec_det d}" definition "mk_dec_det d s = return_pmf (d s)" lemma is_dec_mk_dec_det_iff [simp]: "is_dec (mk_dec_det d) \ is_dec_det d" by (simp add: is_dec_def is_dec_det_def mk_dec_det_def) lemma D_det_to_MR[intro]: "is_dec_det d \ is_dec (mk_dec_det d)" by simp text \ Due to the assumption @{thm A_ne}, a deterministic decision rule always exists. It immediately follows via @{thm is_dec_mk_dec_det_iff} that a randomized decision rule also exists. \ lemma SOME_is_dec_det: "is_dec_det (\s. SOME a. a \ A s)" using A_ne by (simp add: is_dec_det_def some_in_eq) lemma ex_dec_det [simp]: "\d. is_dec_det d" using SOME_is_dec_det by blast lemma D_det_ne [simp]: "D\<^sub>D \ {}" by simp lemma D\<^sub>R_ne [simp]: "D\<^sub>R \ {}" using D_det_ne D_det_to_MR by blast lemma ex_dec[intro, simp]: "\d. is_dec d" using ex_dec_det by blast text \Type synonym for policies.\ type_synonym ('c, 'd) pol = "('c \ 'd) list \ ('c, 'd) dec" text \A policy assigns a decision rule to each observed past.\ definition is_policy :: "('s, 'a) pol \ bool" where "is_policy p \ \hs. is_dec (p hs)" abbreviation "\\<^sub>H\<^sub>R \ {p. is_policy p}" text \Deterministic policies\ definition "is_deterministic p \ is_policy p \ (\h s. \a. p h s = return_pmf a)" definition "mk_det p h s \ return_pmf (p h s)" abbreviation "\\<^sub>H\<^sub>D \ {p. \h. p h \ D\<^sub>D}" text \Markovian policies\ definition "is_markovian p \ is_policy p \ (\h h'. length h = length h' \ p h = p h')" definition mk_markovian :: "(nat \ ('s, 'a) dec) \ ('s, 'a) pol" where "mk_markovian p \ (\h. p (length h))" lemma is_markovian_mk_iff[simp]: "is_markovian (mk_markovian p) \ (\n. is_dec (p n))" unfolding is_markovian_def mk_markovian_def is_policy_def by (metis (mono_tags, opaque_lifting) Ex_list_of_length) lemma is_markovian_mk[intro]: "\n. is_dec (p n) \ is_markovian (mk_markovian p)" unfolding is_markovian_def mk_markovian_def is_policy_def by auto lemma mk_markovian_nil [simp]: "mk_markovian p [] = p 0" unfolding mk_markovian_def by auto definition "mk_markovian_det p \ (\h s. return_pmf (p (length h) s))" abbreviation "\\<^sub>M\<^sub>D \ {p. \n. p n \ D\<^sub>D}" abbreviation "\\<^sub>M\<^sub>R \ {p. \n. p n \ D\<^sub>R}" lemma \\<^sub>M\<^sub>R_imp_policies[intro]: "p \ \\<^sub>M\<^sub>R \ mk_markovian p \ \\<^sub>H\<^sub>R" unfolding is_policy_def mk_markovian_def by auto lemma \\<^sub>M\<^sub>D_MR_iff[simp]: "(\n. mk_dec_det (p n)) \ \\<^sub>M\<^sub>R \ p \ \\<^sub>M\<^sub>D" by auto lemma \\<^sub>M\<^sub>D_to_MR[intro]: "p \ \\<^sub>M\<^sub>D \ (\n. mk_dec_det (p n)) \ \\<^sub>M\<^sub>R" by simp lemma \\<^sub>M\<^sub>D_ne[simp]: "\\<^sub>M\<^sub>D \ {}" by (auto simp: someI_ex[OF ex_dec_det] intro: exI[of _ "\n. (SOME d. is_dec_det d)"]) lemma \\<^sub>M\<^sub>R_ne[simp]: "\\<^sub>M\<^sub>R \ {}" using \\<^sub>M\<^sub>D_ne by fast lemma policies_ne[simp, intro]: "\\<^sub>H\<^sub>R \ {}" using \\<^sub>M\<^sub>R_ne is_policy_def by auto text \Stationary policies\ definition "is_stationary p \ is_policy p \ (\h h'. p h = p h')" lemma is_stationary_const_iff[simp]: "is_stationary (\_. d) = is_dec d" unfolding is_stationary_def is_policy_def by simp lemma is_stationary_const[intro]: "is_dec d \ is_stationary (\_. d)" by simp abbreviation "mk_stationary p \ mk_markovian (\_. p)" abbreviation "mk_stationary_det d \ mk_markovian (\_. mk_dec_det d)" subsubsection \Successor Policy\ text \ -After taking the first step in the MDP, we will know which state and which action got selected -during the initial epoch. To obtain a policy that acts as if the current epoch was the initial one, -we prepend the observed state-action pair to the history. The result is again a policy, +After taking the first step in the MDP, we will know which state and which action got selected +during the initial epoch. To obtain a policy that acts as if the current epoch was the initial one, +we prepend the observed state-action pair to the history. The result is again a policy, i.e. it satisfies @{const is_policy}. \ definition "\_Suc p sa h = p (sa#h)" lemma is_policy_\_Suc [intro]: "is_policy p \ is_policy (\_Suc p sa)" unfolding is_policy_def \_Suc_def by force lemma Suc_mk_markovian[simp]: "\_Suc (mk_markovian p) x = mk_markovian (\n. p (Suc n))" unfolding \_Suc_def mk_markovian_def by auto subsection \Stream Space of the MDP\ subsubsection \Initial State-Action Distribution\ text \ -If we fix a decision rule @{term d} and an initial distribution of states @{term "S0"}, +If we fix a decision rule @{term d} and an initial distribution of states @{term "S0"}, we obtain a distribution over state-action pairs in the following way: -First, the initial state @{term "s"} is sampled from @{term S0}, +First, the initial state @{term "s"} is sampled from @{term S0}, then an action @{term a} is selected from @{term "d s"}. \ definition "K0 d S0 = do { - s \ S0; + s \ S0; a \ d s; return_pmf (s,a) -}" +}" notation K0 ("K\<^sub>0") lemma K0_iff: "K0 d S0 = S0 \ (\s. map_pmf (\a. (s,a)) (d s))" by (simp add: K0_def map_pmf_def) lemma vimage_pair[simp]: "Pair x -` {p} = (if x = fst p then {snd p} else {})" by auto lemma pmf_K0 [simp]: "pmf (K0 d S0) (s,a) = pmf S0 s * pmf (d s) a" unfolding K0_iff pmf_bind by (subst integral_measure_pmf[where A = "{s}"]) (auto simp: pmf_map pmf.rep_eq split: if_splits) lemma set_pmf_K0: "set_pmf (K0 p S0) = {(s,a). s \ S0 \ a \ p s}" by (auto simp add: K0_def) lemma fst_K0[simp]: "map_pmf fst (K0 p S0) = S0" unfolding K0_def by (simp add: map_bind_pmf map_pmf_comp bind_return_pmf') abbreviation "S \ stream_space (count_space UNIV)" text \We inherit the trace space from MDPs with continuous state-action spaces\ interpretation MDP_cont: MDP_cont.discrete_MDP "count_space UNIV" "count_space UNIV" A K proof standard - show "(\x. measure_pmf (K x)) \ + show "(\x. measure_pmf (K x)) \ count_space UNIV \\<^sub>M count_space UNIV \\<^sub>M prob_algebra (count_space UNIV)" using measurable_prob_algebraI by (measurable, auto simp: prob_space_measure_pmf measurable_pair_measure_countable1)+ show "\\\count_space UNIV \\<^sub>M count_space UNIV. \s. \ s \ A s" by (auto simp: A_ne some_in_eq intro: bexI[of _ "\s. SOME a. a \ A s"]) qed (auto simp: A_ne) lemma count_space_M[simp]: "MDP_cont.M = count_space UNIV" by (auto simp: pair_measure_countable) lemma space_M[simp]: "space MDP_cont.M = UNIV" by (auto simp: MDP_cont.space_lim_stream) text \We reuse the stream space provided by @{const MDP_cont.lim_stream}\ definition T :: "('s, 'a) pol \ 's pmf \ ('s \ 'a) stream measure" where "T p = MDP_cont.lim_stream (\n (h,s). p (Omega_Words_Fun.prefix n h) s)" -lemma sets_T[measurable_cong]: +lemma sets_T[measurable_cong]: "sets (T p x) = sets S" by (auto simp: T_def MDP_cont.sets_lim_stream) lemma space_stream_space_ne[simp]: "space S \ {}" by (auto simp: space_stream_space) lemma space_T[simp]: "space (T p S0) = space S" by (simp add: MDP_cont.space_lim_stream T_def space_stream_space) -lemma is_policy_MDP_cont[intro]: +lemma is_policy_MDP_cont[intro]: fixes p :: "('s \ 'a) list \ 's \ 'a pmf" shows "MDP_cont.is_policy (\n (h,s). p (Omega_Words_Fun.prefix n h) s)" unfolding MDP_cont.is_policy_def MDP_cont.is_dec_def using prefix_measurable measurable_pair_swap_iff by (auto simp: prob_space_measure_pmf intro: measurable_pair_measure_countable1 measurable_prob_algebraI) lemma prob_space_T[intro, simp]: "prob_space (T p x)" by (auto simp add: T_def prob_space_measure_pmf space_prob_algebra) -lemma T_subprob[simp]: +lemma T_subprob[simp]: "T p S0 \ space (subprob_algebra S)" by (metis prob_space.M_in_subprob prob_space_T sets_T subprob_algebra_cong) lemma T_subprob_space [simp]: "subprob_space (T p S0)" by (auto intro: prob_space_imp_subprob_space) -lemma K0_MDP_cont_eq: - "MDP_cont.K0 (\x (h,s). measure_pmf (p (Omega_Words_Fun.prefix x h) s)) (measure_pmf S0) = +lemma K0_MDP_cont_eq: + "MDP_cont.K0 (\x (h,s). measure_pmf (p (Omega_Words_Fun.prefix x h) s)) (measure_pmf S0) = K0 (p []) S0" unfolding MDP_cont.K0_def K0_def MDP_cont.K'_def map_pmf_def by (simp add: measure_pmf_bind return_pmf.rep_eq) subsubsection \Decomposition of the Stream Space\ text \ The distribution of traces/walks the MDP allows should intuitively satisfy the following rule: \<^enum> select the initial state @{term s} from @{term S0} \<^enum> pass it to the decision rule @{term "p []"} to determine a distribution over actions \<^enum> select the action @{term a} -\<^item> finally pass the state-action pair @{term "(s,a)"} to the kernel @{term K} to get a new +\<^item> finally pass the state-action pair @{term "(s,a)"} to the kernel @{term K} to get a new distribution over states @{term s0'} Then the iteration repeats with the updated policy @{term "\_Suc p (s,a)"}. The result carries over from @{thm MDP_cont.lim_stream_eq}. \ lemma T_eq: shows "T p S0 = do { sa \ measure_pmf (K0 (p []) S0); \ \ T (\_Suc p sa) (K sa); return S (sa ## \) }" unfolding T_def proof (subst MDP_cont.lim_stream_eq) show "MDP_cont.is_policy (\x xa. measure_pmf (case xa of (h, xa) \ p (Omega_Words_Fun.prefix x h) xa))" by auto -qed (auto simp: space_prob_algebra prob_space_measure_pmf \_Suc_def MDP_cont.Suc_policy_def +qed (auto simp: space_prob_algebra prob_space_measure_pmf \_Suc_def MDP_cont.Suc_policy_def prefix_cons K0_MDP_cont_eq prod.case_distrib) lemma T_eq_distr: shows "T p S0 = measure_pmf (K0 (p []) S0) \ (\sa. distr (T (\_Suc p sa) (K sa)) S ((##) sa))" by (simp add: T_eq[symmetric] bind_return_distr'[symmetric]) text \ -The iteration rule lets us nicely decompose integrals (expected values) over functions on traces of +The iteration rule lets us nicely decompose integrals (expected values) over functions on traces of the MDP. \ lemma integral_T: fixes f :: "('s \ 'a) stream \ real" assumes f_bounded: "\x. \f x\ \ B" assumes f: "f \ borel_measurable S" shows "(\t. f t \T p x) = \sa. \t'. f (sa##t') \T (\_Suc p sa) (K sa) \K0 (p []) x" proof - note T_eq_distr have "(\t. f t \T p x) = (\t. f t \measure_pmf (K\<^sub>0 (p []) x) \ (\sa. distr (T (\_Suc p sa) (K sa)) (stream_space (count_space UNIV)) ((##) sa)))" using T_eq_distr by metis also have "\ = measure_pmf.expectation (K\<^sub>0 (p []) x) (\sa. LINT t'|T (\_Suc p sa) (K sa). f (sa ## t'))" proof (subst integral_bind[OF f f_bounded, where B' = 1], goal_cases) case 1 - then show ?case - by (auto intro!: prob_space_imp_subprob_space prob_space.prob_space_distr + then show ?case + by (auto intro!: prob_space_imp_subprob_space prob_space.prob_space_distr simp: space_subprob_algebra) next case 3 then show ?case by (auto intro!: prob_space.emeasure_le_1 prob_space.prob_space_distr) next case 4 then show ?case by (auto simp: f integral_distr intro: Bochner_Integration.integral_cong) qed auto finally show ?thesis. qed lemma nn_integral_T: assumes f: "f \ borel_measurable S" shows "(\\<^sup>+t. f t \T p x) = (\\<^sup>+sa. \\<^sup>+ t'. f (sa##t') \T (\_Suc p sa) (K sa) \K0 (p []) x)" unfolding T_eq_distr[of p] - by (subst nn_integral_bind[OF f]) - (auto intro!: prob_space_imp_subprob_space prob_space.prob_space_distr + by (subst nn_integral_bind[OF f]) + (auto intro!: prob_space_imp_subprob_space prob_space.prob_space_distr simp: f nn_integral_distr space_subprob_algebra) subsubsection \A Denotational View on the Stochastic Process\ text \ -Many definitions on MDPs do not rely on the individual traces but only on the distribution +Many definitions on MDPs do not rely on the individual traces but only on the distribution of states and actions at each epoch. We define this view on the trace space as the repeated iteration of @{const K0} and @{term K}. It conincides with the definition of @{const T}. \ primrec Pn :: "('s, 'a) pol \ 's pmf \ nat \ ('s \ 'a) pmf" where "Pn p S0 0 = K0 (p []) S0" | "Pn p S0 (Suc n) = K0 (p []) S0 \ (\sa. Pn (\_Suc p sa) (K sa) n)" declare Pn.simps(2)[simp del] lemma Pn_eq_T: shows "measure_pmf (Pn p S0 n) = distr (T p S0) (count_space UNIV) (\t. t !! n)" proof (induction n arbitrary: p S0) case (0 p S0) then show ?case unfolding T_eq[of p] proof (subst distr_bind[where K = S], goal_cases) case 1 - then show ?case + then show ?case by (auto intro!: prob_space_imp_subprob_space subprob_space.bind_in_space) next case 4 - then show ?case + then show ?case by (subst bind_cong[OF refl, where g = "return (count_space UNIV)"]) (auto intro!: bind_const' simp: distr_bind[where K = S] distr_return bind_return'' space_stream_space subprob_space_return_ne) qed auto next case (Suc n) show ?case unfolding T_eq[of p] proof (subst distr_bind[where K = S], goal_cases) case 1 - then show ?case + then show ?case by (auto intro!: prob_space_imp_subprob_space subprob_space.bind_in_space)[1] next case 4 then show ?case by (auto simp: Pn.simps(2) measure_pmf_bind Suc bind_return_distr' distr_distr comp_def intro!: bind_cong) qed auto qed text \ The definition of @{const Pn} also allows us to easily prove that only enabled actions can occur in the traces of the MDP. \ lemma Pn_in_A: "is_policy p \ (s, a) \ Pn p S0 n \ a \ A s" proof (induction n arbitrary: S0 p) case 0 - then show ?case + then show ?case using 0 unfolding is_policy_def is_dec_def by (auto simp: K0_def) next case (Suc n) then show ?case by (auto simp: Pn.simps(2) K0_def) qed lemma T_in_A: assumes "is_policy p" shows "AE t in T p S0. snd (t !! n) \ A (fst (t !! n))" proof - have aux: "AE t in distr (T p S0) (count_space UNIV) (\t. t !! n). snd t \ A (fst t)" using assms Pn_eq_T[symmetric] by (auto simp: Pn_in_A intro!: AE_pmfI cong: AE_cong_simp) show ?thesis by (auto intro!: AE_distrD[OF _ aux]) qed subsubsection \State Process\ text \Alongside @{const Pn}, we also define the state and action distributions as projections.\ definition "Xn p S0 n = map_pmf fst (Pn p S0 n)" lemma X0 [simp]: "Xn p S0 0 = S0" using fst_K0 Xn_def by auto lemma Xn_Suc: "Xn p S0 (Suc n) = Pn p S0 n \ K" proof (induction n arbitrary: p S0) case 0 then show ?case by (simp add: Pn.simps(2) Xn_def map_bind_pmf) next case (Suc n) then show ?case by (simp add: Pn.simps(2) Xn_def map_bind_pmf bind_assoc_pmf) qed lemma Pn_markovian_eq_Xn_bind: "Pn (mk_markovian p) S0 n = K0 (p n) (Xn (mk_markovian p) S0 n)" proof (induction n arbitrary: p S0) case 0 - then show ?case + then show ?case unfolding Xn_def by auto next case (Suc n) then show ?case unfolding Xn_def K0_def by (auto intro!: bind_pmf_cong simp: Pn.simps(2) map_bind_pmf Suc bind_assoc_pmf) qed lemma Xn_Suc': "Xn p S0 (Suc n) = K0 (p []) S0 \ (\sa. Xn (\_Suc p sa) (K sa) n)" unfolding Xn_def by (auto simp: Pn.simps(2) map_bind_pmf) lemma set_pmf_X0 [simp]: "set_pmf (Xn p S0 0) = S0" using X0 by auto -lemma set_pmf_PSuc: "set_pmf (Pn (mk_markovian p) S0 n) = +lemma set_pmf_PSuc: "set_pmf (Pn (mk_markovian p) S0 n) = {(s, a). s \ set_pmf (Xn (mk_markovian p) S0 n) \ a \ p n s}" using set_pmf_K0 Pn_markovian_eq_Xn_bind by auto subsubsection \The Conditional Distribution of Actions\ text \ Actions are selected wrt. the whole history of state-action pairs encountered so far. The following definition defines the expected action selection when only the current state is given.\ definition "Y_cond_X p S0 n x = map_pmf snd (cond_pmf (Pn p S0 n) {(s,a). s = x})" lemma prob_K0_X [simp]: "measure_pmf.prob (K0 p S0) {(s, a). s = x} = pmf S0 x" unfolding K0_iff proof (subst measure_pmf_bind, subst measure_pmf.measure_bind[of _ _ "count_space UNIV"], goal_cases) case 1 then show ?case by (simp add: measure_pmf_in_subprob_algebra) next case 3 - then show ?case + then show ?case by (subst integral_measure_pmf_real[of "{x}"]) (auto split: if_splits) -qed simp +qed simp lemma prob_Pn_X[simp]: "measure_pmf.prob (Pn p S0 n) {(s, a). s = x} = pmf (Xn p S0 n) x" proof (induction n arbitrary: p S0) case 0 then show ?case by auto next case (Suc n) show ?case - unfolding Xn_Suc' Pn.simps(2) measure_pmf_bind + unfolding Xn_Suc' Pn.simps(2) measure_pmf_bind using Suc by (simp add: measure_pmf.measure_bind[of _ _ "count_space UNIV"] K0_def measure_pmf_in_subprob_algebra pmf_bind) qed lemma pmf_Pn_pair: assumes "sa \ set_pmf (Pn p S0 n)" shows "pmf (Pn p S0 n) sa = pmf (Y_cond_X p S0 n (fst sa)) (snd sa) * pmf (Xn p S0 n) (fst sa)" proof - have aux: "set_pmf (Pn p S0 n) \ {(s, a). s = fst sa} \ {}" - using Xn_def assms + using Xn_def assms by auto have aux': "({(s, a). s = fst sa} \ snd -` {snd sa}) = {sa}" by auto show ?thesis using assms unfolding Y_cond_X_def pmf_map cond_pmf.rep_eq[OF aux] by (auto simp: Xn_def pmf_eq_0_set_pmf measure_pmf.emeasure_eq_measure aux' measure_pmf_single) qed lemma pmf_Pn: assumes "x \ set_pmf (Xn p S0 n)" shows "pmf (Pn p S0 n) (x,a) = pmf (Y_cond_X p S0 n x) a * pmf (Xn p S0 n) x" proof - have aux: "set_pmf (Pn p S0 n) \ {(s, a). s = x} \ {}" using Xn_def assms by auto have aux': "({(s, a). s = x} \ snd -` {a}) = {(x, a)}" by auto show ?thesis using assms unfolding Y_cond_X_def cond_pmf.rep_eq[OF aux] pmf_map by (auto simp: pmf_eq_0_set_pmf measure_pmf.emeasure_eq_measure aux' measure_pmf_single) qed lemma pmf_Y_cond_X: assumes "x \ set_pmf (Xn p S0 n)" shows "pmf (Y_cond_X p S0 n x) a = pmf (Pn p S0 n) (x,a) / pmf (Xn p S0 n) x" proof - have aux: "set_pmf (Pn p S0 n) \ {(s, a). s = x} \ {}" using Xn_def assms by auto have aux': "({(s, a). s = x} \ snd -` {a}) = {(x, a)}" by auto show ?thesis using assms aux' unfolding Y_cond_X_def - by (auto simp: cond_pmf.rep_eq[OF aux] pmf_map pmf_eq_0_set_pmf measure_pmf.emeasure_eq_measure + by (auto simp: cond_pmf.rep_eq[OF aux] pmf_map pmf_eq_0_set_pmf measure_pmf.emeasure_eq_measure measure_pmf_single) qed lemma Y_cond_X_0[simp]: assumes "x \ set_pmf S0" shows "Y_cond_X p S0 0 x = p [] x" by (auto intro: pmf_eqI simp: assms pmf_Y_cond_X pmf_eq_0_set_pmf) (* eqn 5.5.3 in Puterman *) lemma Y_cond_X_markovian[simp]: assumes h: "x \ Xn (mk_markovian p) S0 n" shows "Y_cond_X (mk_markovian p) S0 n x = p n x" by (auto intro!: pmf_eqI simp: pmf_Y_cond_X h Pn_markovian_eq_Xn_bind pmf_eq_0_set_pmf) lemma Pn_eq_Xn_Y_cond: "Pn p S0 n = Xn p S0 n \ (\x. map_pmf (\a. (x, a)) (Y_cond_X p S0 n x))" proof (induction n) case 0 - then show ?case + then show ?case by (auto simp: K0_iff intro: bind_pmf_cong) next case (Suc n) show ?case proof (intro pmf_eqI; safe) - fix a :: 's + fix a :: 's fix b :: 'a - have aux': "pmf (Xn p S0 (Suc n) \ (\x. map_pmf (Pair x) (Y_cond_X p S0 (Suc n) x))) (a,b) - = measure_pmf.expectation (Pn p S0 (Suc n)) (\x. + have aux': "pmf (Xn p S0 (Suc n) \ (\x. map_pmf (Pair x) (Y_cond_X p S0 (Suc n) x))) (a,b) + = measure_pmf.expectation (Pn p S0 (Suc n)) (\x. if fst x = a then pmf (Y_cond_X p S0 (Suc n) a) b else 0)" - by (auto intro!: Bochner_Integration.integral_cong[OF refl] + by (auto intro!: Bochner_Integration.integral_cong[OF refl] simp: Xn_def bind_map_pmf pmf_map pmf_bind measure_pmf_single) also have "\ = measure_pmf.expectation (Pn p S0 (Suc n)) (\x. indicator {(s',a'). s' = a} x * (pmf (Pn p S0 (Suc n)) (a, b) / pmf (Xn p S0 (Suc n)) a))" proof (intro Bochner_Integration.integral_cong_AE AE_pmfI) fix y assume h: "y \ set_pmf (Pn p S0 (Suc n))" hence h': "fst y \ set_pmf (Xn p S0 (Suc n))" by (metis mult_eq_0_iff pmf_Pn_pair pmf_eq_0_set_pmf) - show "(if fst y = a then pmf (Y_cond_X p S0 (Suc n) a) b else 0) = - indicat_real {(s', a'). s' = a} y * + show "(if fst y = a then pmf (Y_cond_X p S0 (Suc n) a) b else 0) = + indicat_real {(s', a'). s' = a} y * (pmf (Pn p S0 (Suc n)) (a, b) / pmf (Xn p S0 (Suc n)) a)" by (auto simp: case_prod_beta' pmf_Y_cond_X[of "fst y" p S0 "(Suc n)" b, OF h']) qed auto - also have "\ = measure_pmf.prob (Pn p S0 (Suc n)) {(s',a'). s' = a} * + also have "\ = measure_pmf.prob (Pn p S0 (Suc n)) {(s',a'). s' = a} * pmf (Pn p S0 (Suc n)) (a, b) / pmf (Xn p S0 (Suc n)) a" by auto also have "\ = pmf (Pn p S0 (Suc n)) (a,b)" using prob_Pn_X Xn_def pmf_Pn_pair pmf_eq_0_set_pmf by fastforce - finally show "pmf (Pn p S0 (Suc n)) (a, b) = pmf (Xn p S0 (Suc n) \ + finally show "pmf (Pn p S0 (Suc n)) (a, b) = pmf (Xn p S0 (Suc n) \ (\x. map_pmf (Pair x) (Y_cond_X p S0 (Suc n) x))) (a, b)" by auto qed qed -lemma Pn_eq_Xn_Y_cond': +lemma Pn_eq_Xn_Y_cond': "Pn p S0 n = Xn p S0 n \ (\s. Y_cond_X p S0 n s \ (\a. return_pmf (s,a)))" by (metis K0_def K0_iff Pn_eq_Xn_Y_cond) -lemma Pn_markovian_Suc: "Pn (mk_markovian p) S0 (Suc n) = +lemma Pn_markovian_Suc: "Pn (mk_markovian p) S0 (Suc n) = Pn (mk_markovian p) S0 n \ (\sa. K0 (p (Suc n)) (K sa))" proof (induction n arbitrary: S0 p) case 0 then show ?case by (auto intro: bind_pmf_cong simp: Pn.simps(2) \_Suc_def) next case (Suc n) show ?case by (auto simp add: Suc bind_assoc_pmf Pn.simps(2)[of _ S0] intro: bind_pmf_cong) qed subsubsection \Action Process\ text \The distribution of actions.\ definition "Yn p S0 n = map_pmf snd (Pn p S0 n)" lemma Y0: "Yn p S0 0 = S0 \ p []" by (simp add: Yn_def K0_iff map_bind_pmf map_pmf_comp) text \ For markovian policies, the decision rules at each epoch are independent of each other, hence we may express @{const Yn} solely in terms of @{const Xn} and the current decision rule. \ lemma Yn_markovian: "Yn (mk_markovian p) S0 n = Xn (mk_markovian p) S0 n \ p n" -proof (induction n arbitrary: p S0) +proof (induction n arbitrary: p S0) case 0 then show ?case by (auto simp: Y0) next case (Suc n) then show ?case by (simp add: Xn_def Yn_def map_bind_pmf Suc Pn.simps(2) bind_assoc_pmf) qed subsection \Restriction to Markovian Policies\ -abbreviation "as_markovian p S0 n x \ +abbreviation "as_markovian p S0 n x \ if x \ (Xn p S0 n) then Y_cond_X p S0 n x else return_pmf (SOME a. a \ A x)" text \ For states which cannot occur we choose an arbitrary enabled action, as in this case we cannot make any statements about @{const Y_cond_X} (a distribution conditioned on an event with probability 0). \ lemma is_\\<^sub>M\<^sub>R_as_markovian: - assumes p: "is_policy p" + assumes p: "is_policy p" shows "as_markovian p S0 \ \\<^sub>M\<^sub>R" proof - have aux: "\hs s. s \ set_pmf (Xn p S0 hs) \ set_pmf ((Pn p S0 hs)) \ {(s', a). s' = s} \ {}" by (simp add: measure_pmf_zero_iff[symmetric] pmf_eq_0_set_pmf) thus ?thesis using assms A_ne Pn_in_A unfolding is_dec_def Y_cond_X_def by (auto simp: some_in_eq) qed lemma is_policy_as_markovian: "is_policy p \ is_policy (mk_markovian (as_markovian p S0))" using is_\\<^sub>M\<^sub>R_as_markovian \\<^sub>M\<^sub>R_imp_policies by auto theorem Pn_as_markovian_eq: "Pn (mk_markovian (as_markovian p S0)) S0 = Pn p S0" -proof +proof fix n show "Pn (mk_markovian (as_markovian p S0)) S0 n = Pn p S0 n" proof (induction n) case 0 thus ?case by (auto intro!: map_pmf_cong bind_pmf_cong simp: K0_def) next case (Suc n) - have "\x. x \ Xn p S0 (Suc n) \ + have "\x. x \ Xn p S0 (Suc n) \ Y_cond_X (mk_markovian (as_markovian p S0)) S0 (Suc n) x = Y_cond_X p S0 (Suc n) x" by (auto simp: Suc.IH Xn_Suc) moreover have "Xn (mk_markovian (as_markovian p S0)) S0 (Suc n) = Xn p S0 (Suc n)" by (simp add: Xn_Suc Suc.IH) ultimately show "Pn (mk_markovian (as_markovian p S0)) S0 (Suc n) = Pn p S0 (Suc n)" by (auto intro: bind_pmf_cong simp: Pn_eq_Xn_Y_cond) qed qed subsection \MDPs without Initial Distribution\ text \ From now on, we assume a known, deterministic initial state. All results from the previous discussion carry over as we are now in the special case where we the initial state is of the form @{term "return_pmf s"}. \ definition "\ p s \ T p (return_pmf s)" -lemma \_eq_return_distr: "\ p s = +lemma \_eq_return_distr: "\ p s = measure_pmf (p [] s) \ (\a. distr (T (\_Suc p (s,a)) (K (s,a))) S ((##) (s,a)))" unfolding \_def - by (subst T_eq_distr) (fastforce intro!: bind_distr subprob_space.subprob_space_distr + by (subst T_eq_distr) (fastforce intro!: bind_distr subprob_space.subprob_space_distr simp: K0_iff map_pmf_rep_eq space_subprob_algebra bind_return_pmf)+ lemma \_eq_return: shows "\ p s = do { y \ measure_pmf (p [] s); - \ \ T (\_Suc p (s,y)) (K (s,y)); + \ \ T (\_Suc p (s,y)) (K (s,y)); return S ((s,y) ## \) }" by (auto simp: \_eq_return_distr bind_return_distr' prob_space.not_empty intro!: bind_cong) lemma \_return: shows "T p S0 = measure_pmf S0 \ \ p" proof - - have "T p S0 = measure_pmf S0 \ (\x. measure_pmf (map_pmf (Pair x) (p [] x))) \ + have "T p S0 = measure_pmf S0 \ (\x. measure_pmf (map_pmf (Pair x) (p [] x))) \ (\sa. distr (T (\_Suc p sa) (K sa)) (stream_space (count_space UNIV)) ((##) sa))" unfolding T_eq_distr[of p] K0_iff measure_pmf_bind by auto also have "\ = measure_pmf S0 \ (\x. distr (measure_pmf (p [] x)) (count_space UNIV) (Pair x) \ (\sa. distr (T (\_Suc p sa) (K sa)) (stream_space (count_space UNIV)) ((##) sa)))" - using measurable_measure_pmf + using measurable_measure_pmf by (subst bind_assoc[where N = "count_space UNIV", where R = S]) - (fastforce intro!: prob_space_imp_subprob_space prob_space.prob_space_distr + (fastforce intro!: prob_space_imp_subprob_space prob_space.prob_space_distr simp: space_subprob_algebra prob_space_measure_pmf map_pmf_rep_eq)+ also have "\ = measure_pmf S0 \ \ p" by (subst bind_distr[where K = S]) (auto intro!: prob_space_imp_subprob_space prob_space.prob_space_distr bind_cong simp: space_subprob_algebra \_eq_return_distr) finally show ?thesis. qed lemma \_return_eq: shows "\ p s = do { a \ measure_pmf (p [] s); s' \ measure_pmf (K (s,a)); w \ T (\_Suc p (s,a)) (return_pmf s'); return S ((s,a)##w) }" proof - have "\ p s = do { a \ measure_pmf (p [] s); \ \ T (\_Suc p (s, a)) (K (s, a)); return S ((s, a) ## \)}" - using \_eq_return + using \_eq_return by auto also have "\ = do { a \ measure_pmf (p [] s); s' \ measure_pmf (K (s, a)); \ \ T (\_Suc p (s, a)) (return_pmf s'); return S ((s, a) ## \)}" unfolding \_return by (subst bind_assoc[of _ _ S _ S]) (auto simp add: \_def \_return[symmetric]) finally show ?thesis. qed lemma \_eq: shows "\ p s = do { a \ measure_pmf (p [] s); s' \ measure_pmf (K (s,a)); w \ \ (\_Suc p (s,a)) s'; return S ((s,a)##w) }" by (subst \_return_eq) (auto simp add: \_def ) lemma \_prob_space[intro]: "prob_space (\ p s)" by (metis \_def prob_space_T) -lemma \_sets[measurable_cong]: +lemma \_sets[measurable_cong]: "sets (\ p s) = sets S" by (simp add: \_def sets_T) -lemma measurable_ident_Suc'[measurable]: +lemma measurable_ident_Suc'[measurable]: "(\x. x) \ \ (\_Suc p sa) s' \\<^sub>M S" by (simp add: \_def) -lemma nn_integral_\: +lemma nn_integral_\: fixes f :: "('s \ 'a) stream \ real" assumes f[measurable]: "f \ borel_measurable S" - shows "(\\<^sup>+t. f t \\ p s) + shows "(\\<^sup>+t. f t \\ p s) = \\<^sup>+a. \\<^sup>+s'. \\<^sup>+t'. f ((s,a)##t') \\ (\_Suc p (s,a)) s' \K (s,a) \p [] s" proof - - have "(\\<^sup>+t. f t \\ p s) = + have "(\\<^sup>+t. f t \\ p s) = \\<^sup>+ x. \\<^sup>+ y. (f y) \measure_pmf (K (s, x)) \ (\s'. \ (\_Suc p (s, x)) s' \ (\w. return S ((s, x) ## w))) \(p [] s)" unfolding \_eq[of p] by (subst nn_integral_bind[of _ S]) (auto intro!: measure_pmf.bind_in_space subprob_space.bind_in_space simp: \_prob_space prob_space_imp_subprob_space) also have "\ = \\<^sup>+ x. \\<^sup>+ xa. \\<^sup>+ y. (f y) \\ (\_Suc p (s, x)) xa \ (\w. return S ((s, x) ## w)) \measure_pmf (K (s, x)) \(p [] s)" by (subst nn_integral_bind[of _ S]) (auto intro!: subprob_space.bind_in_space simp: \_prob_space prob_space_imp_subprob_space) also have "\ = \\<^sup>+ x. \\<^sup>+ xa. \\<^sup>+ y. (f y) \distr (\ (\_Suc p (s, x)) xa) S ((##) (s, x)) \measure_pmf (K (s, x)) \(p [] s)" by (auto simp add: bind_return_distr' \_prob_space prob_space.not_empty) also have "\ = \\<^sup>+ x. \\<^sup>+ xa. \\<^sup>+ xa. (f ((s, x) ## xa)) \\ (\_Suc p (s, x)) xa \measure_pmf (K (s, x)) \(p [] s)" by (auto simp: nn_integral_distr) finally show ?thesis. qed -lemma integral_\: +lemma integral_\: fixes f :: "('s \ 'a) stream \ real" assumes f_bounded: "\x. \f x\ \ B" assumes f[measurable]: "f \ borel_measurable S" - shows "(\t. f t \\ p s) + shows "(\t. f t \\ p s) = \a. \s'. \t'. f ((s,a)##t') \\ (\_Suc p (s,a)) s' \K (s,a) \p [] s" unfolding \_def integral_T[OF f_bounded f] K0_iff bind_return_pmf unfolding \_return[of "\_Suc p _"] integral_map_pmf using \_return[of "\_Suc p _", symmetric] - by (subst integral_bind[OF _ f_bounded, where B' = 1, where K = S]) + by (subst integral_bind[OF _ f_bounded, where B' = 1, where K = S]) (auto simp: \_def intro: prob_space.emeasure_le_1) lemma integrable_\_bounded[intro]: fixes f :: "('s \ 'a) stream \ 'd :: {second_countable_topology,banach}" assumes f[measurable]: "f \ borel_measurable S" assumes b: "bounded (range f)" shows "integrable (\ p s) f" using b - by (auto simp: prob_space.finite_measure \_prob_space bounded_iff + by (auto simp: prob_space.finite_measure \_prob_space bounded_iff intro!: finite_measure.integrable_const_bound) definition "Pn' p s = Pn p (return_pmf s)" definition "Xn' p s = Xn p (return_pmf s)" definition "Yn' p s = Yn p (return_pmf s)" definition "K0' d s \ map_pmf (\a. (s, a)) (d s)" definition "K_st d s \ d s \ (\a. K (s,a))" lemma pmf_K_st: "pmf (K_st d s) t = \a. pmf (K(s, a)) t \d s" unfolding K_st_def by (subst pmf_bind) auto text \ @{const K_st} defines the distribution over the successor states for a given decision rule and state. It is mostly useful for markovian policies, as the information which action was selected is lost.\ lemma P0'[simp]: "Pn' p s 0 = K0' (p []) s" by (simp add: Pn'_def K0'_def K0_iff bind_return_pmf) lemma X0'[simp]: "Xn' p s 0 = return_pmf s" using X0 Xn'_def by auto lemma Pn_return_pmf: "S0 \ (\s'. Pn p (return_pmf s') n) = Pn p S0 n" by (induction n arbitrary: p S0) (auto intro: bind_pmf_cong simp add: Pn.simps(2) K0_def bind_assoc_pmf bind_return_pmf) lemma PSuc': "Pn' p s (Suc n) = K0' (p []) s \ (\sa. K sa \ (\s'. Pn' (\_Suc p sa) s' n))" unfolding Pn'_def by (auto intro!: bind_pmf_cong simp: Pn.simps(2) Pn_return_pmf K0_iff K0'_def bind_return_pmf map_bind_pmf bind_map_pmf) -lemma PSuc'_markovian: +lemma PSuc'_markovian: "Pn' (mk_markovian p) s (Suc n) = K_st (p 0) s \ (\s'. Pn' (mk_markovian (p \ Suc)) s' n)" unfolding PSuc' by (auto simp: bind_map_pmf bind_assoc_pmf comp_def K0'_def K_st_def intro!: bind_pmf_cong) lemma Xn'_Suc: "Xn' p s (Suc n) = Pn' p s n \ K" by (auto simp: Xn_Suc Xn'_def Pn'_def) lemma Xn'_Pn': "Xn' p s n = map_pmf fst (Pn' p s n)" by (simp add: Xn_def Xn'_def Pn'_def) lemma Suc_Xn': "Xn' p s (Suc n) = p [] s \ (\a. K (s,a) \ (\s'. Xn' (\_Suc p (s,a)) s' n))" by (auto simp: Xn'_Pn' map_bind_pmf bind_map_pmf PSuc' K0'_def) -lemma Suc_Xn'_markovian: +lemma Suc_Xn'_markovian: "Xn' (mk_markovian p) s (Suc n) = K_st (p 0) s \ (\s'. Xn' (mk_markovian (\n. p (Suc n))) s' n)" by (auto simp: K_st_def bind_assoc_pmf Suc_Xn') -lemma Xn'_split: "Xn' (mk_markovian p) s (n + m) = +lemma Xn'_split: "Xn' (mk_markovian p) s (n + m) = Xn' (mk_markovian p) s n \ (\s. Xn' (mk_markovian (\i. p (i + n))) s m)" by (induction n arbitrary: p s) (auto intro!: bind_pmf_cong simp: bind_assoc_pmf bind_return_pmf Suc_Xn') lemma Yn'_markovian: "Yn' (mk_markovian p) s n = Xn' (mk_markovian p) s n \ p n" - unfolding Yn'_def Xn'_def Yn_markovian + unfolding Yn'_def Xn'_def Yn_markovian by simp lemma Pn'_markovian_eq_Xn'_bind: "Pn' (mk_markovian p) s n = Xn' (mk_markovian p) s n \ K0' (p n)" - unfolding Xn'_def Pn'_def K0'_def K0_iff Pn_markovian_eq_Xn_bind + unfolding Xn'_def Pn'_def K0'_def K0_iff Pn_markovian_eq_Xn_bind by simp lemma Pn'_eq_\: "measure_pmf (Pn' p s n) = distr (\ p s) (count_space UNIV) (\t. t !! n)" by (auto simp: \_def Pn'_def Pn_eq_T) end -end \ No newline at end of file +end diff --git a/thys/Quasi_Borel_Spaces/Pair_QuasiBorel_Measure.thy b/thys/Quasi_Borel_Spaces/Pair_QuasiBorel_Measure.thy --- a/thys/Quasi_Borel_Spaces/Pair_QuasiBorel_Measure.thy +++ b/thys/Quasi_Borel_Spaces/Pair_QuasiBorel_Measure.thy @@ -1,559 +1,559 @@ (* Title: Pair_QuasiBorel_Measure.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) subsection \Binary Product Measure\ theory Pair_QuasiBorel_Measure imports "Monad_QuasiBorel" begin subsubsection \ Binary Product Measure\ text \ Special case of \cite{Heunen_2017} Proposition 23 where $\Omega = \mathbb{R}\times \mathbb{R}$ and $X = X \times Y$. Let $[\alpha,\mu ] \in P(X)$ and $[\beta ,\nu] \in P(Y)$. $\alpha\times\beta$ is the $\alpha$ in Proposition 23. \ definition qbs_prob_pair_measure_t :: "['a qbs_prob_t, 'b qbs_prob_t] \ ('a \ 'b) qbs_prob_t" where "qbs_prob_pair_measure_t p q \ (let (X,\,\) = p; (Y,\,\) = q in (X \\<^sub>Q Y, map_prod \ \ \ real_real.g, distr (\ \\<^sub>M \) real_borel real_real.f))" lift_definition qbs_prob_pair_measure :: "['a qbs_prob_space, 'b qbs_prob_space] \ ('a \ 'b) qbs_prob_space" (infix "\\<^sub>Q\<^sub>m\<^sub>e\<^sub>s" 80) is qbs_prob_pair_measure_t unfolding qbs_prob_pair_measure_t_def proof auto fix X X' :: "'a quasi_borel" fix Y Y' :: "'b quasi_borel" fix \ \' \ \' \ \' \ \' assume h:"qbs_prob_eq (X,\,\) (X',\',\')" "qbs_prob_eq (Y,\,\) (Y',\',\')" then have 1: "X = X'" "Y = Y'" by(auto simp: qbs_prob_eq_def) interpret pqp1: pair_qbs_probs X \ \ Y \ \ by(simp add: pair_qbs_probs_def qbs_prob_eq_dest(1)[OF h(1)] qbs_prob_eq_dest(1)[OF h(2)]) interpret pqp2: pair_qbs_probs X' \' \' Y' \' \' by(simp add: pair_qbs_probs_def qbs_prob_eq_dest(2)[OF h(1)] qbs_prob_eq_dest(2)[OF h(2)]) interpret pqp: pair_qbs_prob "X \\<^sub>Q Y" "map_prod \ \ \ real_real.g" "distr (\ \\<^sub>M \) real_borel real_real.f" "X' \\<^sub>Q Y'" "map_prod \' \' \ real_real.g" "distr (\' \\<^sub>M \') real_borel real_real.f" by(auto intro!: qbs_probI pqp1.P.prob_space_distr pqp2.P.prob_space_distr simp: pair_qbs_prob_def) show "qbs_prob_eq (X \\<^sub>Q Y, map_prod \ \ \ real_real.g, distr (\ \\<^sub>M \) real_borel real_real.f) (X' \\<^sub>Q Y', map_prod \' \' \ real_real.g, distr (\' \\<^sub>M \') real_borel real_real.f)" proof(rule pqp.qbs_prob_space_eq_inverse(1)) show "qbs_prob_space (X \\<^sub>Q Y, map_prod \ \ \ real_real.g, distr (\ \\<^sub>M \) real_borel real_real.f) = qbs_prob_space (X' \\<^sub>Q Y', map_prod \' \' \ real_real.g, distr (\' \\<^sub>M \') real_borel real_real.f)" (is "?lhs = ?rhs") proof - have "?lhs = qbs_prob_space (X, \, \) \ (\x. qbs_prob_space (Y, \, \) \ (\y. qbs_return (X \\<^sub>Q Y) (x, y)))" by(simp add: pqp1.qbs_bind_return_pq) also have "... = qbs_prob_space (X', \', \') \ (\x. qbs_prob_space (Y', \', \') \ (\y. qbs_return (X' \\<^sub>Q Y') (x, y)))" using h by(simp add: qbs_prob_space_eq 1) also have "... = ?rhs" by(simp add: pqp2.qbs_bind_return_pq) finally show ?thesis . qed qed qed lemma(in pair_qbs_probs) qbs_prob_pair_measure_computation: "(qbs_prob_space (X,\,\)) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s (qbs_prob_space (Y,\,\)) = qbs_prob_space (X \\<^sub>Q Y, map_prod \ \ \ real_real.g , distr (\ \\<^sub>M \) real_borel real_real.f)" "qbs_prob (X \\<^sub>Q Y) (map_prod \ \ \ real_real.g) (distr (\ \\<^sub>M \) real_borel real_real.f)" by(simp_all add: qbs_prob_pair_measure.abs_eq qbs_prob_pair_measure_t_def qbs_bind_return_pq) lemma qbs_prob_pair_measure_qbs: "qbs_prob_space_qbs (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) = qbs_prob_space_qbs p \\<^sub>Q qbs_prob_space_qbs q" by(transfer,simp add: qbs_prob_pair_measure_t_def Let_def prod.case_eq_if) lemma(in pair_qbs_probs) qbs_prob_pair_measure_measure: shows "qbs_prob_measure (qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\)) = distr (\ \\<^sub>M \) (qbs_to_measure (X \\<^sub>Q Y)) (map_prod \ \)" by(simp add: qbs_prob_pair_measure_computation distr_distr comp_assoc) lemma qbs_prob_pair_measure_morphism: "case_prod qbs_prob_pair_measure \ monadP_qbs X \\<^sub>Q monadP_qbs Y \\<^sub>Q monadP_qbs (X \\<^sub>Q Y)" proof(rule pair_qbs_morphismI) fix \x \y assume h: "\x \ qbs_Mx (monadP_qbs X)" " \y \ qbs_Mx (monadP_qbs Y)" then obtain \x \y gx gy where ha: "\x \ qbs_Mx X" "gx \ real_borel \\<^sub>M prob_algebra real_borel" "\x = (\r. qbs_prob_space (X, \x, gx r))" "\y \ qbs_Mx Y" "gy \ real_borel \\<^sub>M prob_algebra real_borel" "\y = (\r. qbs_prob_space (Y, \y, gy r))" using rep_monadP_qbs_MPx[of \x X] rep_monadP_qbs_MPx[of \y Y] by auto note [measurable] = ha(2,5) have "(\(x, y). x \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s y) \ (\r. (\x r, \y r)) = (\r. qbs_prob_space (X \\<^sub>Q Y, map_prod \x \y \ real_real.g, distr (gx r \\<^sub>M gy r) real_borel real_real.f))" apply standard using qbs_prob_MPx[OF ha(1,2)] qbs_prob_MPx[OF ha(4,5)] pair_qbs_probs.qbs_prob_pair_measure_computation[of X \x _ Y \y] by (auto simp: ha pair_qbs_probs_def) also have "... \ qbs_Mx (monadP_qbs (X \\<^sub>Q Y))" using qbs_prob_MPx[OF ha(1,2)] qbs_prob_MPx[OF ha(4,5)] pair_qbs_probs.ab_g_in_Mx[of X \x _ Y \y] by(auto intro!: bexI[where x="map_prod \x \y \ real_real.g"] bexI[where x="\r. distr (gx r \\<^sub>M gy r) real_borel real_real.f"] simp: monadP_qbs_MPx_def in_MPx_def pair_qbs_probs_def) finally show "(\(x, y). x \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s y) \ (\r. (\x r, \y r)) \ qbs_Mx (monadP_qbs (X \\<^sub>Q Y))" . qed lemma(in pair_qbs_probs) qbs_prob_pair_measure_nnintegral: assumes "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q z. f z \(qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\))) = (\\<^sup>+ z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ x. ((f \ map_prod \ \) \ real_real.g) x \distr (\ \\<^sub>M \) real_borel real_real.f)" by(simp add: qbs_prob_ennintegral_def[OF assms] qbs_prob_pair_measure_computation) also have "... = (\\<^sup>+ x. ((f \ map_prod \ \) \ real_real.g) (real_real.f x) \(\ \\<^sub>M \))" using assms by(intro nn_integral_distr) auto also have "... = ?rhs" by simp finally show ?thesis . qed lemma(in pair_qbs_probs) qbs_prob_pair_measure_integral: assumes "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" shows "(\\<^sub>Q z. f z \(qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\))) = (\z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" (is "?lhs = ?rhs") proof - have "?lhs = (\x. ((f \ map_prod \ \) \ real_real.g) x \distr (\ \\<^sub>M \) real_borel real_real.f)" by(simp add: qbs_prob_integral_def[OF assms] qbs_prob_pair_measure_computation) also have "... = (\ x. ((f \ map_prod \ \) \ real_real.g) (real_real.f x) \(\ \\<^sub>M \))" using assms by(intro integral_distr) auto also have "... = ?rhs" by simp finally show ?thesis . qed lemma qbs_prob_pair_measure_eq_bind: assumes "p \ monadP_qbs_Px X" and "q \ monadP_qbs_Px Y" shows "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q = p \ (\x. q \ (\y. qbs_return (X \\<^sub>Q Y) (x,y)))" proof - obtain \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_monadP_qbs_Px[OF assms(1)] by auto obtain \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_monadP_qbs_Px[OF assms(2)] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: pair_qbs_probs_def hp hq) show ?thesis by(simp add: hp(1) hq(1) pqp.qbs_prob_pair_measure_computation(1) pqp.qbs_bind_return_pq(1)) qed subsubsection \Fubini Theorem\ lemma qbs_prob_ennintegral_Fubini_fst: assumes "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" and "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. f (x,y) \q \p) = (\\<^sup>+\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - note [simp] = qbs_bind_morphism[OF qbs_morphism_const[of _ "monadP_qbs Y",simplified,OF assms(2)] curry_preserves_morphisms[OF qbs_return_morphism[of "X \\<^sub>Q Y"]],simplified curry_def,simplified] qbs_morphism_Pair1'[OF _ qbs_return_morphism[of "X \\<^sub>Q Y"]] assms(1)[simplified monadP_qbs_Px_def,simplified] assms(2)[simplified monadP_qbs_Px_def,simplified] have "?rhs = (\\<^sup>+\<^sub>Q z. f z \(p \ (\x. q \ (\y. qbs_return (X \\<^sub>Q Y) (x,y)))))" by(simp add: qbs_prob_pair_measure_eq_bind[OF assms(1,2)]) also have "... = (\\<^sup>+\<^sub>Q x. qbs_prob_ennintegral (q \ (\y. qbs_return (X \\<^sub>Q Y) (x, y))) f \p)" by(auto intro!: qbs_prob_ennintegral_bind[OF assms(1) _ assms(3)]) also have "... = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. qbs_prob_ennintegral (qbs_return (X \\<^sub>Q Y) (x, y)) f \q \p)" by(auto intro!: qbs_prob_ennintegral_cong qbs_prob_ennintegral_bind[OF assms(2) _ assms(3)]) also have "... = ?lhs" using assms(3) by(auto intro!: qbs_prob_ennintegral_cong qbs_prob_ennintegral_return) finally show ?thesis by simp qed lemma qbs_prob_ennintegral_Fubini_snd: assumes "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" and "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q y. \\<^sup>+\<^sub>Q x. f (x,y) \p \q) = (\\<^sup>+\<^sub>Q x. f x \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - note [simp] = qbs_bind_morphism[OF qbs_morphism_const[of _ "monadP_qbs X",simplified,OF assms(1)] curry_preserves_morphisms[OF qbs_morphism_pair_swap[OF qbs_return_morphism[of "X \\<^sub>Q Y"]],simplified curry_def,simplified]] qbs_morphism_Pair2'[OF _ qbs_return_morphism[of "X \\<^sub>Q Y"]] assms(1)[simplified monadP_qbs_Px_def,simplified] assms(2)[simplified monadP_qbs_Px_def,simplified] have "?rhs = (\\<^sup>+\<^sub>Q z. f z \(q \ (\y. p \ (\x. qbs_return (X \\<^sub>Q Y) (x,y)))))" by(simp add: qbs_prob_pair_measure_eq_bind[OF assms(1,2)] qbs_bind_return_rotate[OF assms(1,2)]) also have "... = (\\<^sup>+\<^sub>Q y. qbs_prob_ennintegral (p \ (\x. qbs_return (X \\<^sub>Q Y) (x, y))) f \q)" by(auto intro!: qbs_prob_ennintegral_bind[OF assms(2) _ assms(3)]) also have "... = (\\<^sup>+\<^sub>Q y. \\<^sup>+\<^sub>Q x. qbs_prob_ennintegral (qbs_return (X \\<^sub>Q Y) (x, y)) f \p \q)" by(auto intro!: qbs_prob_ennintegral_cong qbs_prob_ennintegral_bind[OF assms(1) _ assms(3)]) also have "... = ?lhs" using assms(3) by(auto intro!: qbs_prob_ennintegral_cong qbs_prob_ennintegral_return) finally show ?thesis by simp qed lemma qbs_prob_ennintegral_indep1: assumes "p \ monadP_qbs_Px X" and "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q z. f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q x. f x \p)" (is "?lhs = _") proof - obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto have "?lhs = (\\<^sup>+\<^sub>Q y. \\<^sup>+\<^sub>Q x. f x \p \q)" using qbs_prob_ennintegral_Fubini_snd[OF assms(1) qbs_prob.qbs_prob_space_in_Px[OF hq(2)] qbs_morphism_fst''[OF assms(2)]] by(simp add: hq(1)) thus ?thesis by(simp add: qbs_prob_ennintegral_const) qed lemma qbs_prob_ennintegral_indep2: assumes "q \ monadP_qbs_Px Y" and "f \ Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q z. f (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q y. f y \q)" (is "?lhs = _") proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto have "?lhs = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y. f y \q \p)" using qbs_prob_ennintegral_Fubini_fst[OF qbs_prob.qbs_prob_space_in_Px[OF hp(2)] assms(1) qbs_morphism_snd''[OF assms(2)]] by(simp add: hp(1)) thus ?thesis by(simp add: qbs_prob_ennintegral_const) qed lemma qbs_ennintegral_indep_mult: assumes "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" and "g \ Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "(\\<^sup>+\<^sub>Q z. f (fst z) * g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sup>+\<^sub>Q x. f x \p) * (\\<^sup>+\<^sub>Q y. g y \q)" (is "?lhs = ?rhs") proof - have h:"(\z. f (fst z) * g (snd z)) \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" - using assms(4,3) + using assms(4,3) by(auto intro!: borel_measurable_subalgebra[OF l_product_sets[of X Y]] simp: space_pair_measure lr_adjunction_correspondence) have "?lhs = (\\<^sup>+\<^sub>Q x. \\<^sup>+\<^sub>Q y .f x * g y \q \p)" using qbs_prob_ennintegral_Fubini_fst[OF assms(1,2) h] by simp also have "... = (\\<^sup>+\<^sub>Q x. f x * \\<^sup>+\<^sub>Q y . g y \q \p)" using qbs_prob_ennintegral_cmult[of q,OF _ assms(4)] assms(2) by(simp add: monadP_qbs_Px_def) also have "... = ?rhs" using qbs_prob_ennintegral_cmult[of p,OF _ assms(3)] assms(1) by(simp add: ab_semigroup_mult_class.mult.commute[where b="qbs_prob_ennintegral q g"] monadP_qbs_Px_def) finally show ?thesis . qed lemma(in pair_qbs_probs) qbs_prob_pair_measure_integrable: assumes "qbs_integrable (qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\)) f" shows "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" "integrable (\ \\<^sub>M \) (f \ (map_prod \ \))" proof - show "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" using qbs_integrable_morphism[OF qbs_prob_pair_measure_qbs assms] by simp next have 1:"integrable (distr (\ \\<^sub>M \) real_borel real_real.f) (f \ (map_prod \ \ \ real_real.g))" using assms[simplified qbs_prob_pair_measure_computation] qbs_integrable_def[of f] by simp have "integrable (\ \\<^sub>M \) (\x. (f \ (map_prod \ \ \ real_real.g)) (real_real.f x))" by(intro integrable_distr[OF _ 1]) simp thus "integrable (\ \\<^sub>M \) (f \ map_prod \ \)" by(simp add: comp_def) qed lemma(in pair_qbs_probs) qbs_prob_pair_measure_integrable': assumes "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" and "integrable (\ \\<^sub>M \) (f \ (map_prod \ \))" - shows "qbs_integrable (qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\)) f" + shows "qbs_integrable (qbs_prob_space (X,\,\) \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s qbs_prob_space (Y,\,\)) f" proof - have "integrable (distr (\ \\<^sub>M \) real_borel real_real.f) (f \ (map_prod \ \ \ real_real.g)) = integrable (\ \\<^sub>M \) (\x. (f \ (map_prod \ \ \ real_real.g)) (real_real.f x))" by(intro integrable_distr_eq) (use assms(1) in auto) thus ?thesis using assms qbs_integrable_def by(simp add: comp_def qbs_prob_pair_measure_computation) qed lemma qbs_integrable_pair_swap: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "qbs_integrable (q \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s p) (\(x,y). f (y,x))" proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: pair_qbs_probs_def hp hq) interpret pqp2: pair_qbs_probs Y \ \ X \ \ by(simp add: pair_qbs_probs_def hp hq) - + have "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" "integrable (\ \\<^sub>M \) (f \ map_prod \ \)" by(auto simp: pqp.qbs_prob_pair_measure_integrable[OF assms[simplified hp(1) hq(1)]]) from qbs_morphism_pair_swap[OF this(1)] pqp.integrable_product_swap[OF this(2)] have "(\(x,y). f (y,x)) \ Y \\<^sub>Q X \\<^sub>Q \\<^sub>Q" "integrable (\ \\<^sub>M \) ((\(x,y). f (y,x)) \ map_prod \ \)" by(simp_all add: map_prod_def comp_def split_beta') from pqp2.qbs_prob_pair_measure_integrable' [OF this] show ?thesis by(simp add: hp(1) hq(1)) qed lemma qbs_integrable_pair1: assumes "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" "qbs_integrable p (\x. \\<^sub>Q y. \f (x,y)\ \q)" and "\x. x \ qbs_space X \ qbs_integrable q (\y. f (x,y))" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" proof - obtain \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_monadP_qbs_Px[OF assms(1)] by auto obtain \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_monadP_qbs_Px[OF assms(2)] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: pair_qbs_probs_def hp hq) have "integrable (\ \\<^sub>M \) (f \ map_prod \ \)" proof(rule pqp.Fubini_integrable) show "f \ map_prod \ \ \ borel_measurable (\ \\<^sub>M \)" using assms(3) by auto next have "(\x. LINT y|\. norm ((f \ map_prod \ \) (x, y))) = (\x. \\<^sub>Q y. \f (x,y)\ \q) \ \" apply standard subgoal for x using qbs_morphism_Pair1'[OF qbs_Mx_to_X(2)[OF pqp.qp1.in_Mx,of x] assms(3)] by(auto intro!: pqp.qp2.qbs_prob_integral_def[symmetric] simp: hq(1)) done moreover have "integrable \ ..." using assms(4) pqp.qp1.qbs_integrable_def by (simp add: hp(1)) ultimately show "integrable \ (\x. LINT y|\. norm ((f \ map_prod \ \) (x, y)))" by simp next have "\x. integrable \ (\y. (f \ map_prod \ \) (x, y))" proof- fix x have "(\y. (f \ map_prod \ \) (x, y)) = (\y. f (\ x,y)) \ \" by auto moreover have "qbs_integrable (qbs_prob_space (Y, \, \)) (\y. f (\ x, y))" by(auto intro!: assms(5)[simplified hq(1)] simp: qbs_Mx_to_X) ultimately show "integrable \ (\y. (f \ map_prod \ \) (x, y))" by(simp add: pqp.qp2.qbs_integrable_def) qed thus "AE x in \. integrable \ (\y. (f \ map_prod \ \) (x, y))" by simp qed thus ?thesis using pqp.qbs_prob_pair_measure_integrable'[OF assms(3)] by(simp add: hp(1) hq(1)) qed lemma qbs_integrable_pair2: assumes "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" "qbs_integrable q (\y. \\<^sub>Q x. \f (x,y)\ \p)" and "\y. y \ qbs_space Y \ qbs_integrable p (\x. f (x,y))" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" using qbs_integrable_pair_swap[OF qbs_integrable_pair1[OF assms(2,1) qbs_morphism_pair_swap[OF assms(3)],simplified,OF assms(4,5)]] by simp lemma qbs_integrable_fst: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "qbs_integrable p (\x. \\<^sub>Q y. f (x,y) \q)" proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: hp hq pair_qbs_probs_def) have h0: "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" using qbs_integrable_morphism[OF _ assms,simplified qbs_prob_pair_measure_qbs] by(simp_all add: monadP_qbs_Px_def hp(1) hq(1)) show "qbs_integrable p (\x. \\<^sub>Q y. f (x, y) \q)" proof(auto simp add: pqp.qp1.qbs_integrable_def hp(1)) show "(\x. \\<^sub>Q y. f (x, y) \q) \ borel_measurable (qbs_to_measure X)" using qbs_morphism_integral_fst[OF h0(2,3)] by auto next have "integrable \ (\x. LINT y|\. (f \ map_prod \ \) (x, y))" by(intro pqp.integrable_fst') (rule pqp.qbs_prob_pair_measure_integrable(2)[OF assms[simplified hp(1) hq(1)]]) moreover have "\x. ((\x. \\<^sub>Q y. f (x, y) \q) \ \) x = LINT y|\. (f \ map_prod \ \) (x, y)" by(auto intro!: pqp.qp2.qbs_prob_integral_def qbs_morphism_Pair1'[OF qbs_Mx_to_X(2)[OF pqp.qp1.in_Mx] h0(3)] simp: hq) ultimately show "integrable \ ((\x. \\<^sub>Q y. f (x, y) \q) \ \)" - using integrable_cong[of \ \ "(\x. \\<^sub>Q y. f (x, y) \q) \ \" " (\x. LINT y|\. (f \ map_prod \ \) (x, y))"] + using Bochner_Integration.integrable_cong[of \ \ "(\x. \\<^sub>Q y. f (x, y) \q) \ \" " (\x. LINT y|\. (f \ map_prod \ \) (x, y))"] by simp qed qed lemma qbs_integrable_snd: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "qbs_integrable q (\y. \\<^sub>Q x. f (x,y) \p)" using qbs_integrable_fst[OF qbs_integrable_pair_swap[OF assms]] by simp lemma qbs_integrable_indep_mult: assumes "qbs_integrable p f" and "qbs_integrable q g" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. f (fst x) * g (snd x))" proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: hp hq pair_qbs_probs_def) have h0: "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q \\<^sub>Q" "g \ Y \\<^sub>Q \\<^sub>Q" using qbs_integrable_morphism[OF _ assms(1)] qbs_integrable_morphism[OF _ assms(2)] by(simp_all add: monadP_qbs_Px_def hp(1) hq(1)) show ?thesis proof(rule qbs_integrable_pair1[OF h0(1,2)],simp_all add: assms(2)) show "(\z. f (fst z) * g (snd z)) \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" using h0(3,4) by(auto intro!: borel_measurable_subalgebra[OF l_product_sets[of X Y]] simp: space_pair_measure lr_adjunction_correspondence) next show "qbs_integrable p (\x. \\<^sub>Q y. \f x * g y\ \q)" by(auto intro!: qbs_integrable_mult[OF qbs_integrable_abs[OF assms(1)]] simp only: idom_abs_sgn_class.abs_mult qbs_prob_integral_cmult ab_semigroup_mult_class.mult.commute[where b="\\<^sub>Q y. \g y\ \q"]) qed qed lemma qbs_integrable_indep1: assumes "qbs_integrable p f" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. f (fst x))" using qbs_integrable_indep_mult[OF assms qbs_integrable_const[of q 1]] by simp lemma qbs_integrable_indep2: assumes "qbs_integrable q g" shows "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\x. g (snd x))" using qbs_integrable_pair_swap[OF qbs_integrable_indep1[OF assms],of p] by(simp add: split_beta') lemma qbs_prob_integral_Fubini_fst: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "(\\<^sub>Q x. \\<^sub>Q y. f (x,y) \q \p) = (\\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: hp hq pair_qbs_probs_def) have h0: "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" using qbs_integrable_morphism[OF _ assms,simplified qbs_prob_pair_measure_qbs] by(simp_all add: monadP_qbs_Px_def hp(1) hq(1)) have "?lhs = (\ x. \\<^sub>Q y. f (\ x, y) \q \\)" using qbs_morphism_integral_fst[OF h0(2) h0(3)] by(auto intro!: pqp.qp1.qbs_prob_integral_def simp: hp(1)) also have "... = (\x. \y. f (\ x, \ y) \\ \\)" using qbs_morphism_Pair1'[OF qbs_Mx_to_X(2)[OF pqp.qp1.in_Mx] h0(3)] by(auto intro!: Bochner_Integration.integral_cong pqp.qp2.qbs_prob_integral_def simp: hq(1)) also have "... = (\z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" using pqp.integral_fst'[OF pqp.qbs_prob_pair_measure_integrable(2)[OF assms[simplified hp(1) hq(1)]]] by(simp add: map_prod_def comp_def) also have "... = ?rhs" by(simp add: pqp.qbs_prob_pair_measure_integral[OF h0(3)] hp(1) hq(1)) finally show ?thesis . qed lemma qbs_prob_integral_Fubini_snd: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" shows "(\\<^sub>Q y. \\<^sub>Q x. f (x,y) \p \q) = (\\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - obtain X \ \ where hp: "p = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of p] by auto obtain Y \ \ where hq: - "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" + "q = qbs_prob_space (Y, \, \)" "qbs_prob Y \ \" using rep_qbs_prob_space[of q] by auto interpret pqp: pair_qbs_probs X \ \ Y \ \ by(simp add: hp hq pair_qbs_probs_def) have h0: "p \ monadP_qbs_Px X" "q \ monadP_qbs_Px Y" "f \ X \\<^sub>Q Y \\<^sub>Q \\<^sub>Q" using qbs_integrable_morphism[OF _ assms,simplified qbs_prob_pair_measure_qbs] by(simp_all add: monadP_qbs_Px_def hp(1) hq(1)) have "?lhs = (\ y. \\<^sub>Q x. f (x,\ y) \p \\)" using qbs_morphism_integral_snd[OF h0(1) h0(3)] by(auto intro!: pqp.qp2.qbs_prob_integral_def simp: hq(1)) also have "... = (\y. \x. f (\ x, \ y) \\ \\)" using qbs_morphism_Pair2'[OF qbs_Mx_to_X(2)[OF pqp.qp2.in_Mx] h0(3)] by(auto intro!: Bochner_Integration.integral_cong pqp.qp1.qbs_prob_integral_def simp: hp(1)) also have "... = (\z. (f \ map_prod \ \) z \(\ \\<^sub>M \))" using pqp.integral_snd[of "curry (f \ map_prod \ \)"] pqp.qbs_prob_pair_measure_integrable(2)[OF assms[simplified hp(1) hq(1)]] by(simp add: map_prod_def comp_def split_beta') also have "... = ?rhs" by(simp add: pqp.qbs_prob_pair_measure_integral[OF h0(3)] hp(1) hq(1)) finally show ?thesis . qed lemma qbs_prob_integral_indep1: assumes "qbs_integrable p f" shows "(\\<^sub>Q z. f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q x. f x \p)" using qbs_prob_integral_Fubini_snd[OF qbs_integrable_indep1[OF assms],of q] by(simp add: qbs_prob_integral_const) lemma qbs_prob_integral_indep2: assumes "qbs_integrable q g" shows "(\\<^sub>Q z. g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q y. g y \q)" using qbs_prob_integral_Fubini_fst[OF qbs_integrable_indep2[OF assms],of p] by(simp add: qbs_prob_integral_const) lemma qbs_prob_integral_indep_mult: assumes "qbs_integrable p f" and "qbs_integrable q g" shows "(\\<^sub>Q z. f (fst z) * g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q x. f x \p) * (\\<^sub>Q y. g y \q)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sub>Q x. \\<^sub>Q y. f x * g y \q \p)" using qbs_prob_integral_Fubini_fst[OF qbs_integrable_indep_mult[OF assms]] by simp also have "... = (\\<^sub>Q x. f x * (\\<^sub>Q y. g y \q) \p)" by(simp add: qbs_prob_integral_cmult) also have "... = ?rhs" by(simp add: qbs_prob_integral_cmult ab_semigroup_mult_class.mult.commute[where b="\\<^sub>Q y. g y \q"]) finally show ?thesis . qed lemma qbs_prob_var_indep_plus: assumes "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f" "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\z. (f z)\<^sup>2)" "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g" "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\z. (g z)\<^sup>2)" "qbs_integrable (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\z. (f z) * (g z))" and "(\\<^sub>Q z. f z * g z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q z. f z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) * (\\<^sub>Q z. g z \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" shows "qbs_prob_var (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\z. f z + g z) = qbs_prob_var (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f + qbs_prob_var (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g" unfolding qbs_prob_var_def proof - show "(\\<^sub>Q z. (f z + g z - \\<^sub>Q w. f w + g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))\<^sup>2 \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) = (\\<^sub>Q z. (f z - qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f)\<^sup>2 \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) + (\\<^sub>Q z. (g z - qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g)\<^sup>2 \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sub>Q z. ((f z - (\\<^sub>Q w. f w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))) + (g z - (\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))))\<^sup>2 \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" by(simp add: qbs_prob_integral_add[OF assms(1,3)] add_diff_add) also have "... = (\\<^sub>Q z. (f z - (\\<^sub>Q w. f w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)))\<^sup>2 + (g z - (\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)))\<^sup>2 + (2 * f z * g z - 2 * (\\<^sub>Q w. f w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) * g z - (2 * f z * (\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) - 2 * (\\<^sub>Q w. f w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)) * (\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)))) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" by(simp add: comm_semiring_1_class.power2_sum comm_semiring_1_cancel_class.left_diff_distrib' ring_class.right_diff_distrib) also have "... = ?rhs" using qbs_prob_integral_add[OF qbs_integrable_add[OF qbs_integrable_sq[OF assms(1,2)] qbs_integrable_sq[OF assms(3,4)]] qbs_integrable_diff[OF qbs_integrable_diff[OF qbs_integrable_mult[OF assms(5),of 2,simplified comm_semiring_1_class.semiring_normalization_rules(18)] qbs_integrable_mult[OF assms(3),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f"]] qbs_integrable_diff[OF qbs_integrable_mult[OF assms(1),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g",simplified ab_semigroup_mult_class.mult_ac(1)[where b="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] ab_semigroup_mult_class.mult.commute[where a="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] comm_semiring_1_class.semiring_normalization_rules(18)[of _ _ "qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]] qbs_integrable_const[of _ "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]]]] qbs_prob_integral_add[OF qbs_integrable_sq[OF assms(1,2)] qbs_integrable_sq[OF assms(3,4)]] qbs_prob_integral_diff[OF qbs_integrable_diff[OF qbs_integrable_mult[OF assms(5),of 2,simplified comm_semiring_1_class.semiring_normalization_rules(18)] qbs_integrable_mult[OF assms(3),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f"]] qbs_integrable_diff[OF qbs_integrable_mult[OF assms(1),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g",simplified ab_semigroup_mult_class.mult_ac(1)[where b="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] ab_semigroup_mult_class.mult.commute[where a="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] comm_semiring_1_class.semiring_normalization_rules(18)[of _ _ "qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]] qbs_integrable_const[of _ "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]]] qbs_prob_integral_diff[OF qbs_integrable_mult[OF assms(5),of 2,simplified comm_semiring_1_class.semiring_normalization_rules(18)] qbs_integrable_mult[OF assms(3),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f"]] qbs_prob_integral_diff[OF qbs_integrable_mult[OF assms(1),of "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g",simplified ab_semigroup_mult_class.mult_ac(1)[where b="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] ab_semigroup_mult_class.mult.commute[where a="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] comm_semiring_1_class.semiring_normalization_rules(18)[of _ _ "qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]] qbs_integrable_const[of _ "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"]] qbs_prob_integral_cmult[of "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q" 2 "\z. f z * g z",simplified assms(6) comm_semiring_1_class.semiring_normalization_rules(18)] qbs_prob_integral_cmult[of "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q" "2 * (\\<^sub>Q w. f w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" g] qbs_prob_integral_cmult[of "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q" "2 * (\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q))" f,simplified semigroup_mult_class.mult.assoc[of 2 "\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)"] ab_semigroup_mult_class.mult.commute[where a="qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] comm_semiring_1_class.semiring_normalization_rules(18)[of 2 _ "\\<^sub>Q w. g w \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)"]] qbs_prob_integral_const[of "p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q" "2 * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) f * qbs_prob_integral (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) g"] by simp finally show ?thesis . qed qed lemma qbs_prob_var_indep_plus': assumes "qbs_integrable p f" "qbs_integrable p (\x. (f x)\<^sup>2)" "qbs_integrable q g" and "qbs_integrable q (\x. (g x)\<^sup>2)" shows "qbs_prob_var (p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q) (\z. f (fst z) + g (snd z)) = qbs_prob_var p f + qbs_prob_var q g" using qbs_prob_var_indep_plus[OF qbs_integrable_indep1[OF assms(1)] qbs_integrable_indep1[OF assms(2)] qbs_integrable_indep2[OF assms(3)] qbs_integrable_indep2[OF assms(4)] qbs_integrable_indep_mult[OF assms(1) assms(3)] qbs_prob_integral_indep_mult[OF assms(1) assms(3),simplified qbs_prob_integral_indep1[OF assms(1),of q,symmetric] qbs_prob_integral_indep2[OF assms(3),of p,symmetric]]] qbs_prob_integral_indep1[OF qbs_integrable_sq[OF assms(1,2)],of q "\\<^sub>Q z. f (fst z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)"] qbs_prob_integral_indep2[OF qbs_integrable_sq[OF assms(3,4)],of p "\\<^sub>Q z. g (snd z) \(p \\<^sub>Q\<^sub>m\<^sub>e\<^sub>s q)"] by(simp add: qbs_prob_var_def qbs_prob_integral_indep1[OF assms(1)] qbs_prob_integral_indep2[OF assms(3)]) -end \ No newline at end of file +end diff --git a/thys/Quasi_Borel_Spaces/Probability_Space_QuasiBorel.thy b/thys/Quasi_Borel_Spaces/Probability_Space_QuasiBorel.thy --- a/thys/Quasi_Borel_Spaces/Probability_Space_QuasiBorel.thy +++ b/thys/Quasi_Borel_Spaces/Probability_Space_QuasiBorel.thy @@ -1,1125 +1,1122 @@ (* Title: Probability_Space_QuasiBorel.thy Author: Michikazu Hirata, Yasuhiko Minamide, Tokyo Institute of Technology *) section \Probability Spaces\ subsection \Probability Measures \ theory Probability_Space_QuasiBorel imports Exponent_QuasiBorel begin subsubsection \ Probability Measures \ type_synonym 'a qbs_prob_t = "'a quasi_borel * (real \ 'a) * real measure" locale in_Mx = fixes X :: "'a quasi_borel" and \ :: "real \ 'a" assumes in_Mx[simp]:"\ \ qbs_Mx X" -locale qbs_prob = in_Mx X \ + real_distribution \ +locale qbs_prob = in_Mx X \ + real_distribution \ for X :: "'a quasi_borel" and \ and \ begin declare prob_space_axioms[simp] lemma m_in_space_prob_algebra[simp]: "\ \ space (prob_algebra real_borel)" using space_prob_algebra[of real_borel] by simp end locale pair_qbs_probs = qp1:qbs_prob X \ \ + qp2:qbs_prob Y \ \ for X :: "'a quasi_borel"and \ \ and Y :: "'b quasi_borel" and \ \ begin sublocale pair_prob_space \ \ by standard lemma ab_measurable[measurable]: "map_prod \ \ \ real_borel \\<^sub>M real_borel \\<^sub>M qbs_to_measure (X \\<^sub>Q Y)" using qbs_morphism_map_prod[of \ "\\<^sub>Q" X \ "\\<^sub>Q" Y] qp1.in_Mx qp2.in_Mx l_preserves_morphisms[of "\\<^sub>Q \\<^sub>Q \\<^sub>Q" "X \\<^sub>Q Y"] by(auto simp: qbs_Mx_is_morphisms) lemma ab_g_in_Mx[simp]: "map_prod \ \ \ real_real.g \ pair_qbs_Mx X Y" using qbs_closed1_dest[OF qp1.in_Mx] qbs_closed1_dest[OF qp2.in_Mx] by(auto simp add: pair_qbs_Mx_def comp_def) sublocale qbs_prob "X \\<^sub>Q Y" "map_prod \ \ \ real_real.g" "distr (\ \\<^sub>M \) real_borel real_real.f" by(auto simp: qbs_prob_def in_Mx_def) end locale pair_qbs_prob = qp1:qbs_prob X \ \ + qp2:qbs_prob Y \ \ for X :: "'a quasi_borel"and \ \ and Y :: "'a quasi_borel" and \ \ begin sublocale pair_qbs_probs by standard lemma same_spaces[simp]: assumes "Y = X" shows "\ \ qbs_Mx X" by(simp add: assms[symmetric]) end lemma prob_algebra_real_prob_measure: "p \ space (prob_algebra (real_borel)) = real_distribution p" proof assume "p \ space (prob_algebra real_borel)" then show "real_distribution p" unfolding real_distribution_def real_distribution_axioms_def by(simp add: space_prob_algebra sets_eq_imp_space_eq) next assume "real_distribution p" then interpret rd: real_distribution p . show "p \ space (prob_algebra real_borel)" by (simp add: space_prob_algebra rd.prob_space_axioms) qed lemma qbs_probI: assumes "\ \ qbs_Mx X" and "sets \ = sets borel" and "prob_space \" shows "qbs_prob X \ \" using assms by(auto intro!: qbs_prob.intro simp: in_Mx_def real_distribution_def real_distribution_axioms_def) lemma qbs_empty_not_qbs_prob :"\ qbs_prob (empty_quasi_borel) f M" by(simp add: qbs_prob_def in_Mx_def) definition qbs_prob_eq :: "['a qbs_prob_t, 'a qbs_prob_t] \ bool" where "qbs_prob_eq p1 p2 \ (let (qbs1, a1, m1) = p1; (qbs2, a2, m2) = p2 in qbs_prob qbs1 a1 m1 \ qbs_prob qbs2 a2 m2 \ qbs1 = qbs2 \ distr m1 (qbs_to_measure qbs1) a1 = distr m2 (qbs_to_measure qbs2) a2)" definition qbs_prob_eq2 :: "['a qbs_prob_t, 'a qbs_prob_t] \ bool" where "qbs_prob_eq2 p1 p2 \ (let (qbs1, a1, m1) = p1; (qbs2, a2, m2) = p2 in qbs_prob qbs1 a1 m1 \ qbs_prob qbs2 a2 m2 \ qbs1 = qbs2 \ - (\f \ qbs1 \\<^sub>Q real_quasi_borel. + (\f \ qbs1 \\<^sub>Q real_quasi_borel. (\x. f (a1 x) \ m1) = (\x. f (a2 x) \ m2)))" -definition qbs_prob_eq3 :: "['a qbs_prob_t, 'a qbs_prob_t] \ bool" where - "qbs_prob_eq3 p1 p2 \ +definition qbs_prob_eq3 :: "['a qbs_prob_t, 'a qbs_prob_t] \ bool" where + "qbs_prob_eq3 p1 p2 \ (let (qbs1, a1, m1) = p1; (qbs2, a2, m2) = p2 in (qbs_prob qbs1 a1 m1 \ qbs_prob qbs2 a2 m2 \ qbs1 = qbs2 \ (\f \ qbs1 \\<^sub>Q real_quasi_borel. (\ k \ qbs_space qbs1. 0 \ f k) \ (\x. f (a1 x) \ m1) = (\x. f (a2 x) \ m2))))" definition qbs_prob_eq4 :: "['a qbs_prob_t, 'a qbs_prob_t] \ bool" where "qbs_prob_eq4 p1 p2 \ (let (qbs1, a1, m1) = p1; (qbs2, a2, m2) = p2 in (qbs_prob qbs1 a1 m1 \ qbs_prob qbs2 a2 m2 \ qbs1 = qbs2 \ - (\f \ qbs1 \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0. + (\f \ qbs1 \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0. (\\<^sup>+x. f (a1 x) \ m1) = (\\<^sup>+x. f (a2 x) \ m2))))" lemma(in qbs_prob) qbs_prob_eq_refl[simp]: "qbs_prob_eq (X,\,\) (X,\,\)" by(simp add: qbs_prob_eq_def qbs_prob_axioms) lemma(in qbs_prob) qbs_prob_eq2_refl[simp]: "qbs_prob_eq2 (X,\,\) (X,\,\)" by(simp add: qbs_prob_eq2_def qbs_prob_axioms) lemma(in qbs_prob) qbs_prob_eq3_refl[simp]: "qbs_prob_eq3 (X,\,\) (X,\,\)" by(simp add: qbs_prob_eq3_def qbs_prob_axioms) lemma(in qbs_prob) qbs_prob_eq4_refl[simp]: "qbs_prob_eq4 (X,\,\) (X,\,\)" by(simp add: qbs_prob_eq4_def qbs_prob_axioms) lemma(in pair_qbs_prob) qbs_prob_eq_intro: assumes "X = Y" and "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" shows "qbs_prob_eq (X,\,\) (Y,\,\)" - using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms + using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms by(auto simp add: qbs_prob_eq_def) lemma(in pair_qbs_prob) qbs_prob_eq2_intro: assumes "X = Y" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" shows "qbs_prob_eq2 (X,\,\) (Y,\,\)" - using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms + using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms by(auto simp add: qbs_prob_eq2_def) lemma(in pair_qbs_prob) qbs_prob_eq3_intro: assumes "X = Y" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\ k \ qbs_space X. 0 \ f k) \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" shows "qbs_prob_eq3 (X,\,\) (Y,\,\)" - using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms + using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms by(auto simp add: qbs_prob_eq3_def) lemma(in pair_qbs_prob) qbs_prob_eq4_intro: assumes "X = Y" and "\f. f \ qbs_to_measure X \\<^sub>M ennreal_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)" shows "qbs_prob_eq4 (X,\,\) (Y,\,\)" - using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms + using assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms by(auto simp add: qbs_prob_eq4_def) lemma qbs_prob_eq_dest: assumes "qbs_prob_eq (X,\,\) (Y,\,\)" shows "qbs_prob X \ \" "qbs_prob Y \ \" "Y = X" - and "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" + and "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" using assms by(auto simp: qbs_prob_eq_def) lemma qbs_prob_eq2_dest: assumes "qbs_prob_eq2 (X,\,\) (Y,\,\)" shows "qbs_prob X \ \" "qbs_prob Y \ \" "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" using assms by(auto simp: qbs_prob_eq2_def) lemma qbs_prob_eq3_dest: assumes "qbs_prob_eq3 (X,\,\) (Y,\,\)" shows "qbs_prob X \ \" "qbs_prob Y \ \" "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\ k \ qbs_space X. 0 \ f k) \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" using assms by(auto simp: qbs_prob_eq3_def) lemma qbs_prob_eq4_dest: assumes "qbs_prob_eq4 (X,\,\) (Y,\,\)" shows "qbs_prob X \ \" "qbs_prob Y \ \" "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M ennreal_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)" using assms by(auto simp: qbs_prob_eq4_def) definition qbs_prob_t_ennintegral :: "['a qbs_prob_t, 'a \ ennreal] \ ennreal" where "qbs_prob_t_ennintegral p f \ - (if f \ (fst p) \\<^sub>Q ennreal_quasi_borel + (if f \ (fst p) \\<^sub>Q ennreal_quasi_borel then (\\<^sup>+x. f (fst (snd p) x) \ (snd (snd p))) else 0)" definition qbs_prob_t_integral :: "['a qbs_prob_t, 'a \ real] \ real" where -"qbs_prob_t_integral p f \ +"qbs_prob_t_integral p f \ (if f \ (fst p) \\<^sub>Q \\<^sub>Q then (\x. f (fst (snd p) x) \ (snd (snd p))) else 0)" definition qbs_prob_t_integrable :: "['a qbs_prob_t, 'a \ real] \ bool" where "qbs_prob_t_integrable p f \ f \ fst p \\<^sub>Q real_quasi_borel \ integrable (snd (snd p)) (f \ (fst (snd p)))" definition qbs_prob_t_measure :: "'a qbs_prob_t \ 'a measure" where "qbs_prob_t_measure p \ distr (snd (snd p)) (qbs_to_measure (fst p)) (fst (snd p))" lemma qbs_prob_eq_symp: "symp qbs_prob_eq" by(simp add: symp_def qbs_prob_eq_def) lemma qbs_prob_eq_transp: "transp qbs_prob_eq" by(simp add: transp_def qbs_prob_eq_def) quotient_type 'a qbs_prob_space = "'a qbs_prob_t" / partial: qbs_prob_eq morphisms rep_qbs_prob_space qbs_prob_space proof(rule part_equivpI) let ?U = "UNIV :: 'a set" let ?Uf = "UNIV :: (real \ 'a) set" let ?f = "(\_. undefined) :: real \ 'a" have "qbs_prob (Abs_quasi_borel (?U,?Uf)) ?f (return borel 0)" proof(auto simp add: qbs_prob_def in_Mx_def) have "Rep_quasi_borel (Abs_quasi_borel (?U,?Uf)) = (?U, ?Uf)" using Abs_quasi_borel_inverse by (auto simp add: qbs_closed1_def qbs_closed2_def qbs_closed3_def is_quasi_borel_def) thus "(\_. undefined) \ qbs_Mx (Abs_quasi_borel (?U, ?Uf))" by(simp add: qbs_Mx_def) next show "real_distribution (return borel 0)" by (simp add: prob_space_return real_distribution_axioms_def real_distribution_def) qed thus "\x :: 'a qbs_prob_t . qbs_prob_eq x x" unfolding qbs_prob_eq_def by(auto intro!: exI[where x="(Abs_quasi_borel (?U,?Uf), ?f, return borel 0)"]) qed (simp_all add: qbs_prob_eq_symp qbs_prob_eq_transp) interpretation qbs_prob_space : quot_type "qbs_prob_eq" "Abs_qbs_prob_space" "Rep_qbs_prob_space" using Abs_qbs_prob_space_inverse Rep_qbs_prob_space by(simp add: quot_type_def equivp_implies_part_equivp qbs_prob_space_equivp Rep_qbs_prob_space_inverse Rep_qbs_prob_space_inject) blast lemma qbs_prob_space_induct: assumes "\X \ \. qbs_prob X \ \ \ P (qbs_prob_space (X,\,\))" shows "P s" apply(rule qbs_prob_space.abs_induct) using assms by(auto simp: qbs_prob_eq_def) lemma qbs_prob_space_induct': assumes "\X \ \. qbs_prob X \ \ \ s = qbs_prob_space (X,\,\)\ P (qbs_prob_space (X,\,\))" shows "P s" by (metis (no_types, lifting) Rep_qbs_prob_space_inverse assms case_prodE qbs_prob_eq_def qbs_prob_space.abs_def qbs_prob_space.rep_prop qbs_prob_space_def) lemma rep_qbs_prob_space: "\X \ \. p = qbs_prob_space (X, \, \) \ qbs_prob X \ \" by(rule qbs_prob_space.abs_induct,auto simp add: qbs_prob_eq_def) lemma(in qbs_prob) in_Rep: "(X, \, \) \ Rep_qbs_prob_space (qbs_prob_space (X,\,\))" by (metis mem_Collect_eq qbs_prob_eq_refl qbs_prob_space.abs_def qbs_prob_space.abs_inverse qbs_prob_space_def) lemma(in qbs_prob) if_in_Rep: assumes "(X',\',\') \ Rep_qbs_prob_space (qbs_prob_space (X,\,\))" shows "X' = X" "qbs_prob X' \' \'" "qbs_prob_eq (X,\,\) (X',\',\')" proof - have h:"X' = X" by (metis assms mem_Collect_eq qbs_prob_eq_dest(3) qbs_prob_eq_refl qbs_prob_space.abs_def qbs_prob_space.abs_inverse qbs_prob_space_def) have [simp]:"qbs_prob X' \' \'" by (metis assms mem_Collect_eq prod_cases3 qbs_prob_eq_dest(2) qbs_prob_space.rep_prop) have [simp]:"qbs_prob_eq (X,\,\) (X',\',\')" by (metis assms mem_Collect_eq qbs_prob_eq_refl qbs_prob_space.abs_def qbs_prob_space.abs_inverse qbs_prob_space_def) show "X' = X" "qbs_prob X' \' \'" "qbs_prob_eq (X,\,\) (X',\',\')" by simp_all (simp add: h) qed lemma(in qbs_prob) in_Rep_induct: assumes "\Y \ \. (Y,\,\) \ Rep_qbs_prob_space (qbs_prob_space (X,\,\)) \ P (Y,\,\)" shows "P (rep_qbs_prob_space (qbs_prob_space (X,\,\)))" unfolding rep_qbs_prob_space_def qbs_prob_space.rep_def by(rule someI2[where a="(X,\,\)"]) (use in_Rep assms in auto) (* qbs_prob_eq[1-4] are equivalent. *) lemma qbs_prob_eq_2_implies_3 : assumes "qbs_prob_eq2 p1 p2" shows "qbs_prob_eq3 p1 p2" using assms by(auto simp: qbs_prob_eq2_def qbs_prob_eq3_def) lemma qbs_prob_eq_3_implies_1 : assumes "qbs_prob_eq3 (p1 :: 'a qbs_prob_t) p2" shows "qbs_prob_eq p1 p2" proof(rule prod_cases3[where y=p1],rule prod_cases3[where y=p2],simp) fix X Y :: "'a quasi_borel" fix \ \ \ \ assume "p1 = (X,\,\)" "p2 = (Y,\,\)" then have h:"qbs_prob_eq3 (X,\,\) (Y,\,\)" using assms by simp then interpret qp : pair_qbs_prob X \ \ Y \ \ by(auto intro!: pair_qbs_prob.intro simp: qbs_prob_eq3_def) note [simp] = qbs_prob_eq3_dest(3)[OF h] show "qbs_prob_eq (X,\,\) (Y,\,\)" proof(rule qp.qbs_prob_eq_intro) show "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" proof(rule measure_eqI) fix U assume hu:"U \ sets (distr \ (qbs_to_measure X) \)" have "measure (distr \ (qbs_to_measure X) \) U = measure (distr \ (qbs_to_measure X) \) U" (is "?lhs = ?rhs") proof - have "?lhs = measure \ (\ -` U \ space \)" by(rule measure_distr) (use hu in simp_all) also have "... = integral\<^sup>L \ (indicat_real (\ -` U))" by simp also have "... = (\x. indicat_real U (\ x) \\)" using indicator_vimage[of \ U] Bochner_Integration.integral_cong[of \ _ "indicat_real (\ -` U)" "\x. indicat_real U (\ x)"] by auto also have "... = (\x. indicat_real U (\ x) \\)" using qbs_prob_eq3_dest(4)[OF h,of "indicat_real U"] hu by simp also have "... = integral\<^sup>L \ (indicat_real (\ -` U))" using indicator_vimage[of \ U,symmetric] Bochner_Integration.integral_cong[of \ _ "\x. indicat_real U (\ x)" "indicat_real (\ -` U)"] by blast also have "... = measure \ (\ -` U \ space \)" by simp also have "... = ?rhs" by(rule measure_distr[symmetric]) (use hu in simp_all) finally show ?thesis . qed thus "emeasure (distr \ (qbs_to_measure X) \) U = emeasure (distr \ (qbs_to_measure X) \) U" using qp.qp2.finite_measure_distr[of \] qp.qp1.finite_measure_distr[of \] by(simp add: finite_measure.emeasure_eq_measure) qed simp qed simp qed lemma qbs_prob_eq_1_implies_2 : assumes "qbs_prob_eq p1 (p2 :: 'a qbs_prob_t)" shows "qbs_prob_eq2 p1 p2" proof(rule prod_cases3[where y=p1],rule prod_cases3[where y=p2],simp) fix X Y :: "'a quasi_borel" fix \ \ \ \ assume "p1 = (X,\,\)" "p2 = (Y,\,\)" then have h:"qbs_prob_eq (X,\,\) (Y,\,\)" using assms by simp then interpret qp : pair_qbs_prob X \ \ Y \ \ by(auto intro!: pair_qbs_prob.intro simp: qbs_prob_eq_def) note [simp] = qbs_prob_eq_dest(3)[OF h] show "qbs_prob_eq2 (X,\,\) (Y,\,\)" proof(rule qp.qbs_prob_eq2_intro) fix f :: "'a \ real" assume [measurable]:"f \ borel_measurable (qbs_to_measure X)" show "(\r. f (\ r) \ \) = (\r. f (\ r) \ \)" (is "?lhs = ?rhs") proof - have "?lhs = (\x. f x \(distr \ (qbs_to_measure X) \))" by(simp add: Bochner_Integration.integral_distr[symmetric]) also have "... = (\x. f x \(distr \ (qbs_to_measure X) \))" by(simp add: qbs_prob_eq_dest(4)[OF h]) also have "... = ?rhs" by(simp add: Bochner_Integration.integral_distr) finally show ?thesis . qed qed simp qed - + lemma qbs_prob_eq_1_implies_4 : assumes "qbs_prob_eq p1 p2" shows "qbs_prob_eq4 p1 p2" proof(rule prod_cases3[where y=p1],rule prod_cases3[where y=p2],simp) fix X Y :: "'a quasi_borel" fix \ \ \ \ assume "p1 = (X,\,\)" "p2 = (Y,\,\)" then have h:"qbs_prob_eq (X,\,\) (Y,\,\)" using assms by simp then interpret qp : pair_qbs_prob X \ \ Y \ \ by(auto intro!: pair_qbs_prob.intro simp: qbs_prob_eq_def) note [simp] = qbs_prob_eq_dest(3)[OF h] show "qbs_prob_eq4 (X,\,\) (Y,\,\)" proof(rule qp.qbs_prob_eq4_intro) fix f ::"'a \ ennreal" assume [measurable]:"f \ borel_measurable (qbs_to_measure X)" show "(\\<^sup>+ x. f (\ x) \\) = (\\<^sup>+ x. f (\ x) \\)" (is "?lhs = ?rhs") proof - have "?lhs = integral\<^sup>N (distr \ (qbs_to_measure X) \) f" by(simp add: nn_integral_distr) also have "... = integral\<^sup>N (distr \ (qbs_to_measure X) \) f" by(simp add: qbs_prob_eq_dest(4)[OF h]) also have "... = ?rhs" by(simp add: nn_integral_distr) finally show ?thesis . - qed + qed qed simp qed lemma qbs_prob_eq_4_implies_3 : assumes "qbs_prob_eq4 p1 p2" shows "qbs_prob_eq3 p1 p2" proof(rule prod_cases3[where y=p1],rule prod_cases3[where y=p2],simp) fix X Y :: "'a quasi_borel" fix \ \ \ \ assume "p1 = (X,\,\)" "p2 = (Y,\,\)" then have h:"qbs_prob_eq4 (X,\,\) (Y,\,\)" using assms by simp then interpret qp : pair_qbs_prob X \ \ Y \ \ by(auto intro!: pair_qbs_prob.intro simp: qbs_prob_eq4_def) note [simp] = qbs_prob_eq4_dest(3)[OF h] show "qbs_prob_eq3 (X,\,\) (Y,\,\)" proof(rule qp.qbs_prob_eq3_intro) fix f :: "'a \ real" assume [measurable]:"f \ borel_measurable (qbs_to_measure X)" and h': "\k\qbs_space X. 0 \ f k" show "(\ x. f (\ x) \\) = (\ x. f (\ x) \\)" (is "?lhs = ?rhs") proof - have "?lhs = enn2real (\\<^sup>+ x. ennreal (f (\ x)) \\)" using h' by(auto simp: integral_eq_nn_integral[where f="(\x. f (\ x))"] qbs_Mx_to_X(2)) also have "... = enn2real (\\<^sup>+ x. (ennreal \ f) (\ x) \\)" by simp also have "... = enn2real (\\<^sup>+ x. (ennreal \ f) (\ x) \\)" using qbs_prob_eq4_dest(4)[OF h,of "ennreal \ f"] by simp also have "... = enn2real (\\<^sup>+ x. ennreal (f (\ x)) \\)" by simp also have "... = ?rhs" using h' by(auto simp: integral_eq_nn_integral[where f="(\x. f (\ x))"] qbs_Mx_to_X(2)) finally show ?thesis . qed qed simp qed lemma qbs_prob_eq_equiv12 : "qbs_prob_eq = qbs_prob_eq2" using qbs_prob_eq_1_implies_2 qbs_prob_eq_2_implies_3 qbs_prob_eq_3_implies_1 by blast lemma qbs_prob_eq_equiv13 : "qbs_prob_eq = qbs_prob_eq3" using qbs_prob_eq_1_implies_2 qbs_prob_eq_2_implies_3 qbs_prob_eq_3_implies_1 by blast lemma qbs_prob_eq_equiv14 : "qbs_prob_eq = qbs_prob_eq4" using qbs_prob_eq_2_implies_3 qbs_prob_eq_3_implies_1 qbs_prob_eq_1_implies_4 qbs_prob_eq_4_implies_3 qbs_prob_eq_1_implies_2 by blast lemma qbs_prob_eq_equiv23 : "qbs_prob_eq2 = qbs_prob_eq3" using qbs_prob_eq_1_implies_2 qbs_prob_eq_2_implies_3 qbs_prob_eq_3_implies_1 by blast lemma qbs_prob_eq_equiv24 : "qbs_prob_eq2 = qbs_prob_eq4" using qbs_prob_eq_2_implies_3 qbs_prob_eq_4_implies_3 qbs_prob_eq_3_implies_1 qbs_prob_eq_1_implies_4 qbs_prob_eq_1_implies_2 by blast lemma qbs_prob_eq_equiv34: "qbs_prob_eq3 = qbs_prob_eq4" using qbs_prob_eq_3_implies_1 qbs_prob_eq_1_implies_4 qbs_prob_eq_4_implies_3 by blast lemma qbs_prob_eq_equiv31 : "qbs_prob_eq = qbs_prob_eq3" using qbs_prob_eq_1_implies_2 qbs_prob_eq_2_implies_3 qbs_prob_eq_3_implies_1 by blast lemma qbs_prob_space_eq: assumes "qbs_prob_eq (X,\,\) (Y,\,\)" shows "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" using Quotient3_rel[OF Quotient3_qbs_prob_space] assms by blast lemma(in pair_qbs_prob) qbs_prob_space_eq: assumes "Y = X" and "distr \ (qbs_to_measure X) \ = distr \ (qbs_to_measure X) \" shows "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" using assms qbs_prob_eq_intro qbs_prob_space_eq by auto lemma(in pair_qbs_prob) qbs_prob_space_eq2: assumes "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" shows "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" using qbs_prob_space_eq assms qbs_prob_eq_2_implies_3[of "(X,\,\)" "(Y,\,\)"] qbs_prob_eq_3_implies_1[of "(X,\,\)" "(Y,\,\)"] qbs_prob_eq2_intro qbs_prob_eq_dest(4) by blast lemma(in pair_qbs_prob) qbs_prob_space_eq3: assumes "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M real_borel \ (\k\ qbs_space X. 0 \ f k) \ (\x. f (\ x) \ \) = (\x. f (\ x) \ \)" shows "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" using assms qbs_prob_eq_3_implies_1[of "(X,\,\)" "(Y,\,\)"] qbs_prob_eq3_intro qbs_prob_space_eq qbs_prob_eq_dest(4) by blast lemma(in pair_qbs_prob) qbs_prob_space_eq4: assumes "Y = X" and "\f. f \ qbs_to_measure X \\<^sub>M ennreal_borel \ (\\<^sup>+x. f (\ x) \ \) = (\\<^sup>+x. f (\ x) \ \)" shows "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" using assms qbs_prob_eq_4_implies_3[of "(X,\,\)" "(Y,\,\)"] qbs_prob_space_eq3[OF assms(1)] qbs_prob_eq3_dest(4) qbs_prob_eq4_intro - by blast + by blast lemma(in pair_qbs_prob) qbs_prob_space_eq_inverse: assumes "qbs_prob_space (X,\,\) = qbs_prob_space (Y,\,\)" shows "qbs_prob_eq (X,\,\) (Y,\,\)" and "qbs_prob_eq2 (X,\,\) (Y,\,\)" and "qbs_prob_eq3 (X,\,\) (Y,\,\)" and "qbs_prob_eq4 (X,\,\) (Y,\,\)" using Quotient3_rel[OF Quotient3_qbs_prob_space,of "(X, \, \)" "(Y,\,\)",simplified] assms qp1.qbs_prob_axioms qp2.qbs_prob_axioms by(simp_all add: qbs_prob_eq_equiv13[symmetric] qbs_prob_eq_equiv12[symmetric] qbs_prob_eq_equiv14[symmetric]) lift_definition qbs_prob_space_qbs :: "'a qbs_prob_space \ 'a quasi_borel" is fst by(auto simp add: qbs_prob_eq_def) lemma(in qbs_prob) qbs_prob_space_qbs_computation[simp]: "qbs_prob_space_qbs (qbs_prob_space (X,\,\)) = X" by(simp add: qbs_prob_space_qbs.abs_eq) lemma rep_qbs_prob_space': assumes "qbs_prob_space_qbs s = X" shows "\\ \. s = qbs_prob_space (X,\,\) \ qbs_prob X \ \" proof - obtain X' \ \ where hs: "s = qbs_prob_space (X', \, \)" "qbs_prob X' \ \" using rep_qbs_prob_space[of s] by auto then interpret qp:qbs_prob X' \ \ by simp show ?thesis using assms hs(2) by(auto simp add: hs(1)) qed lift_definition qbs_prob_ennintegral :: "['a qbs_prob_space, 'a \ ennreal] \ ennreal" is qbs_prob_t_ennintegral by(auto simp add: qbs_prob_t_ennintegral_def qbs_prob_eq_equiv14 qbs_prob_eq4_def) lift_definition qbs_prob_integral :: "['a qbs_prob_space, 'a \ real] \ real" is qbs_prob_t_integral by(auto simp add: qbs_prob_eq_equiv12 qbs_prob_t_integral_def qbs_prob_eq2_def) syntax "_qbs_prob_ennintegral" :: "pttrn \ ennreal \ 'a qbs_prob_space \ ennreal" ("\\<^sup>+\<^sub>Q((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sup>+\<^sub>Q x. f \p" \ "CONST qbs_prob_ennintegral p (\x. f)" syntax "_qbs_prob_integral" :: "pttrn \ real \ 'a qbs_prob_space \ real" ("\\<^sub>Q((2 _./ _)/ \_)" [60,61] 110) translations "\\<^sub>Q x. f \p" \ "CONST qbs_prob_integral p (\x. f)" text \ We define the function \l\<^sub>X \ L(P(X)) \\<^sub>M G(X)\. \ lift_definition qbs_prob_measure :: "'a qbs_prob_space \ 'a measure" is qbs_prob_t_measure by(auto simp add: qbs_prob_eq_def qbs_prob_t_measure_def) declare [[coercion qbs_prob_measure]] lemma(in qbs_prob) qbs_prob_measure_computation[simp]: "qbs_prob_measure (qbs_prob_space (X,\,\)) = distr \ (qbs_to_measure X) \" by (simp add: qbs_prob_measure.abs_eq qbs_prob_t_measure_def) definition qbs_emeasure ::"'a qbs_prob_space \ 'a set \ ennreal" where "qbs_emeasure s \ emeasure (qbs_prob_measure s)" lemma(in qbs_prob) qbs_emeasure_computation[simp]: assumes "U \ sets (qbs_to_measure X)" shows "qbs_emeasure (qbs_prob_space (X,\,\)) U = emeasure \ (\ -` U)" by(simp add: qbs_emeasure_def emeasure_distr[OF _ assms]) definition qbs_measure ::"'a qbs_prob_space \ 'a set \ real" where "qbs_measure s \ measure (qbs_prob_measure s)" interpretation qbs_prob_measure_prob_space : prob_space "qbs_prob_measure (s::'a qbs_prob_space)" for s proof(transfer,auto) fix X :: "'a quasi_borel" fix \ \ assume "qbs_prob_eq (X,\,\) (X,\,\)" then interpret qp: qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "prob_space (qbs_prob_t_measure (X,\,\))" by(simp add: qbs_prob_t_measure_def qp.prob_space_distr) qed lemma qbs_prob_measure_space: "qbs_space (qbs_prob_space_qbs s) = space (qbs_prob_measure s)" by(transfer,simp add: qbs_prob_t_measure_def) lemma qbs_prob_measure_sets[measurable_cong]: "sets (qbs_to_measure (qbs_prob_space_qbs s)) = sets (qbs_prob_measure s)" by(transfer,simp add: qbs_prob_t_measure_def) lemma(in qbs_prob) qbs_prob_ennintegral_def: assumes "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral (qbs_prob_space (X,\,\)) f = (\\<^sup>+x. f (\ x) \ \)" by (simp add: assms qbs_prob_ennintegral.abs_eq qbs_prob_t_ennintegral_def) lemma(in qbs_prob) qbs_prob_ennintegral_def2: assumes "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral (qbs_prob_space (X,\,\)) f = integral\<^sup>N (distr \ (qbs_to_measure X) \) f" using assms by(auto simp add: qbs_prob_ennintegral.abs_eq qbs_prob_t_ennintegral_def qbs_prob_t_measure_def nn_integral_distr) lemma (in qbs_prob) qbs_prob_ennintegral_not_morphism: assumes "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral (qbs_prob_space (X,\,\)) f = 0" by(simp add: assms qbs_prob_ennintegral.abs_eq qbs_prob_t_ennintegral_def) lemma qbs_prob_ennintegral_def2: assumes "qbs_prob_space_qbs s = (X :: 'a quasi_borel)" and "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral s f = integral\<^sup>N (qbs_prob_measure s) f" using assms proof(transfer,auto) fix X :: "'a quasi_borel" and \ \ f assume "qbs_prob_eq (X,\,\) (X,\,\)" and h:"f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" then interpret qp : qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "qbs_prob_t_ennintegral (X, \, \) f = integral\<^sup>N (qbs_prob_t_measure (X, \, \)) f" using qp.qbs_prob_ennintegral_def2[OF h] by(simp add: qbs_prob_ennintegral.abs_eq qbs_prob_t_measure_def) qed lemma(in qbs_prob) qbs_prob_integral_def: assumes "f \ X \\<^sub>Q real_quasi_borel" shows "qbs_prob_integral (qbs_prob_space (X,\,\)) f = (\x. f (\ x) \ \)" by (simp add: assms qbs_prob_integral.abs_eq qbs_prob_t_integral_def) lemma(in qbs_prob) qbs_prob_integral_def2: "qbs_prob_integral (qbs_prob_space (X,\,\)) f = integral\<^sup>L (distr \ (qbs_to_measure X) \) f" proof - consider "f \ X \\<^sub>Q \\<^sub>Q" | "f \ X \\<^sub>Q \\<^sub>Q" by auto thus ?thesis proof cases case h:2 then have "\ integrable (qbs_prob_measure (qbs_prob_space (X,\,\))) f" by auto thus ?thesis using h by(simp add: qbs_prob_integral.abs_eq qbs_prob_t_integral_def not_integrable_integral_eq) qed (auto simp add: qbs_prob_integral.abs_eq qbs_prob_t_integral_def integral_distr ) qed lemma qbs_prob_integral_def2: "qbs_prob_integral (s::'a qbs_prob_space) f = integral\<^sup>L (qbs_prob_measure s) f" proof(transfer,auto) fix X :: "'a quasi_borel" and \ \ f assume "qbs_prob_eq (X,\,\) (X,\,\)" then interpret qp : qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "qbs_prob_t_integral (X,\,\) f = integral\<^sup>L (qbs_prob_t_measure (X,\,\)) f" using qp.qbs_prob_integral_def2 by(simp add: qbs_prob_t_measure_def qbs_prob_integral.abs_eq) qed definition qbs_prob_var :: "'a qbs_prob_space \ ('a \ real) \ real" where "qbs_prob_var s f \ qbs_prob_integral s (\x. (f x - qbs_prob_integral s f)\<^sup>2)" lemma(in qbs_prob) qbs_prob_var_computation: assumes "f \ X \\<^sub>Q real_quasi_borel" shows "qbs_prob_var (qbs_prob_space (X,\,\)) f = (\x. (f (\ x) - (\x. f (\ x) \ \))\<^sup>2 \ \)" proof - have "(\x. (f x - qbs_prob_integral (qbs_prob_space (X, \, \)) f)\<^sup>2) \ X \\<^sub>Q \\<^sub>Q" using assms by auto thus ?thesis using assms qbs_prob_integral_def[of "\x. (f x - qbs_prob_integral (qbs_prob_space (X,\,\)) f)\<^sup>2"] by(simp add: qbs_prob_var_def qbs_prob_integral_def) qed lift_definition qbs_integrable :: "['a qbs_prob_space, 'a \ real] \ bool" is qbs_prob_t_integrable proof auto have H:"\ (X::'a quasi_borel) Y \ \ \ \ f. qbs_prob_eq (X,\,\) (Y,\,\) \ qbs_prob_t_integrable (X,\,\) f \ qbs_prob_t_integrable (Y,\,\) f" proof - fix X Y :: "'a quasi_borel" fix \ \ \ \ f assume H0:"qbs_prob_eq (X, \, \) (Y, \, \)" "qbs_prob_t_integrable (X, \, \) f" then interpret qp: pair_qbs_prob X \ \ Y \ \ by(auto intro!: pair_qbs_prob.intro simp: qbs_prob_eq_def) have [measurable]: "f \ qbs_to_measure X \\<^sub>M real_borel" and h: "integrable \ (f \ \)" using H0 by(auto simp: qbs_prob_t_integrable_def) note [simp] = qbs_prob_eq_dest(3)[OF H0(1)] - + show "qbs_prob_t_integrable (Y, \, \) f" unfolding qbs_prob_t_integrable_def proof auto have "integrable (distr \ (qbs_to_measure X) \) f" using h by(simp add: comp_def integrable_distr_eq) hence "integrable (distr \ (qbs_to_measure X) \) f" using qbs_prob_eq_dest(4)[OF H0(1)] by simp thus "integrable \ (f \ \)" by(simp add: comp_def integrable_distr_eq) qed qed fix X Y :: "'a quasi_borel" fix \ \ \ \ assume H0:"qbs_prob_eq (X, \, \) (Y, \, \)" then have H1:"qbs_prob_eq (Y, \, \) (X, \, \)" by(auto simp add: qbs_prob_eq_def) show "qbs_prob_t_integrable (X, \, \) = qbs_prob_t_integrable (Y, \, \)" using H[OF H0] H[OF H1] by auto qed lemma(in qbs_prob) qbs_integrable_def: "qbs_integrable (qbs_prob_space (X, \, \)) f = (f \ X \\<^sub>Q \\<^sub>Q \ integrable \ (f \ \))" by (simp add: qbs_integrable.abs_eq qbs_prob_t_integrable_def) lemma qbs_integrable_morphism: assumes "qbs_prob_space_qbs s = X" and "qbs_integrable s f" shows "f \ X \\<^sub>Q \\<^sub>Q" using assms by(transfer,auto simp: qbs_prob_t_integrable_def) lemma(in qbs_prob) qbs_integrable_measurable[simp,measurable]: assumes "qbs_integrable (qbs_prob_space (X,\,\)) f" shows "f \ qbs_to_measure X \\<^sub>M real_borel" using assms by(auto simp add: qbs_integrable_def) lemma qbs_integrable_iff_integrable: "(qbs_integrable (s::'a qbs_prob_space) f) = (integrable (qbs_prob_measure s) f)" apply transfer subgoal for s f proof(rule prod_cases3[where y=s],simp) fix X :: "'a quasi_borel" fix \ \ assume "qbs_prob_eq (X,\,\) (X,\,\)" then interpret qp: qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "qbs_prob_t_integrable (X,\,\) f = integrable (qbs_prob_t_measure (X,\,\)) f" (is "?lhs = ?rhs") using integrable_distr_eq[of \ \ "qbs_to_measure X" f] by(auto simp add: qbs_prob_t_integrable_def qbs_prob_t_measure_def comp_def) qed done lemma(in qbs_prob) qbs_integrable_iff_integrable_distr: "qbs_integrable (qbs_prob_space (X,\,\)) f = integrable (distr \ (qbs_to_measure X) \) f" by(simp add: qbs_integrable_iff_integrable) lemma(in qbs_prob) qbs_integrable_iff_integrable: assumes "f \ qbs_to_measure X \\<^sub>M real_borel" shows "qbs_integrable (qbs_prob_space (X,\,\)) f = integrable \ (\x. f (\ x))" by(auto intro!: integrable_distr_eq[of \ \ "qbs_to_measure X" f] simp: assms qbs_integrable_iff_integrable_distr) lemma qbs_integrable_if_integrable: assumes "integrable (qbs_prob_measure s) f" shows "qbs_integrable (s::'a qbs_prob_space) f" using assms by(simp add: qbs_integrable_iff_integrable) lemma integrable_if_qbs_integrable: assumes "qbs_integrable (s::'a qbs_prob_space) f" shows "integrable (qbs_prob_measure s) f" using assms by(simp add: qbs_integrable_iff_integrable) lemma qbs_integrable_iff_bounded: assumes "qbs_prob_space_qbs s = X" shows "qbs_integrable s f \ f \ X \\<^sub>Q \\<^sub>Q \ qbs_prob_ennintegral s (\x. ennreal \f x \) < \" (is "?lhs = ?rhs") proof - obtain \ \ where hs: "qbs_prob X \ \" "s = qbs_prob_space (X,\,\)" using rep_qbs_prob_space'[OF assms] by auto then interpret qp:qbs_prob X \ \ by simp have "?lhs = integrable (distr \ (qbs_to_measure X) \) f" by (simp add: hs(2) qbs_integrable_iff_integrable) also have "... = (f \ borel_measurable (distr \ (qbs_to_measure X) \) \ ((\\<^sup>+ x. ennreal (norm (f x)) \(distr \ (qbs_to_measure X) \)) < \))" by(rule integrable_iff_bounded) also have "... = ?rhs" proof - have [simp]:"f \ X \\<^sub>Q \\<^sub>Q \(\x. ennreal \f x\) \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" by auto have "f \ X \\<^sub>Q \\<^sub>Q \ qbs_prob_ennintegral s (\x. ennreal \f x\) = (\\<^sup>+ x. ennreal (norm (f x)) \qbs_prob_measure s)" using qp.qbs_prob_ennintegral_def2[of "\x. ennreal \f x\"] by(auto simp: hs(2)) thus ?thesis by(simp add: hs(2)) fastforce qed finally show ?thesis . qed lemma qbs_integrable_cong: assumes "qbs_prob_space_qbs s = X" "\x. x \ qbs_space X \ f x = g x" and "qbs_integrable s f" shows "qbs_integrable s g" - apply(rule qbs_integrable_if_integrable) - using integrable_cong[OF refl, of "qbs_prob_measure s" f g,simplified qbs_prob_measure_space[symmetric] assms(1),OF assms(2)] - qbs_integrable_iff_integrable[of s f] assms(3) - by simp + by (metis Bochner_Integration.integrable_cong assms integrable_if_qbs_integrable qbs_integrable_if_integrable qbs_prob_measure_space) lemma qbs_integrable_const[simp]: "qbs_integrable s (\x. c)" using qbs_integrable_iff_integrable[of s "\x. c"] by simp lemma qbs_integrable_add[simp]: assumes "qbs_integrable s f" and "qbs_integrable s g" shows "qbs_integrable s (\x. f x + g x)" - by(rule qbs_integrable_if_integrable[OF Bochner_Integration.integrable_add[OF integrable_if_qbs_integrable[OF assms(1)] integrable_if_qbs_integrable[OF assms(2)]]]) + using assms qbs_integrable_iff_integrable by blast lemma qbs_integrable_diff[simp]: assumes "qbs_integrable s f" and "qbs_integrable s g" shows "qbs_integrable s (\x. f x - g x)" by(rule qbs_integrable_if_integrable[OF Bochner_Integration.integrable_diff[OF integrable_if_qbs_integrable[OF assms(1)] integrable_if_qbs_integrable[OF assms(2)]]]) - + lemma qbs_integrable_mult_iff[simp]: "(qbs_integrable s (\x. c * f x)) = (c = 0 \ qbs_integrable s f)" - using qbs_integrable_iff_integrable[of s "\x. c * f x"] integrable_mult_left_iff[of "qbs_prob_measure s" c f] qbs_integrable_iff_integrable[of s f] + using qbs_integrable_iff_integrable[of s "\x. c * f x"] integrable_mult_left_iff[of "qbs_prob_measure s" c f] qbs_integrable_iff_integrable[of s f] by simp lemma qbs_integrable_mult[simp]: assumes "qbs_integrable s f" shows "qbs_integrable s (\x. c * f x)" using assms qbs_integrable_mult_iff by auto lemma qbs_integrable_abs[simp]: assumes "qbs_integrable s f" shows "qbs_integrable s (\x. \f x\)" by(rule qbs_integrable_if_integrable[OF integrable_abs[OF integrable_if_qbs_integrable[OF assms]]]) lemma qbs_integrable_sq[simp]: assumes "qbs_integrable s f" and "qbs_integrable s (\x. (f x)\<^sup>2)" shows "qbs_integrable s (\x. (f x - c)\<^sup>2)" by(simp add: comm_ring_1_class.power2_diff,rule qbs_integrable_diff,rule qbs_integrable_add) (simp_all add: comm_semiring_1_class.semiring_normalization_rules(16)[of 2] assms) lemma qbs_ennintegral_eq_qbs_integral: assumes "qbs_prob_space_qbs s = X" "qbs_integrable s f" and "\x. x \ qbs_space X \ 0 \ f x" shows "qbs_prob_ennintegral s (\x. ennreal (f x)) = ennreal (qbs_prob_integral s f)" using nn_integral_eq_integral[OF integrable_if_qbs_integrable[OF assms(2)]] assms qbs_prob_ennintegral_def2[OF assms(1) qbs_morphism_comp[OF qbs_integrable_morphism[OF assms(1,2)],of ennreal "\\<^sub>Q\<^sub>\\<^sub>0",simplified comp_def]] measurable_ennreal by (metis AE_I2 qbs_prob_integral_def2 qbs_prob_measure_space real.standard_borel_r_full_faithful) lemma qbs_prob_ennintegral_cong: assumes "qbs_prob_space_qbs s = X" and "\x. x \ qbs_space X \ f x = g x" shows "qbs_prob_ennintegral s f = qbs_prob_ennintegral s g" proof - obtain \ \ where hs: "s = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space'[OF assms(1)] by auto then interpret qp : qbs_prob X \ \ by simp have H1:"f \ \ = g \ \" using assms(2) unfolding comp_def apply standard using assms(2)[of "\ _"] by (simp add: qbs_Mx_to_X(2)) consider "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" | "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" by auto then have "qbs_prob_t_ennintegral (X,\,\) f = qbs_prob_t_ennintegral (X,\,\) g" proof cases case h:1 then have "g \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" using qbs_morphism_cong[of X f g "\\<^sub>Q\<^sub>\\<^sub>0"] assms by simp then show ?thesis using h H1 by(simp add: qbs_prob_t_ennintegral_def comp_def) next case h:2 then have "g \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" using assms qbs_morphism_cong[of X g f "\\<^sub>Q\<^sub>\\<^sub>0"] by auto then show ?thesis using h by(simp add: qbs_prob_t_ennintegral_def) qed thus ?thesis using hs(1) by (simp add: qbs_prob_ennintegral.abs_eq) qed lemma qbs_prob_ennintegral_const: "qbs_prob_ennintegral (s::'a qbs_prob_space) (\x. c) = c" using qbs_prob_ennintegral_def2[OF _ qbs_morphism_const[of c "\\<^sub>Q\<^sub>\\<^sub>0" "qbs_prob_space_qbs s",simplified]] by (simp add: qbs_prob_measure_prob_space.emeasure_space_1) lemma qbs_prob_ennintegral_add: assumes "qbs_prob_space_qbs s = X" "f \ (X::'a quasi_borel) \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" and "g \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral s (\x. f x + g x) = qbs_prob_ennintegral s f + qbs_prob_ennintegral s g" using qbs_prob_ennintegral_def2[of s X "\x. f x + g x"] qbs_prob_ennintegral_def2[OF assms(1,2)] qbs_prob_ennintegral_def2[OF assms(1,3)] assms nn_integral_add[of f "qbs_prob_measure s" g] by fastforce lemma qbs_prob_ennintegral_cmult: assumes "qbs_prob_space_qbs s = X" and "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" shows "qbs_prob_ennintegral s (\x. c * f x) = c * qbs_prob_ennintegral s f" using qbs_prob_ennintegral_def2[OF assms(1),of "\x. c * f x"] qbs_prob_ennintegral_def2[OF assms(1,2)] nn_integral_cmult[of f "qbs_prob_measure s"] assms by fastforce lemma qbs_prob_ennintegral_cmult_noninfty: assumes "c \ \" shows "qbs_prob_ennintegral s (\x. c * f x) = c * qbs_prob_ennintegral s f" proof - obtain X \ \ where hs: "s = qbs_prob_space (X, \, \)" "qbs_prob X \ \" using rep_qbs_prob_space[of s] by auto then interpret qp: qbs_prob X \ \ by simp consider "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" | "f \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" by auto then show ?thesis proof cases case 1 then show ?thesis by(auto intro!: qbs_prob_ennintegral_cmult[where X=X] simp: hs(1)) next case 2 consider "c = 0" | "c \ 0 \ c \ \" using assms by auto then show ?thesis proof cases case 1 then show ?thesis by(simp add: hs qbs_prob_ennintegral.abs_eq qbs_prob_t_ennintegral_def) next case h:2 have "(\x. c * f x) \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" proof(rule ccontr) assume "\ (\x. c * f x) \ X \\<^sub>Q \\<^sub>Q\<^sub>\\<^sub>0" hence 3:"(\x. c * f x) \ qbs_to_measure X \\<^sub>M ennreal_borel" by auto have "f = (\x. (1/c) * (c * f x))" using h by(simp add: divide_eq_1_ennreal ennreal_divide_times mult.assoc mult.commute[of c] mult_divide_eq_ennreal) also have "... \ qbs_to_measure X \\<^sub>M ennreal_borel" using 3 by simp finally show False using 2 by auto qed thus ?thesis using qp.qbs_prob_ennintegral_not_morphism 2 by(simp add: hs(1)) qed qed qed lemma qbs_prob_integral_cong: assumes "qbs_prob_space_qbs s = X" and "\x. x \ qbs_space X \ f x = g x" shows "qbs_prob_integral s f = qbs_prob_integral s g" by(simp add: qbs_prob_integral_def2) (metis Bochner_Integration.integral_cong assms(1) assms(2) qbs_prob_measure_space) lemma qbs_prob_integral_nonneg: assumes "qbs_prob_space_qbs s = X" and "\x. x \ qbs_space X \ 0 \ f x" shows "0 \ qbs_prob_integral s f" using qbs_prob_measure_space[of s] assms by(simp add: qbs_prob_integral_def2) lemma qbs_prob_integral_mono: assumes "qbs_prob_space_qbs s = X" "qbs_integrable (s :: 'a qbs_prob_space) f" "qbs_integrable s g" and "\x. x \ qbs_space X \ f x \ g x" shows "qbs_prob_integral s f \ qbs_prob_integral s g" using integral_mono[OF integrable_if_qbs_integrable[OF assms(2)] integrable_if_qbs_integrable[OF assms(3)] assms(4)[simplified qbs_prob_measure_space]] by(simp add: qbs_prob_integral_def2 assms(1) qbs_prob_measure_space[symmetric]) lemma qbs_prob_integral_const: "qbs_prob_integral (s::'a qbs_prob_space) (\x. c) = c" by(simp add: qbs_prob_integral_def2 qbs_prob_measure_prob_space.prob_space) lemma qbs_prob_integral_add: assumes "qbs_integrable (s::'a qbs_prob_space) f" and "qbs_integrable s g" shows "qbs_prob_integral s (\x. f x + g x) = qbs_prob_integral s f + qbs_prob_integral s g" using Bochner_Integration.integral_add[OF integrable_if_qbs_integrable[OF assms(1)] integrable_if_qbs_integrable[OF assms(2)]] by(simp add: qbs_prob_integral_def2) lemma qbs_prob_integral_diff: assumes "qbs_integrable (s::'a qbs_prob_space) f" and "qbs_integrable s g" shows "qbs_prob_integral s (\x. f x - g x) = qbs_prob_integral s f - qbs_prob_integral s g" using Bochner_Integration.integral_diff[OF integrable_if_qbs_integrable[OF assms(1)] integrable_if_qbs_integrable[OF assms(2)]] by(simp add: qbs_prob_integral_def2) lemma qbs_prob_integral_cmult: "qbs_prob_integral s (\x. c * f x) = c * qbs_prob_integral s f" by(simp add: qbs_prob_integral_def2) lemma real_qbs_prob_integral_def: assumes "qbs_integrable (s::'a qbs_prob_space) f" shows "qbs_prob_integral s f = enn2real (qbs_prob_ennintegral s (\x. ennreal (f x))) - enn2real (qbs_prob_ennintegral s (\x. ennreal (- f x)))" using assms proof(transfer,auto) fix X :: "'a quasi_borel" fix \ \ f assume H:"qbs_prob_eq (X,\,\) (X,\,\)" "qbs_prob_t_integrable (X,\,\) f" then interpret qp: qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "qbs_prob_t_integral (X,\,\) f = enn2real (qbs_prob_t_ennintegral (X,\,\) (\x. ennreal (f x))) - enn2real (qbs_prob_t_ennintegral (X,\,\) (\x. ennreal (- f x)))" using H(2) real_lebesgue_integral_def[of \ "f \ \"] by(auto simp add: comp_def qbs_prob_t_integrable_def qbs_prob_t_integral_def qbs_prob_t_ennintegral_def) qed lemma qbs_prob_var_eq: assumes "qbs_integrable (s::'a qbs_prob_space) f" and "qbs_integrable s (\x. (f x)\<^sup>2)" shows "qbs_prob_var s f = qbs_prob_integral s (\x. (f x)\<^sup>2) - (qbs_prob_integral s f)\<^sup>2" unfolding qbs_prob_var_def using assms proof(transfer,auto) fix X :: "'a quasi_borel" fix \ \ f assume H:"qbs_prob_eq (X,\,\) (X,\,\)" "qbs_prob_t_integrable (X,\,\) f" "qbs_prob_t_integrable (X,\,\) (\x. (f x)\<^sup>2)" then interpret qp: qbs_prob X \ \ by(simp add: qbs_prob_eq_def) show "qbs_prob_t_integral (X,\,\) (\x. (f x - qbs_prob_t_integral (X,\,\) f)\<^sup>2) = qbs_prob_t_integral (X,\,\) (\x. (f x)\<^sup>2) - (qbs_prob_t_integral (X,\,\) f)\<^sup>2" using H(2,3) prob_space.variance_eq[of \ "f \ \"] by(auto simp add: qbs_prob_t_integral_def qbs_prob_def qbs_prob_t_integrable_def comp_def qbs_prob_eq_def) qed lemma qbs_prob_var_affine: assumes "qbs_integrable s f" shows "qbs_prob_var s (\x. a * f x + b) = a\<^sup>2 * qbs_prob_var s f" (is "?lhs = ?rhs") proof - have "?lhs = qbs_prob_integral s (\x. (a * f x + b - (a * qbs_prob_integral s f + b))\<^sup>2)" using qbs_prob_integral_add[OF qbs_integrable_mult[OF assms,of a] qbs_integrable_const[of s b]] by (simp add: qbs_prob_integral_cmult qbs_prob_integral_const qbs_prob_var_def) also have "... = qbs_prob_integral s (\x. (a * f x - a * qbs_prob_integral s f)\<^sup>2)" by simp also have "... = qbs_prob_integral s (\x. a\<^sup>2 * (f x - qbs_prob_integral s f)\<^sup>2)" by (metis power_mult_distrib right_diff_distrib) also have "... = ?rhs" by(simp add: qbs_prob_var_def qbs_prob_integral_cmult[of s "a\<^sup>2"]) finally show ?thesis . qed lemma qbs_prob_integral_Markov_inequality: assumes "qbs_prob_space_qbs s = X" and "qbs_integrable s f" "\x. x \ qbs_space X \ 0 \ f x" and "0 < c" shows "qbs_emeasure s {x \ qbs_space X. c \ f x} \ ennreal (1/c * qbs_prob_integral s f)" using integral_Markov_inequality[OF integrable_if_qbs_integrable[OF assms(2)] _ assms(4)] assms(3) by(force simp add: qbs_prob_integral_def2 qbs_prob_measure_space qbs_emeasure_def assms(1) qbs_prob_measure_space[symmetric]) lemma qbs_prob_integral_Markov_inequality': assumes "qbs_prob_space_qbs s = X" "qbs_integrable s f" "\x. x \ qbs_space (qbs_prob_space_qbs s) \ 0 \ f x" and "0 < c" shows "qbs_measure s {x \ qbs_space (qbs_prob_space_qbs s). c \ f x} \ (1/c * qbs_prob_integral s f)" using qbs_prob_integral_Markov_inequality[OF assms] ennreal_le_iff[of "1/c * qbs_prob_integral s f" "qbs_measure s {x \ qbs_space (qbs_prob_space_qbs s). c \ f x}"] qbs_prob_integral_nonneg[of s X f,OF assms(1,3)] assms(4) by(simp add: qbs_measure_def qbs_emeasure_def qbs_prob_measure_prob_space.emeasure_eq_measure assms(1)) lemma qbs_prob_integral_Markov_inequality_abs: assumes "qbs_prob_space_qbs s = X" "qbs_integrable s f" and "0 < c" shows "qbs_emeasure s {x \ qbs_space X. c \ \f x\} \ ennreal (1/c * qbs_prob_integral s (\x. \f x\))" using qbs_prob_integral_Markov_inequality[OF assms(1) qbs_integrable_abs[OF assms(2)] _ assms(3)] by(simp add: assms(1)) lemma qbs_prob_integral_Markov_inequality_abs': assumes "qbs_prob_space_qbs s = X" "qbs_integrable s f" and "0 < c" shows "qbs_measure s {x \ qbs_space X. c \ \f x\} \ (1/c * qbs_prob_integral s (\x. \f x\))" using qbs_prob_integral_Markov_inequality'[OF assms(1) qbs_integrable_abs[OF assms(2)] _ assms(3)] by(simp add: assms(1)) lemma qbs_prob_integral_real_Markov_inequality: assumes "qbs_prob_space_qbs s = \\<^sub>Q" "qbs_integrable s f" and "0 < c" shows "qbs_emeasure s {r. c \ \f r\} \ ennreal (1/c * qbs_prob_integral s (\x. \f x\))" using qbs_prob_integral_Markov_inequality_abs[OF assms] by simp lemma qbs_prob_integral_real_Markov_inequality': assumes "qbs_prob_space_qbs s = \\<^sub>Q" "qbs_integrable s f" and "0 < c" shows "qbs_measure s {r. c \ \f r\} \ 1/c * qbs_prob_integral s (\x. \f x\)" using qbs_prob_integral_Markov_inequality_abs'[OF assms] by simp lemma qbs_prob_integral_Chebyshev_inequality: assumes "qbs_prob_space_qbs s = X" "qbs_integrable s f" "qbs_integrable s (\x. (f x)\<^sup>2)" and "0 < b" shows "qbs_measure s {x \ qbs_space X. b \ \f x - qbs_prob_integral s f\} \ 1 / b\<^sup>2 * qbs_prob_var s f" proof - have "qbs_integrable s (\x. \f x - qbs_prob_integral s f\\<^sup>2)" by(simp, rule qbs_integrable_sq[OF assms(2,3)]) moreover have "{x \ qbs_space X. b\<^sup>2 \ \f x - qbs_prob_integral s f\\<^sup>2} = {x \ qbs_space X. b \ \f x - qbs_prob_integral s f\}" by (metis (mono_tags, opaque_lifting) abs_le_square_iff abs_of_nonneg assms(4) less_imp_le power2_abs) ultimately show ?thesis using qbs_prob_integral_Markov_inequality'[OF assms(1),of "\x. \f x - qbs_prob_integral s f\^2" "b^2"] assms(4) by(simp add: qbs_prob_var_def assms(1)) qed -end \ No newline at end of file +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) + by (intro Bochner_Integration.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) + by (intro Bochner_Integration.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 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_real_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_real_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) + by (subst lborel_pair.integrable_product_swap_iff [symmetric], intro Bochner_Integration.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