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,3522 +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" + 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" + 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\ + 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 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 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/Disintegration/Disintegration.thy b/thys/Disintegration/Disintegration.thy --- a/thys/Disintegration/Disintegration.thy +++ b/thys/Disintegration/Disintegration.thy @@ -1,2524 +1,2525 @@ (* Title: Disintegration.thy Author: Michikazu Hirata, Tokyo Institute of Technology *) section \ Disintegration Theorem \ theory Disintegration imports "S_Finite_Measure_Monad.Kernels" "Lemmas_Disintegration" begin subsection \ Definition 14.D.2. (Mixture and Disintegration) \ context measure_kernel begin definition mixture_of :: "[('a \'b) measure, 'a measure] \ bool" where "mixture_of \ \ \ sets \ = sets (X \\<^sub>M Y) \ sets \ = sets X \ (\C\sets (X \\<^sub>M Y). \ C = (\\<^sup>+x. \\<^sup>+y. indicator C (x,y) \(\ x) \\))" definition disintegration :: "[('a \'b) measure, 'a measure] \ bool" where "disintegration \ \ \ sets \ = sets (X \\<^sub>M Y) \ sets \ = sets X \ (\A\sets X. \B\sets Y. \ (A \ B) = (\\<^sup>+x\A. (\ x B) \\))" lemma disintegrationI: assumes "sets \ = sets (X \\<^sub>M Y)" "sets \ = sets X" and "\A B. A \ sets X \ B \ sets Y \ \ (A \ B) = (\\<^sup>+x\A. (\ x B) \\)" shows "disintegration \ \" by(simp add: disintegration_def assms) lemma mixture_of_disintegration: assumes "mixture_of \ \" shows "disintegration \ \" unfolding disintegration_def proof safe fix A B assume [simp]:"A \ sets X" "B \ sets Y" have [simp,measurable_cong]: "sets \ = sets X" "space \ = space X" using assms by(auto simp: mixture_of_def intro!: sets_eq_imp_space_eq) have "A \ B \ sets (X \\<^sub>M Y)" by simp with assms have "\ (A \ B) = (\\<^sup>+x. \\<^sup>+y. indicator (A \ B) (x,y) \(\ x) \\)" by(simp add: mixture_of_def) also have "... = (\\<^sup>+x. \\<^sup>+y. indicator A x * indicator B y \(\ x) \\)" by(simp add: indicator_times) also have "... = (\\<^sup>+x\A. (\ x B) \\)" by(auto intro!: nn_integral_cong simp: kernel_sets nn_integral_cmult_indicator mult.commute) - finally show "emeasure \ (A \ B) = \\<^sup>+x\A. emeasure (\ x) B\\" . + finally show "emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B\\)" . qed(use assms[simplified mixture_of_def] in auto) lemma shows mixture_of_sets_eq: "mixture_of \ \ \ sets \ = sets (X \\<^sub>M Y)" "mixture_of \ \ \ sets \ = sets X" and mixture_of_space_eq: "mixture_of \ \ \ space \ = space (X \\<^sub>M Y)" "mixture_of \ \ \ space \ = space X" and disintegration_sets_eq: "disintegration \ \ \ sets \ = sets (X \\<^sub>M Y)" "disintegration \ \ \ sets \ = sets X" and disintegration_space_eq: "disintegration \ \ \ space \ = space (X \\<^sub>M Y)" "disintegration \ \ \ space \ = space X" by(auto simp: mixture_of_def disintegration_def intro!: sets_eq_imp_space_eq) lemma shows mixture_ofD: "mixture_of \ \ \ C \ sets (X \\<^sub>M Y) \ \ C = (\\<^sup>+x. \\<^sup>+y. indicator C (x,y) \(\ x) \\)" and disintegrationD: "disintegration \ \ \ A \ sets X \ B \ sets Y \ \ (A \ B) = (\\<^sup>+x\A. (\ x B) \\)" by(auto simp: mixture_of_def disintegration_def) lemma disintegration_restrict_space: assumes "disintegration \ \" "A \ space X \ sets X" shows "measure_kernel.disintegration (restrict_space X A) Y \ (restrict_space \ (A \ space Y)) (restrict_space \ A)" proof(rule measure_kernel.disintegrationI[OF restrict_measure_kernel[of A]]) have "sets (restrict_space \ (A \ space Y)) = sets (restrict_space (X \\<^sub>M Y) (A \ space Y))" by(auto simp: disintegration_sets_eq[OF assms(1)] intro!: sets_restrict_space_cong) also have "... = sets (restrict_space X A \\<^sub>M Y)" using sets_pair_restrict_space[of X A Y "space Y"] by simp finally show "sets (restrict_space \ (A \ space Y)) = sets (restrict_space X A \\<^sub>M Y)" . next show "sets (restrict_space \ A) = sets (restrict_space X A)" by(auto simp: disintegration_sets_eq[OF assms(1)] intro!: sets_restrict_space_cong) next fix a b assume h:"a \ sets (restrict_space X A)" "b \ sets Y" then have "restrict_space \ (A \ space Y) (a \ b) = \ (a \ b)" using sets.sets_into_space by(auto intro!: emeasure_restrict_space simp: disintegration_space_eq[OF assms(1)] disintegration_sets_eq[OF assms(1)] sets_restrict_space) (metis Sigma_Int_distrib1 assms(2) pair_measureI sets.top space_pair_measure) also have "... = (\\<^sup>+x\a. emeasure (\ x) b \\)" using sets_restrict_space_iff[OF assms(2)] h assms(1) by(auto simp: disintegration_def) also have "... = (\\<^sup>+x\A. (emeasure (\ x) b * indicator a x) \\)" using h(1) by(auto intro!: nn_integral_cong simp: sets_restrict_space) (metis IntD1 indicator_simps(1) indicator_simps(2) mult.comm_neutral mult_zero_right) also have "... = (\\<^sup>+x\a. emeasure (\ x) b \restrict_space \ A)" by (metis (no_types, lifting) assms disintegration_sets_eq(2) disintegration_space_eq(2) nn_integral_cong nn_integral_restrict_space) finally show "restrict_space \ (A \ space Y) (a \ b) = (\\<^sup>+x\a. emeasure (\ x) b \restrict_space \ A)" . qed end context subprob_kernel begin lemma countable_disintegration_AE_unique: assumes "countable (space Y)" and [measurable_cong]:"sets Y = Pow (space Y)" and "subprob_kernel X Y \'" "sigma_finite_measure \" and "disintegration \ \" "measure_kernel.disintegration X Y \' \ \" shows "AE x in \. \ x = \' x" proof - interpret \': subprob_kernel X Y \' by fact interpret s: sigma_finite_measure \ by fact have sets_eq[measurable_cong]: "sets \ = sets X" "sets \ = sets (X \\<^sub>M Y)" using assms(5) by(auto simp: disintegration_def) have 1:"AE x in \. \y \ space Y. \ x {y} = \' x {y}" unfolding AE_ball_countable[OF assms(1)] proof fix y assume y: "y \ space Y" show "AE x in \. emeasure (\ x) {y} = emeasure (\' x) {y}" proof(rule s.sigma_finite) fix J :: "nat \ _" assume J:"range J \ sets \" " \ (range J) = space \" "\i. emeasure \ (J i) \ \" from y have [measurable]: "(\x. \ x {y}) \ borel_measurable X" "(\x. \' x {y}) \ borel_measurable X" using emeasure_measurable \'.emeasure_measurable by auto define A where "A \ {x \ space \. \ x {y} \ \' x {y}}" have [measurable]:"A \ sets \" by(auto simp: A_def) have A: "\x. x \ space \ \ x \ A \ \' x {y} \ \ x {y}" by(auto simp: A_def) have 1: "AE x\A in \. \ x {y} = \' x {y}" proof - have "AE x in \. \n. (x \ A \ J n \ \ x {y} = \' x {y})" unfolding AE_all_countable proof fix n have ninf:"(\\<^sup>+x\A \ J n. (\ x) {y}\\) < \" proof - have "(\\<^sup>+x\A \ J n. (\ x) {y}\\) \ (\\<^sup>+x\A \ J n. (\ x) (space Y) \\)" using kernel_sets y by(auto intro!: nn_integral_mono emeasure_mono simp: indicator_def disintegration_space_eq(2)[OF assms(5)]) also have "... \ (\\<^sup>+x\A \ J n. 1 \\)" using subprob_space by(auto intro!: nn_integral_mono simp: indicator_def disintegration_space_eq(2)[OF assms(5)]) also have "... = \ (A \ J n)" using J by simp also have "... \ \ (J n)" using J by (auto intro!: emeasure_mono) also have "... < \" using J(3)[of n] by (simp add: top.not_eq_extremum) finally show ?thesis . qed have "(\\<^sup>+x. (\' x) {y} * indicator (A \ J n) x - (\ x) {y} * indicator (A \ J n) x \\) = (\\<^sup>+x\A \ J n. (\' x) {y} \\) - (\\<^sup>+x\A \ J n. (\ x) {y} \\)" using J ninf by(auto intro!: nn_integral_diff simp: indicator_def A_def) also have "... = 0" proof - have 0: "\ ((A \ J n) \ {y}) = (\\<^sup>+x\A \ J n. (\ x) {y} \\)" using J y sets_eq by(auto intro!: disintegrationD[OF assms(5),of "A \ J n" "{y}"]) have [simp]: "(\\<^sup>+x\A \ J n. (\' x) {y} \\) = \ ((A \ J n) \ {y})" using J y sets_eq by(auto intro!: \'.disintegrationD[OF assms(6),of "A \ J n" "{y}",symmetric]) show ?thesis using ninf by (simp add: 0 diff_eq_0_iff_ennreal) qed finally have assm:"AE x in \. (\' x) {y} * indicator (A \ J n) x - (\ x) {y} * indicator (A \ J n) x = 0" using J by(simp add: nn_integral_0_iff_AE) show "AE x\A \ J n in \. (\ x) {y} = (\' x) {y}" proof(rule AE_mp[OF assm]) show "AE x in \. emeasure (\' x) {y} * indicator (A \ J n) x - emeasure (\ x) {y} * indicator (A \ J n) x = 0 \ x \ A \ J n \ emeasure (\ x) {y} = emeasure (\' x) {y}" proof - { fix x assume h: "(\' x) {y} - (\ x) {y} = 0" "x \ A" have "(\ x) {y} = (\' x) {y}" using h(2) by(auto intro!: antisym ennreal_minus_eq_0[OF h(1)] simp: A_def) } thus ?thesis by(auto simp: indicator_def) qed qed qed hence "AE x\A \ (\ (range J))in \. \ x {y} = \' x {y}" by auto thus ?thesis using J(2) by auto qed have 2: "AE x\ (space \ - A) in \. \ x {y} = \' x {y}" proof - have "AE x in \. \n. x \ (space \ - A) \ J n \ \ x {y} = \' x {y}" unfolding AE_all_countable proof fix n have ninf:"(\\<^sup>+x\(space \ - A) \ J n. (\' x) {y}\\) < \" proof - have "(\\<^sup>+x\(space \ - A) \ J n. (\' x) {y}\\) \ (\\<^sup>+x\(space \ - A) \ J n. (\' x) (space Y) \\)" using kernel_sets y by(auto intro!: nn_integral_mono emeasure_mono simp: indicator_def disintegration_space_eq(2)[OF assms(5)]) also have "... \ (\\<^sup>+x\(space \ - A) \ J n. 1 \\)" using \'.subprob_space by(auto intro!: nn_integral_mono simp: indicator_def disintegration_space_eq(2)[OF assms(5)]) also have "... = \ ((space \ - A) \ J n)" using J by simp also have "... \ \ (J n)" using J by (auto intro!: emeasure_mono) also have "... < \" using J(3)[of n] by (simp add: top.not_eq_extremum) finally show ?thesis . qed have "(\\<^sup>+x. (\ x) {y} * indicator ((space \ - A) \ J n) x - (\' x) {y} * indicator ((space \ - A) \ J n) x \\) = (\\<^sup>+x\(space \ - A) \ J n. (\ x) {y} \\) - (\\<^sup>+x\(space \ - A) \ J n. (\' x) {y} \\)" using J ninf A by(auto intro!: nn_integral_diff simp: indicator_def) also have "... = 0" proof - have 0: "(\\<^sup>+x\(space \ - A) \ J n. (\' x) {y} \\) = \ (((space \ - A) \ J n) \ {y})" using J y sets_eq by(auto intro!: \'.disintegrationD[OF assms(6),of "(space \ - A) \ J n" "{y}",symmetric]) have [simp]: "\ (((space \ - A) \ J n) \ {y}) = (\\<^sup>+x\(space \ - A) \ J n. (\ x) {y} \\)" using J y sets_eq by(auto intro!: disintegrationD[OF assms(5),of "(space \ - A) \ J n" "{y}"]) show ?thesis using ninf by (simp add: 0 diff_eq_0_iff_ennreal) qed finally have assm:"AE x in \. (\ x) {y} * indicator ((space \ - A) \ J n) x - (\' x) {y} * indicator ((space \ - A) \ J n) x = 0" using J by(simp add: nn_integral_0_iff_AE) show "AE x\(space \ - A) \ J n in \. (\ x) {y} = (\' x) {y}" proof(rule AE_mp[OF assm]) show "AE x in \. emeasure (\ x) {y} * indicator ((space \ - A) \ J n) x - emeasure (\' x) {y} * indicator ((space \ - A) \ J n) x = 0 \ x \ (space \ - A) \ J n \ emeasure (\ x) {y} = emeasure (\' x) {y}" proof - { fix x assume h: "(\ x) {y} - (\' x) {y} = 0" "x \ space \" "x \ A" have "(\ x) {y} = (\' x) {y}" using A[OF h(2,3)] by(auto intro!: antisym ennreal_minus_eq_0[OF h(1)] simp: A_def) } thus ?thesis by(auto simp: indicator_def) qed qed qed hence "AE x\(space \ - A) \ (\ (range J))in \. \ x {y} = \' x {y}" by auto thus ?thesis using J(2) by auto qed show "AE x in \. \ x {y} = \' x {y}" using 1 2 by(auto simp: A_def) qed qed show ?thesis proof(rule AE_mp[OF 1]) { fix x assume x: "x \ space X" and h: "\y\space Y. \ x {y} = \' x {y}" have "\ x = \' x " by (simp add: \'.kernel_sets assms h kernel_sets measure_eqI_countable x) } thus " AE x in \. (\y\space Y. emeasure (\ x) {y} = emeasure (\' x) {y}) \ \ x = \' x" by(auto simp: sets_eq_imp_space_eq[OF sets_eq(1)]) qed qed end lemma(in subprob_kernel) nu_mu_spaceY_le: assumes "disintegration \ \" "A \ sets X" shows "\ (A \ space Y) \ \ A" proof - have "\ (A \ space Y) = (\\<^sup>+x\A. (\ x (space Y)) \\)" using assms by(simp add: disintegration_def) also have "... \ (\\<^sup>+x\A. 1 \\)" using assms subprob_space by(auto intro!: nn_integral_mono simp: disintegration_space_eq) (metis dual_order.refl indicator_simps(1) indicator_simps(2) mult.commute mult_1 mult_zero_right) also have "... = \ A" using assms by (simp add: disintegration_def) finally show ?thesis . qed context prob_kernel begin lemma countable_disintegration_AE_unique_prob: assumes "countable (space Y)" and [measurable_cong]:"sets Y = Pow (space Y)" and "prob_kernel X Y \'" "sigma_finite_measure \" and "disintegration \ \" "measure_kernel.disintegration X Y \' \ \" shows "AE x in \. \ x = \' x" by(auto intro!: countable_disintegration_AE_unique[OF assms(1,2) _ assms(4-6)] prob_kernel.subprob_kernel assms(3)) end subsection \ Lemma 14.D.3. \ lemma(in prob_kernel) nu_mu_spaceY: assumes "disintegration \ \" "A \ sets X" shows "\ (A \ space Y) = \ A" proof - have "\ (A \ space Y) = (\\<^sup>+x\A. (\ x (space Y)) \\)" using assms by(simp add: disintegration_def) also have "... = (\\<^sup>+x\A. 1 \\)" using assms by(auto intro!: nn_integral_cong simp: prob_space disintegration_space_eq) also have "... = \ A" using assms by (simp add: disintegration_def) finally show ?thesis . qed corollary(in subprob_kernel) nu_finite: assumes "disintegration \ \" "finite_measure \" shows "finite_measure \" proof have "\ (space \) = \ (space (X \\<^sub>M Y))" using assms by(simp add: disintegration_space_eq) also have "... \ \ (space \)" using assms by(simp add: nu_mu_spaceY_le disintegration_space_eq space_pair_measure) finally show "\ (space \) \ \" using assms(2) by (metis finite_measure.emeasure_finite infinity_ennreal_def neq_top_trans) qed corollary(in subprob_kernel) nu_subprob_space: assumes "disintegration \ \" "subprob_space \" shows "subprob_space \" proof have "\ (space \) = \ (space (X \\<^sub>M Y))" using assms by(simp add: disintegration_space_eq) also have "... \ \ (space \)" using assms by(simp add: nu_mu_spaceY_le disintegration_space_eq space_pair_measure) finally show "\ (space \) \ 1" using assms(2) order.trans subprob_space.emeasure_space_le_1 by auto next show "space \ \ {}" using Y_not_empty assms by(auto simp: disintegration_space_eq subprob_space_def subprob_space_axioms_def space_pair_measure) qed corollary(in prob_kernel) nu_prob_space: assumes "disintegration \ \" "prob_space \" shows "prob_space \" proof have "\ (space \) = \ (space (X \\<^sub>M Y))" using assms by(simp add: disintegration_space_eq) also have "... = \ (space \)" using assms by(simp add: nu_mu_spaceY disintegration_space_eq space_pair_measure) finally show "\ (space \) = 1" by (simp add: assms(2) prob_space.emeasure_space_1) qed lemma(in subprob_kernel) nu_sigma_finite: assumes "disintegration \ \" "sigma_finite_measure \" shows "sigma_finite_measure \" proof obtain A where A:"countable A" "A \ sets \" "\ A = space \" "\a\A. emeasure \ a \ \" using assms(2) by (meson sigma_finite_measure.sigma_finite_countable) have "countable {a \ space Y |a. a \ A}" using countable_image[OF A(1),of "\a. a \ space Y"] by (simp add: Setcompr_eq_image) moreover have "{a \ space Y |a. a \ A} \ sets \" using A(2) assms(1) disintegration_def by auto moreover have "\ {a \ space Y |a. a \ A} = space \" using assms A(3) by(simp add: disintegration_space_eq space_pair_measure) blast moreover have "\b\{a \ space Y |a. a \ A}. emeasure \ b \ \" using neq_top_trans[OF _ nu_mu_spaceY_le[OF assms(1)]] A(2,4) assms disintegration_sets_eq(2) by auto ultimately show "\B. countable B \ B \ sets \ \ \ B = space \ \ (\b\B. emeasure \ b \ \)" by blast qed subsection \ Theorem 14.D.4. (Measure Mixture Theorem) \ lemma(in measure_kernel) exist_nu: assumes "sets \ = sets X" shows "\\. disintegration \ \" proof - define \ where "\ = extend_measure (space X \ space Y) {(a, b). a \ sets X \ b \ sets Y} (\(a, b). a \ b) (\(a, b). \\<^sup>+x\a. emeasure (\ x) b\\) " have 1: "sets \ = sets (X \\<^sub>M Y)" proof - have "sets \ = sigma_sets (space X \ space Y) ((\(a, b). a \ b) ` {(a, b). a \ sets X \ b \ sets Y})" unfolding \_def by(rule sets_extend_measure) (use sets.space_closed[of X] sets.space_closed[of Y] in blast) also have "... = sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" by(auto intro!: sigma_sets_eqI) also have "... = sets (X \\<^sub>M Y)" by(simp add: sets_pair_measure) finally show ?thesis . qed have 2: "\ (A \ B) = (\\<^sup>+x\A. (\ x B) \\)" if "A \ sets X" "B \ sets Y" for A B proof(rule extend_measure_caratheodory_pair[OF \_def]) fix i j k l assume "i \ sets X \ j \ sets Y" "k \ sets X \ l \ sets Y" "i \ j = k \ l" then consider "i = k" "j = l" | "i \ j = {}" "k \ l = {}" by blast thus "(\\<^sup>+x\i. emeasure (\ x) j \\) = (\\<^sup>+x\k. emeasure (\ x) l \\)" by cases auto next fix A B j k assume h: "\n::nat. A n \ sets X \ B n \ sets Y" "j \ sets X \ k \ sets Y" "disjoint_family (\n. A n \ B n)" "(\i. A i \ B i) = j \ k" show "(\n. \\<^sup>+x\A n. emeasure (\ x) (B n)\\) = (\\<^sup>+x\j. emeasure (\ x) k\\)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+x. (\n. \ x (B n) * indicator (A n) x)\\)" proof(rule nn_integral_suminf[symmetric]) fix n have [measurable]:"(\x. emeasure (\ x) (B n)) \ borel_measurable \" "indicator (A n) \ borel_measurable \" using h(1)[of n] emeasure_measurable[of "B n"] assms(1) by auto thus "(\x. emeasure (\ x) (B n) * indicator (A n) x) \ borel_measurable \" by simp qed also have "... = ?rhs" proof(safe intro!: nn_integral_cong) fix x assume "x \ space \" consider "j = {}" | "k = {}" | "j \ {}" "k \ {}" by auto then show "(\n. emeasure (\ x) (B n) * indicator (A n) x) = emeasure (\ x) k * indicator j x" proof cases case 1 then have "\n. A n \ B n = {}" using h(4) by auto have "emeasure (\ x) (B n) * indicator (A n) x = 0" for n using \A n \ B n = {}\ by(auto simp: Sigma_empty_iff) thus ?thesis by(simp only: 1,simp) next case 2 then have "\n. A n \ B n = {}" using h(4) by auto have "emeasure (\ x) (B n) * indicator (A n) x = 0" for n using \A n \ B n = {}\ by(auto simp: Sigma_empty_iff) thus ?thesis by(simp only: 2,simp) next case 3 then have xinjiff:"x \ j \ (\i. \y\B i. (x,y) \ A i \ B i)" using h(4) by blast have bunk:"\ (B ` {i. x \ A i}) = k" if "x \ j" using that 3 h(4) by blast show ?thesis proof(cases "x \ j") case False then have "\n. x \ A n \ B n = {}" using h(4) 3 xinjiff by auto have "emeasure (\ x) (B n) * indicator (A n) x = 0" for n using \x \ A n \ B n = {}\ by auto thus ?thesis by(simp only:)(simp add: False) next case True then have [simp]: "emeasure (\ x) k * indicator j x = emeasure (\ x) k" by simp have "(\n. emeasure (\ x) (B n) * indicator (A n) x) = (\n. emeasure (\ x) (if x \ A n then B n else {}))" by(auto intro!: suminf_cong) also have "... = emeasure (\ x) (\n. if x \ A n then B n else {})" proof(rule suminf_emeasure) show "disjoint_family (\i. if x \ A i then B i else {})" using disjoint_family_onD[OF h(3)] by (auto simp: disjoint_family_on_def) next show "range (\i. if x \ A i then B i else {}) \ sets (\ x)" using h(1) kernel_sets[of x] \x \ space \\ sets_eq_imp_space_eq[OF assms(1)] by auto qed also have "... = emeasure (\ x) k" using True by(simp add: bunk) finally show ?thesis by simp qed qed qed finally show ?thesis . qed qed(use that in auto) show ?thesis using 1 2 assms by(auto simp: disintegration_def) qed lemma(in subprob_kernel) exist_unique_nu_sigma_finite': assumes "sets \ = sets X" "sigma_finite_measure \" shows "\!\. disintegration \ \" proof - obtain \ where disi: "disintegration \ \" using exist_nu[OF assms(1)] by auto with assms(2) interpret sf: sigma_finite_measure \ by(simp add: nu_sigma_finite) interpret \: sigma_finite_measure \ by fact show ?thesis proof(rule ex1I[where a=\]) fix \' assume disi':"disintegration \' \" show "\' = \" proof(rule \.sigma_finite_disjoint) fix A :: "nat \ _" assume A: "range A \ sets \" "\ (range A) = space \" "\i. emeasure \ (A i) \ \" "disjoint_family A" define B where "B \ \i. A i \ space Y" show "\' = \" proof(rule measure_eqI_generator_eq[where E=" {a \ b|a b. a \ sets X \ b \ sets Y}" and A=B and \="space X \ space Y"]) show "\C. C \ {a \ b |a b. a \ sets X \ b \ sets Y} \ emeasure \' C = emeasure \ C" "sets \' = sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" "sets \ = sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" using disi disi' by(auto simp: disintegration_def sets_pair_measure) next show "range B \ {a \ b |a b. a \ sets X \ b \ sets Y}" "\ (range B) = space X \ space Y" using A(1,2) by(auto simp: B_def assms(1) sets_eq_imp_space_eq[OF assms(1)]) next fix i show "emeasure \' (B i) \ \" using A(1) nu_mu_spaceY_le[OF disi',of "A i"] A(3)[of i] by(auto simp: B_def assms top.extremum_uniqueI) qed(simp_all add: Int_stable_pair_measure_generator pair_measure_closed) qed qed fact qed lemma(in subprob_kernel) exist_unique_nu_sigma_finite: assumes "sets \ = sets X" "sigma_finite_measure \" shows "\!\. disintegration \ \ \ sigma_finite_measure \" using assms exist_unique_nu_sigma_finite' nu_sigma_finite by blast lemma(in subprob_kernel) exist_unique_nu_finite: assumes "sets \ = sets X" "finite_measure \" shows "\!\. disintegration \ \ \ finite_measure \" using assms nu_finite finite_measure.sigma_finite_measure[OF assms(2)] exist_unique_nu_sigma_finite' by blast lemma(in subprob_kernel) exist_unique_nu_sub_prob_space: assumes "sets \ = sets X" "subprob_space \" shows "\!\. disintegration \ \ \ subprob_space \" using assms nu_subprob_space subprob_space_imp_sigma_finite[OF assms(2)] exist_unique_nu_sigma_finite' by blast lemma(in prob_kernel) exist_unique_nu_prob_space: assumes "sets \ = sets X" "prob_space \" shows "\!\. disintegration \ \ \ prob_space \" using assms nu_prob_space prob_space_imp_sigma_finite[OF assms(2)] exist_unique_nu_sigma_finite' by blast lemma(in subprob_kernel) nn_integral_fst_finite': assumes "f \ borel_measurable (X \\<^sub>M Y)" "disintegration \ \" "finite_measure \" shows "(\\<^sup>+z. f z \\) = (\\<^sup>+x. \\<^sup>+y. f (x,y) \(\ x) \\)" using assms(1) proof induction case (cong f g) have "integral\<^sup>N \ f = integral\<^sup>N \ g" using cong(3) by(auto intro!: nn_integral_cong simp: disintegration_space_eq(1)[OF assms(2)]) with cong(3) show ?case by(auto simp: cong(4) kernel_space disintegration_space_eq(2)[OF assms(2)] space_pair_measure intro!: nn_integral_cong) next case (set A) show ?case proof(rule sigma_sets_induct_disjoint[of "{a \ b|a b. a \ sets X \ b \ sets Y}" "space X \ space Y"]) show "A \ sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" using set by(simp add: sets_pair_measure) next fix A assume "A \ {a \ b |a b. a \ sets X \ b \ sets Y}" then obtain a b where ab: "A = a \ b" "a \ sets X" "b \ sets Y" by auto with assms(2) have "integral\<^sup>N \ (indicator A) = (\\<^sup>+x\a. (\ x b) \\)" by(simp add: disintegration_def) also have "... = (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" by(auto simp: ab(1) indicator_times disintegration_space_eq(2)[OF assms(2)] ab(3) kernel_sets mult.commute nn_integral_cmult_indicator intro!: nn_integral_cong) finally show "integral\<^sup>N \ (indicator A) = (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" . next fix A assume h: "A \ sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" "integral\<^sup>N \ (indicator A) = (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" show "integral\<^sup>N \ (indicator (space X \ space Y - A)) = (\\<^sup>+ x. \\<^sup>+ y. indicator (space X \ space Y - A) (x, y) \\ x \\)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ z. 1 - indicator A z \\)" by(auto intro!: nn_integral_cong simp: disintegration_space_eq(1)[OF assms(2)] space_pair_measure indicator_def) also have "... = (\\<^sup>+ z. 1 \\) - (\\<^sup>+ z. indicator A z \\)" proof(rule nn_integral_diff) show "integral\<^sup>N \ (indicator A) \ \" using h(1)[simplified sets_pair_measure[symmetric]] disintegration_sets_eq(1)[OF assms(2)] finite_measure.emeasure_finite[OF nu_finite[OF assms(2,3)]] by auto next show "indicator A \ borel_measurable \" using h(1)[simplified sets_pair_measure[symmetric]] disintegration_sets_eq(1)[OF assms(2)] by simp qed(simp_all add: indicator_def) also have "... = (\\<^sup>+ z. 1 \\) - (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" by(simp add: h(2)) also have "... = \ (space X \ space Y) - (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" using nn_integral_indicator[OF sets.top[of \]] by(simp add: space_pair_measure disintegration_space_eq(1)[OF assms(2)]) also have "... = (\\<^sup>+ x. \ x (space Y) \\) - (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" proof - have "\ (space X \ space Y) = (\\<^sup>+ x. \ x (space Y) \\)" using assms(2) by(auto simp: disintegration_def disintegration_space_eq(2)[OF assms(2)] intro!: nn_integral_cong) thus ?thesis by simp qed also have "... = (\\<^sup>+ x. \\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x \\) - (\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\)" proof - have "(\\<^sup>+ x. \ x (space Y) \\) = (\\<^sup>+ x. \\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x \\)" using kernel_sets by(auto intro!: nn_integral_cong simp: indicator_times disintegration_space_eq(2)[OF assms(2)] ) thus ?thesis by simp qed also have "... = (\\<^sup>+ x. (\\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x) - (\\<^sup>+ y. indicator A (x, y) \\ x) \\)" proof(rule nn_integral_diff[symmetric]) show "(\x. \\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x) \ borel_measurable \" "(\x. \\<^sup>+ y. indicator A (x, y) \\ x) \ borel_measurable \" by(use disintegration_sets_eq[OF assms(2)] nn_integral_measurable_f h(1)[simplified sets_pair_measure[symmetric]] in auto)+ next have "(\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\) \ (\\<^sup>+ x. \\<^sup>+ y. 1 \\ x \\)" by(rule nn_integral_mono)+ (simp add: indicator_def) also have "... \ (\\<^sup>+ x. 1 \\)" by(rule nn_integral_mono) (simp add: subprob_spaces disintegration_space_eq(2)[OF assms(2)] subprob_space.subprob_emeasure_le_1) also have "... < \" using finite_measure.emeasure_finite[OF assms(3)] by (simp add: top.not_eq_extremum) finally show "(\\<^sup>+ x. \\<^sup>+ y. indicator A (x, y) \\ x \\) \ \" by auto next have "A \ space X \ space Y" by (metis h(1) sets.sets_into_space sets_pair_measure space_pair_measure) hence "\x. (\\<^sup>+ y. indicator A (x, y) \\ x) \ (\\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x)" by(auto intro!: nn_integral_mono) thus "AE x in \. (\\<^sup>+ y. indicator A (x, y) \\ x) \ (\\<^sup>+ y. indicator (space X \ space Y) (x, y) \\ x)" by simp qed also have "... = (\\<^sup>+ x. (\\<^sup>+ y. indicator (space X \ space Y) (x, y) - indicator A (x, y) \\ x) \\)" proof(intro nn_integral_cong nn_integral_diff[symmetric]) fix x assume "x \ space \" then have "x \ space X" by(auto simp: disintegration_space_eq(2)[OF assms(2)]) with kernel_sets[OF this] h(1)[simplified sets_pair_measure[symmetric]] show "(\y. indicator (space X \ space Y) (x, y)) \ borel_measurable (\ x)" "(\y. indicator A (x, y)) \ borel_measurable (\ x)" by auto next fix x assume "x \ space \" then have "x \ space X" by(auto simp: disintegration_space_eq(2)[OF assms(2)]) have "(\\<^sup>+ y. indicator A (x, y) \\ x) \ (\\<^sup>+ y. 1 \\ x)" by(rule nn_integral_mono) (simp add: indicator_def) also have "... \ 1" using subprob_spaces[OF \x \ space X\] by (simp add: subprob_space.subprob_emeasure_le_1) also have "... < \" by auto finally show "(\\<^sup>+ y. indicator A (x, y) \\ x) \ \" by simp have "A \ space X \ space Y" by (metis h(1) sets.sets_into_space sets_pair_measure space_pair_measure) thus "AE y in \ x. indicator A (x, y) \ (indicator (space X \ space Y) (x, y) :: ennreal)" by auto qed also have "... = ?rhs" by(auto simp: indicator_def intro!: nn_integral_cong) finally show ?thesis . qed next fix A assume h:"disjoint_family A" "range A \ sigma_sets (space X \ space Y) {a \ b |a b. a \ sets X \ b \ sets Y}" "\i::nat. integral\<^sup>N \ (indicator (A i)) = (\\<^sup>+ x. \\<^sup>+ y. indicator (A i) (x, y) \\ x \\)" show "integral\<^sup>N \ (indicator (\ (range A))) = (\\<^sup>+ x. \\<^sup>+ y. indicator (\ (range A)) (x, y) \\ x \\)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ z. (\i. indicator (A i) z) \\)" by(simp add: suminf_indicator[OF h(1)]) also have "... = (\i. (\\<^sup>+ z. indicator (A i) z \\))" by(rule nn_integral_suminf) (use disintegration_sets_eq(1)[OF assms(2)] h(2)[simplified sets_pair_measure[symmetric]] in simp) also have "... = (\i. (\\<^sup>+ x. \\<^sup>+ y. indicator (A i) (x, y) \\ x \\))" by(simp add: h) also have "... = (\\<^sup>+ x. (\i. (\\<^sup>+ y. indicator (A i) (x, y) \\ x)) \\)" by(rule nn_integral_suminf[symmetric]) (use h(2)[simplified sets_pair_measure[symmetric]] disintegration_sets_eq(2)[OF assms(2)] nn_integral_measurable_f in simp) also have "... = (\\<^sup>+ x. \\<^sup>+ y. (\i. indicator (A i) (x, y)) \\ x \\)" using h(2)[simplified sets_pair_measure[symmetric]] kernel_sets by(auto intro!: nn_integral_cong nn_integral_suminf[symmetric] simp: disintegration_space_eq(2)[OF assms(2)]) also have "... = ?rhs" by(simp add: suminf_indicator[OF h(1)]) finally show ?thesis . qed qed(simp_all add: Int_stable_pair_measure_generator pair_measure_closed) next case (mult u c) show ?case (is "?lhs = ?rhs") proof - have "?lhs = c * (\\<^sup>+ z. u z \\)" using disintegration_sets_eq(1)[OF assms(2)] mult by(simp add: nn_integral_cmult) also have "... = c * ( \\<^sup>+ x. \\<^sup>+ y. u (x, y) \\ x \\)" by(simp add: mult) also have "... = ( \\<^sup>+ x. c * (\\<^sup>+ y. u (x, y) \\ x) \\)" using nn_integral_measurable_f'[OF mult(2)] disintegration_sets_eq(2)[OF assms(2)] by(simp add: nn_integral_cmult) also have "... = ?rhs" using mult by(auto intro!: nn_integral_cong nn_integral_cmult[symmetric] simp: disintegration_space_eq(2)[OF assms(2)]) finally show ?thesis . qed next case (add u v) show ?case (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ z. v z \\) + (\\<^sup>+ z. u z \\)" using add disintegration_sets_eq(1)[OF assms(2)] by (simp add: nn_integral_add) also have "... = (\\<^sup>+ x. \\<^sup>+ y. v (x, y) \\ x \\) + (\\<^sup>+ x. \\<^sup>+ y. u (x, y) \\ x \\)" by(simp add: add) also have "... = (\\<^sup>+ x. (\\<^sup>+ y. v (x, y) \\ x) + (\\<^sup>+ y. u (x, y) \\ x) \\)" using nn_integral_measurable_f'[OF add(1)] nn_integral_measurable_f'[OF add(3)] disintegration_sets_eq[OF assms(2)] by(auto intro!: nn_integral_add[symmetric]) also have "... = (\\<^sup>+ x. (\\<^sup>+ y. v (x, y) + u (x, y) \\ x) \\)" using add by(auto intro!: nn_integral_add[symmetric] nn_integral_cong simp: disintegration_space_eq(2)[OF assms(2)]) finally show ?thesis . qed next case (seq fi) have "(\\<^sup>+ y. (\ range fi) (x, y) \\ x) = (\i. \\<^sup>+ y. fi i (x, y) \\ x)" (is "?lhs = ?rhs") if "x \ space X" for x proof - have "?lhs = (\\<^sup>+ y. (\i. fi i (x, y)) \\ x)" by(metis SUP_apply) also have "... = ?rhs" proof(rule nn_integral_monotone_convergence_SUP) show "incseq (\i y. fi i (x, y))" using seq mono_compose by blast next fix i show "(\y. fi i (x, y)) \ borel_measurable (\ x)" using seq(1)[of i] that kernel_sets[OF that] by simp qed finally show ?thesis . qed have "integral\<^sup>N \ (\ range fi) = (\\<^sup>+ x. (\i. fi i x) \\)" by (metis SUP_apply) also have "... = (\i. integral\<^sup>N \ (fi i))" using disintegration_sets_eq(1)[OF assms(2)] seq(1,3) by(auto intro!: nn_integral_monotone_convergence_SUP) also have "... = (\i. \\<^sup>+ x. \\<^sup>+ y. fi i (x, y) \\ x \\)" by(simp add: seq) also have "... = (\\<^sup>+ x. (\i. \\<^sup>+ y. fi i (x, y) \\ x) \\)" proof(safe intro!: nn_integral_monotone_convergence_SUP[symmetric]) show "incseq (\i x. \\<^sup>+ y. fi i (x, y) \\ x)" using le_funD[OF incseq_SucD[OF seq(3)]] by(auto intro!: incseq_SucI le_funI nn_integral_mono) qed(use disintegration_sets_eq(2)[OF assms(2)] nn_integral_measurable_f'[OF seq(1)] in auto) also have "... = (\\<^sup>+ x. \\<^sup>+ y. (\i. fi i (x, y)) \\ x \\)" using kernel_sets seq(1) by(auto intro!: nn_integral_cong nn_integral_monotone_convergence_SUP[symmetric] simp: disintegration_space_eq(2)[OF assms(2)] mono_compose seq(3)) also have "... = (\\<^sup>+ x. \\<^sup>+ y. (\ range fi) (x, y) \\ x \\)" by(auto intro!: nn_integral_cong simp: image_image) finally show ?case . qed lemma(in prob_kernel) nn_integral_fst: assumes "f \ borel_measurable (X \\<^sub>M Y)" "disintegration \ \" "sigma_finite_measure \" shows "(\\<^sup>+z. f z \\) = (\\<^sup>+x. \\<^sup>+y. f (x,y) \(\ x) \\)" proof(rule sigma_finite_measure.sigma_finite_disjoint[OF assms(3)]) fix A assume A:"range A \ sets \" "\ (range A) = space \" "\i::nat. emeasure \ (A i) \ \" "disjoint_family A" then have A': "range (\i. A i \ space Y) \ sets \" "\ (range (\i. A i \ space Y)) = space \" "disjoint_family (\i. A i \ space Y)" by(auto simp: disintegration_sets_eq[OF assms(2)] disjoint_family_on_def disintegration_space_eq[OF assms(2)] space_pair_measure) blast show ?thesis (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+z\\ (range (\i. A i \ space Y)). f z \\)" using A'(2) by auto also have "... = (\i. \\<^sup>+z\ A i \ space Y. f z \\)" using A'(1,3) assms(1) disintegration_sets_eq[OF assms(2)] by(auto intro!: nn_integral_disjoint_family) also have "... = (\i. \\<^sup>+z. f z \restrict_space \ (A i \ space Y))" using A'(1) by(auto intro!: suminf_cong nn_integral_restrict_space[symmetric]) also have "... = (\i. \\<^sup>+x. \\<^sup>+y. f (x,y) \(\ x) \restrict_space \ (A i))" proof(safe intro!: suminf_cong) fix n interpret pk: prob_kernel "restrict_space X (A n)" Y \ by(rule restrict_probability_kernel) have An:"A n \ space X \ sets X" "A n \ space X = A n" using A(1) by(auto simp: disintegration_sets_eq[OF assms(2)]) have f:"f \ borel_measurable (restrict_space X (A n) \\<^sub>M Y)" proof - have 1:"sets (restrict_space X (A n) \\<^sub>M Y) = sets (restrict_space (X \\<^sub>M Y) (A n \ space Y))" using sets_pair_restrict_space[where Y=Y and B="space Y"] by simp show ?thesis using assms(1) by(simp add: measurable_cong_sets[OF 1 refl] measurable_restrict_space1) qed have fin: "finite_measure (restrict_space \ (A n))" by (metis A(1) A(3) UNIV_I emeasure_restrict_space finite_measureI image_subset_iff space_restrict_space space_restrict_space2 subset_eq) show "(\\<^sup>+z. f z \restrict_space \ (A n \ space Y)) = (\\<^sup>+x. \\<^sup>+y. f (x,y) \(\ x) \restrict_space \ (A n))" by(rule pk.nn_integral_fst_finite'[OF f disintegration_restrict_space[OF assms(2) An(1)] fin]) qed also have "... = (\i. \\<^sup>+x\A i. \\<^sup>+y. f (x,y) \(\ x) \\)" using A(1) by(auto intro!: suminf_cong nn_integral_restrict_space) also have "... = (\\<^sup>+x\\ (range A). \\<^sup>+y. f (x,y) \(\ x) \\)" using A(1,4) nn_integral_measurable_f'[OF assms(1)] disintegration_sets_eq[OF assms(2)] by(auto intro!: nn_integral_disjoint_family[symmetric]) also have "... = ?rhs" using A(2) by simp finally show ?thesis . qed qed lemma(in prob_kernel) integrable_eq1: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes [measurable]:"f \ borel_measurable (X \\<^sub>M Y)" and "disintegration \ \" "sigma_finite_measure \" shows "(\\<^sup>+ z. ennreal (norm (f z)) \\) < \ \ (\\<^sup>+x. \\<^sup>+y. ennreal (norm (f (x,y))) \(\ x) \\) < \" by(simp add: nn_integral_fst[OF _ assms(2,3)]) lemma(in prob_kernel) integrable_kernel_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable \ f" "disintegration \ \" "sigma_finite_measure \" shows "AE x in \. integrable (\ x) (\y. f (x,y))" proof - have [measurable]:"f \ borel_measurable (X \\<^sub>M Y)" using integrable_iff_bounded assms(1) disintegration_sets_eq[OF assms(2)] by simp show ?thesis unfolding integrable_iff_bounded proof - have 1:"(\\<^sup>+ x. \\<^sup>+ y. ennreal (norm (f (x,y))) \\ x \\) < \" using assms(1) integrable_eq1[OF _ assms(2,3),of f] by(simp add: integrable_iff_bounded) have "AE x in \. (\\<^sup>+ y. ennreal (norm (f (x,y))) \\ x) \ \" by(rule nn_integral_PInf_AE) (use 1 disintegration_sets_eq[OF assms(2)] nn_integral_measurable_f in auto) thus "AE x in \. (\y. f (x, y)) \ borel_measurable (\ x) \ (\\<^sup>+ y. ennreal (norm (f (x,y))) \\ x) < \" using top.not_eq_extremum by(fastforce simp: disintegration_space_eq[OF assms(2)]) qed qed lemma(in prob_kernel) integrable_lebesgue_integral_integrable': fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable \ f" "disintegration \ \" "sigma_finite_measure \" shows "integrable \ (\x. \y. f (x,y) \(\ x))" unfolding integrable_iff_bounded proof show "(\x. \y. f (x,y) \(\ x)) \ borel_measurable \" using disintegration_sets_eq[OF assms(2)] assms(1) integral_measurable_f'[of f] by(auto simp: integrable_iff_bounded) next have "(\\<^sup>+ x. ennreal (norm (\y. f (x,y) \(\ x))) \\) \ (\\<^sup>+ x. \\<^sup>+y. ennreal (norm (f (x,y))) \(\ x) \\)" using integral_norm_bound_ennreal integrable_kernel_integrable[OF assms] by(auto intro!: nn_integral_mono_AE) also have "... < \" using integrable_eq1[OF _ assms(2,3),of f] assms(1) disintegration_sets_eq[OF assms(2)] by(simp add: integrable_iff_bounded) finally show "(\\<^sup>+ x. ennreal (norm (\y. f (x,y) \(\ x))) \\) < \" . qed lemma(in prob_kernel) integrable_lebesgue_integral_integrable: fixes f :: "_ \_ \ _::{banach, second_countable_topology}" assumes "integrable \ (\(x,y). f x y)" "disintegration \ \" "sigma_finite_measure \" shows "integrable \ (\x. \y. f x y \(\ x))" using integrable_lebesgue_integral_integrable'[OF assms] by simp lemma(in prob_kernel) integral_fst: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "integrable \ f" "disintegration \ \" "sigma_finite_measure \" shows "(\z. f z \\) = (\x. \y. f (x,y) \(\ x) \\)" using assms(1) proof induct case b:(base A c) then have 0:"integrable \ (indicat_real A)" by blast then have 1[measurable]: "indicat_real A \ borel_measurable (X \\<^sub>M Y)" using disintegration_sets_eq[OF assms(2)] by auto have eq:"(\z. indicat_real A z \\) = (\x. \y. indicat_real A (x,y) \\ x \\)" (is "?lhs = ?rhs") proof - have "?lhs = enn2real (\\<^sup>+ z. ennreal (indicat_real A z) \\)" by(rule integral_eq_nn_integral) (use b in auto) also have "... = enn2real (\\<^sup>+ x. \\<^sup>+ y. ennreal (indicat_real A (x,y)) \\ x \\)" using nn_integral_fst[OF _ assms(2,3)] b disintegration_sets_eq[OF assms(2)] by auto also have "... = enn2real (\\<^sup>+ x. ennreal (\ y. indicat_real A (x,y) \\ x) \\)" proof - have "(\\<^sup>+ x. \\<^sup>+ y. ennreal (indicat_real A (x,y)) \\ x \\) = (\\<^sup>+ x. ennreal (\ y. indicat_real A (x,y) \\ x) \\)" proof(safe intro!: nn_integral_cong nn_integral_eq_integral) fix x assume "x \ space \" then have "x \ space X" by(simp add: disintegration_space_eq[OF assms(2)]) hence [simp]:"prob_space (\ x)" "sets (\ x) = sets Y" "space (\ x) = space Y" by(auto intro!: prob_spaces sets_eq_imp_space_eq kernel_sets) have [simp]:"{y. (x, y) \ A} \ sets Y" proof - have "{y. (x, y) \ A} = (\y. (x,y)) -` A \ space Y" using b(1)[simplified disintegration_sets_eq[OF assms(2)]] by auto also have "... \ sets Y" using b(1)[simplified disintegration_sets_eq[OF assms(2)]] \x \ space X\ by auto finally show ?thesis . qed have [simp]: "(\y. indicat_real A (x, y)) = indicat_real {y. (x,y) \ A}" by (auto simp: indicator_def) show "integrable (\ x) (\y. indicat_real A (x, y))" using prob_space.emeasure_le_1[of "\ x" "{y. (x, y) \ A}"] by(auto simp add: integrable_indicator_iff order_le_less_trans) qed simp thus ?thesis by simp qed also have "... = ?rhs" using disintegration_sets_eq[OF assms(2)] integral_measurable_f'[OF 1] by(auto intro!: integral_eq_nn_integral[symmetric]) finally show ?thesis . qed show ?case (is "?lhs = ?rhs") proof - have "?lhs = (\z. indicat_real A z \\) *\<^sub>R c" using 0 by auto also have "... = (\x. \y. indicat_real A (x,y) \\ x \\) *\<^sub>R c" by(simp only: eq) also have "... = (\x. (\y. indicat_real A (x,y) \\ x) *\<^sub>R c \\)" using integrable_lebesgue_integral_integrable'[OF 0 assms(2,3)] by(auto intro!: integral_scaleR_left[symmetric]) also have "... = ?rhs" using integrable_kernel_integrable[OF 0 assms(2,3)] integral_measurable_f'[of " indicat_real A"] integral_measurable_f'[of "\z. indicat_real A z *\<^sub>R c"] disintegration_sets_eq[OF assms(2)] by(auto intro!: integral_cong_AE) finally show ?thesis . qed next case fg:(add f g) note [measurable] = integrable_lebesgue_integral_integrable'[OF fg(1) assms(2,3)] integrable_lebesgue_integral_integrable'[OF fg(3) assms(2,3)] integrable_lebesgue_integral_integrable'[OF Bochner_Integration.integrable_add[OF fg(1,3)] assms(2,3)] show ?case (is "?lhs = ?rhs") proof - have "?lhs = (\x. (\y. f (x,y) \(\ x)) + (\y. g (x,y) \(\ x)) \\)" by(simp add: Bochner_Integration.integral_add[OF fg(1,3)] fg Bochner_Integration.integral_add[OF integrable_lebesgue_integral_integrable'[OF fg(1) assms(2,3)] integrable_lebesgue_integral_integrable'[OF fg(3) assms(2,3)]]) also have "... = ?rhs" using integrable_kernel_integrable[OF fg(1) assms(2,3)] integrable_kernel_integrable[OF fg(3) assms(2,3)] by(auto intro!: integral_cong_AE) finally show ?thesis . qed next case (lim f s) then have [measurable]: "f \ borel_measurable \" "\i. s i \ borel_measurable \" by auto show ?case proof (rule LIMSEQ_unique) show "(\i. integral\<^sup>L \ (s i)) \ integral\<^sup>L \ f" proof (rule integral_dominated_convergence) show "integrable \ (\x. 2 * norm (f x))" using lim(5) by auto qed(use lim in auto) next have "(\i. \ x. \ y. s i (x, y) \(\ x) \\) \ \ x. \ y. f (x, y) \(\ x) \\" proof (rule integral_dominated_convergence) have "AE x in \. \i. integrable (\ x) (\y. s i (x, y))" unfolding AE_all_countable using integrable_kernel_integrable[OF lim(1) assms(2,3)] .. with AE_space integrable_kernel_integrable[OF lim(5) assms(2,3)] show "AE x in \. (\i. \ y. s i (x, y) \(\ x)) \ \ y. f (x, y) \(\ x)" proof eventually_elim fix x assume x: "x \ space \" and s: "\i. integrable (\ x) (\y. s i (x, y))" and f: "integrable (\ x) (\y. f (x, y))" show "(\i. \ y. s i (x, y) \(\ x)) \ \ y. f (x, y) \(\ x)" proof (rule integral_dominated_convergence) show "integrable (\ x) (\y. 2 * norm (f (x, y)))" using f by auto show "AE xa in (\ x). (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) kernel_space by (auto simp: space_pair_measure disintegration_space_eq[OF assms(2)]) show "\i. AE xa in (\ x). norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) kernel_space by (auto simp: space_pair_measure disintegration_space_eq[OF assms(2)]) qed (use x disintegration_sets_eq[OF assms(2)] disintegration_space_eq[OF assms(2)] in auto) qed next show "integrable \ (\x. (\ y. 2 * norm (f (x, y)) \(\ x)))" using integrable_lebesgue_integral_integrable'[OF _ assms(2,3),of "\z. 2 * norm (f (fst z, snd z))"] lim(5) by auto next fix i show "AE x in \. norm (\ y. s i (x, y) \(\ x)) \ (\ y. 2 * norm (f (x, y)) \(\ x))" using AE_space integrable_kernel_integrable[OF lim(1) assms(2,3),of i] integrable_kernel_integrable[OF lim(5) assms(2,3)] proof eventually_elim case sf:(elim x) from sf(2) have "norm (\ y. s i (x, y) \(\ x)) \ (\\<^sup>+y. norm (s i (x, y)) \(\ x))" by (rule integral_norm_bound_ennreal) also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \(\ x))" using sf lim kernel_space by (auto intro!: nn_integral_mono simp: space_pair_measure disintegration_space_eq[OF assms(2)]) also have "\ = (\y. 2 * norm (f (x, y)) \(\ x))" using sf by (intro nn_integral_eq_integral) auto finally show "norm (\ y. s i (x, y) \(\ x)) \ (\ y. 2 * norm (f (x, y)) \(\ x))" by simp qed qed(use integrable_lebesgue_integral_integrable'[OF lim(1) assms(2,3)] integrable_lebesgue_integral_integrable'[OF lim(5) assms(2,3)] disintegration_sets_eq[OF assms(2)] in auto) then show "(\i. integral\<^sup>L \ (s i)) \ \ x. \ y. f (x, y) \(\ x) \\" using lim by simp qed qed subsection \ Marginal Measure \ definition marginal_measure_on :: "['a measure, 'b measure, ('a \ 'b) measure, 'b set] \ 'a measure" where "marginal_measure_on X Y \ B = measure_of (space X) (sets X) (\A. \ (A \ B))" abbreviation marginal_measure :: "['a measure, 'b measure, ('a \ 'b) measure] \ 'a measure" where "marginal_measure X Y \ \ marginal_measure_on X Y \ (space Y)" lemma space_marginal_measure: "space (marginal_measure_on X Y \ B) = space X" and sets_marginal_measure: "sets (marginal_measure_on X Y \ B) = sets X" by(simp_all add: marginal_measure_on_def) lemma emeasure_marginal_measure_on: assumes "sets \ = sets (X \\<^sub>M Y)" "B \ sets Y" "A \ sets X" shows "marginal_measure_on X Y \ B A = \ (A \ B)" unfolding marginal_measure_on_def proof(rule emeasure_measure_of_sigma) show "countably_additive (sets X) (\A. emeasure \ (A \ B))" proof(rule countably_additiveI) fix A :: "nat \ _" assume h:"range A \ sets X" "disjoint_family A" "\ (range A) \ sets X" have [simp]: "(\i. A i \ B) = (\ (range A) \ B)" by blast have "range (\i. A i \ B) \ sets \" "disjoint_family (\i. A i \ B)" using h assms(1,2) by(auto simp: disjoint_family_on_def) from suminf_emeasure[OF this] show " (\i. \ (A i \ B)) = \ (\ (range A) \ B)" by simp qed qed(insert assms, auto simp: positive_def sets.sigma_algebra_axioms) lemma emeasure_marginal_measure: assumes "sets \ = sets (X \\<^sub>M Y)" "A \ sets X" shows "marginal_measure X Y \ A = \ (A \ space Y)" using emeasure_marginal_measure_on[OF assms(1) _ assms(2)] by simp lemma finite_measure_marginal_measure_on_finite: assumes "finite_measure \" "sets \ = sets (X \\<^sub>M Y)" "B \ sets Y" shows "finite_measure (marginal_measure_on X Y \ B)" by (simp add: assms emeasure_marginal_measure_on finite_measure.emeasure_finite finite_measureI space_marginal_measure) lemma finite_measure_marginal_measure_finite: assumes "finite_measure \" "sets \ = sets (X \\<^sub>M Y)" shows "finite_measure (marginal_measure X Y \)" by(rule finite_measure_marginal_measure_on_finite[OF assms sets.top]) lemma marginal_measure_restrict_space: assumes "sets \ = sets (X \\<^sub>M Y)" "B \ sets Y" shows "marginal_measure X (restrict_space Y B) (restrict_space \ (space X \ B)) = marginal_measure_on X Y \ B" proof(rule measure_eqI) fix A assume "A \ sets (marginal_measure X (restrict_space Y B) (restrict_space \ (space X \ B)))" then have "A \ sets X" by(simp add: sets_marginal_measure) have 1:"sets (restrict_space \ (space X \ B)) = sets (X \\<^sub>M restrict_space Y B)" by (metis assms(1) restrict_space_space sets_pair_restrict_space sets_restrict_space_cong) show "emeasure (marginal_measure X (restrict_space Y B) (restrict_space \ (space X \ B))) A = emeasure (marginal_measure_on X Y \ B) A" apply(simp add: emeasure_marginal_measure_on[OF assms(1) assms(2) \A \ sets X\] emeasure_marginal_measure[OF 1 \A \ sets X\]) apply(simp add: space_restrict_space) by (metis Sigma_cong Sigma_mono \A \ sets X\ assms(1) assms(2) emeasure_restrict_space inf_le1 pair_measureI sets.Int_space_eq2 sets.sets_into_space sets.top) qed(simp add: sets_marginal_measure) lemma restrict_space_marginal_measure_on: assumes "sets \ = sets (X \\<^sub>M Y)" "B \ sets Y" "A \ sets X" shows "restrict_space (marginal_measure_on X Y \ B) A = marginal_measure_on (restrict_space X A) Y (restrict_space \ (A \ space Y)) B" proof(rule measure_eqI) fix A' assume "A' \ sets (restrict_space (marginal_measure_on X Y \ B) A)" then have h:"A' \ sets.restricted_space X A" by(simp add: sets_marginal_measure sets_restrict_space) show "emeasure (restrict_space (marginal_measure_on X Y \ B) A) A' = emeasure (marginal_measure_on (restrict_space X A) Y (restrict_space \ (A \ space Y)) B) A'" (is "?lhs = ?rhs") proof - have 1:"sets (restrict_space \ (A \ space Y)) = sets (restrict_space X A \\<^sub>M Y)" by (metis assms(1) restrict_space_space sets_pair_restrict_space sets_restrict_space_cong) have "?lhs = emeasure (marginal_measure_on X Y \ B) A'" using h by(auto intro!: emeasure_restrict_space simp: space_marginal_measure sets_marginal_measure assms) also have "... = \ (A' \ B)" using emeasure_marginal_measure_on[OF assms(1,2),of A'] h assms(3) by auto also have "... = restrict_space \ (A \ space Y) (A' \ B)" using h assms sets.sets_into_space by(auto intro!: emeasure_restrict_space[symmetric]) also have "... = ?rhs" using emeasure_marginal_measure_on[OF 1 assms(2),simplified sets_restrict_space,OF h] .. finally show ?thesis . qed qed(simp add: sets_marginal_measure sets_restrict_space) lemma restrict_space_marginal_measure: assumes "sets \ = sets (X \\<^sub>M Y)" "A \ sets X" shows "restrict_space (marginal_measure X Y \) A = marginal_measure (restrict_space X A) Y (restrict_space \ (A \ space Y))" using restrict_space_marginal_measure_on[OF assms(1) _ assms(2)] by simp lemma marginal_measure_mono: assumes "sets \ = sets (X \\<^sub>M Y)" "A \ sets Y" "B \ sets Y" "A \ B" shows "emeasure (marginal_measure_on X Y \ A) \ emeasure (marginal_measure_on X Y \ B)" proof(rule le_funI) fix U show "emeasure (marginal_measure_on X Y \ A) U \ emeasure (marginal_measure_on X Y \ B) U" proof - have 1:"U \ A \ U \ B" using assms(4) by auto show ?thesis proof(cases "U \ sets X") case True then show ?thesis by (simp add: "1" assms emeasure_marginal_measure_on emeasure_mono) next case False then show ?thesis by (simp add: emeasure_notin_sets sets_marginal_measure) qed qed qed lemma marginal_measure_absolutely_countinuous: assumes "sets \ = sets (X \\<^sub>M Y)" "A \ sets Y" "B \ sets Y" "A \ B" shows "absolutely_continuous (marginal_measure_on X Y \ B) (marginal_measure_on X Y \ A)" using emeasure_marginal_measure[OF assms(1)] assms(2,3) le_funD[OF marginal_measure_mono[OF assms]] by(auto intro!: mono_absolutely_continuous simp: sets_marginal_measure) lemma marginal_measure_absolutely_continuous': assumes "sets \ = sets (X \\<^sub>M Y)" "A \ sets Y" shows "absolutely_continuous (marginal_measure X Y \) (marginal_measure_on X Y \ A)" by(rule marginal_measure_absolutely_countinuous[OF assms sets.top sets.sets_into_space[OF assms(2)]]) subsection \ Lemma 14.D.6. \ locale sigma_finite_measure_on_pair = fixes X :: "'a measure" and Y :: "'b measure" and \ :: "('a \ 'b) measure" assumes nu_sets[measurable_cong]: "sets \ = sets (X \\<^sub>M Y)" and sigma_finite: "sigma_finite_measure \" begin abbreviation "\x \ marginal_measure X Y \" end locale projection_sigma_finite = fixes X :: "'a measure" and Y :: "'b measure" and \ :: "('a \ 'b) measure" assumes nu_sets[measurable_cong]: "sets \ = sets (X \\<^sub>M Y)" and marginal_sigma_finite: "sigma_finite_measure (marginal_measure X Y \)" begin sublocale \x : sigma_finite_measure "marginal_measure X Y \" by(rule marginal_sigma_finite) lemma \_sigma_finite: "sigma_finite_measure \" proof(rule \x.sigma_finite[simplified sets_marginal_measure space_marginal_measure]) fix A :: "nat \ _" assume A: "range A \ sets X" "\ (range A) = space X" "\i. marginal_measure X Y \ (A i) \ \" define C where "C \ range (\n. A n \ space Y)" have 1:"C \ sets \" "countable C" "\ C = space \" using nu_sets A(1,2) by(auto simp: C_def sets_eq_imp_space_eq[OF nu_sets] space_pair_measure) show "sigma_finite_measure \" unfolding sigma_finite_measure_def proof(safe intro!: exI[where x=C,simplified C_def]) fix n assume "\ (A n \ space Y) = \" moreover have "\ (A n \ space Y) \ \" using A(3)[of n] emeasure_marginal_measure[OF nu_sets,of "A n"] A(1) by auto ultimately show False by auto qed (use 1 C_def in auto) qed sublocale sigma_finite_measure_on_pair using \_sigma_finite by(auto simp: sigma_finite_measure_on_pair_def nu_sets) definition \' :: "'a \ 'b set \ ennreal" where "\' x B \ RN_deriv \x (marginal_measure_on X Y \ B) x" (* Notice that \' has the type "'a \ 'b set \ ennreal" not "'a \ 'b measure" because \' x is not a measure in general. *) lemma kernel_measurable[measurable]: "(\x. RN_deriv (marginal_measure X Y \) (marginal_measure_on X Y \ B) x) \ borel_measurable \x" by simp corollary \'_measurable[measurable]: "(\x. \' x B) \ borel_measurable X" using sets_marginal_measure[of X Y \ "space Y"] by(auto simp: \'_def) lemma kernel_RN_deriv: assumes "A \ sets X" "B \ sets Y" shows "\ (A \ B) = (\\<^sup>+x\A. \' x B \\x)" unfolding \'_def proof - have "emeasure \ (A \ B) = emeasure (density \x (RN_deriv \x (marginal_measure_on X Y \ B))) A" by (simp add: \x.density_RN_deriv assms emeasure_marginal_measure_on marginal_measure_absolutely_continuous' nu_sets sets_marginal_measure) then show "emeasure \ (A \ B) = set_nn_integral \x A (RN_deriv \x (marginal_measure_on X Y \ B))" by (simp add: assms(1) emeasure_density sets_marginal_measure) qed lemma empty_Y_bot: assumes "space Y = {}" shows "\ = \" proof - have "sets \ = {{}}" using nu_sets space_empty_iff[of "X \\<^sub>M Y",simplified space_pair_measure] assms by simp thus ?thesis by(simp add: sets_eq_bot) qed lemma empty_Y_nux: assumes "space Y = {}" shows "\x A = 0" proof(cases "A \ sets X") case True from emeasure_marginal_measure[OF nu_sets this] show ?thesis by(simp add: assms) next case False with sets_marginal_measure[of X Y \ "space Y"] show ?thesis by(auto intro!: emeasure_notin_sets) qed lemma kernel_empty0_AE: "AE x in \x. \' x {} = 0" unfolding \'_def by(rule AE_symmetric[OF \x.RN_deriv_unique]) (auto intro!: measure_eqI simp: sets_marginal_measure emeasure_density emeasure_marginal_measure_on[OF nu_sets]) lemma kernel_Y1_AE: "AE x in \x. \' x (space Y) = 1" unfolding \'_def by(rule AE_symmetric[OF \x.RN_deriv_unique]) (auto intro!: measure_eqI simp: emeasure_density) lemma kernel_suminf_AE: assumes "disjoint_family F" and "\i. F i \ sets Y" shows "AE x in \x. (\i. \' x (F i)) = \' x (\ (range F))" unfolding \'_def proof(rule \x.RN_deriv_unique) show "density \x (\x. \i. RN_deriv local.\x (marginal_measure_on X Y \ (F i)) x) = marginal_measure_on X Y \ (\ (range F))" proof(rule measure_eqI) fix A assume [measurable]:"A \ sets (density \x (\x. \i. RN_deriv \x (marginal_measure_on X Y \ (F i)) x))" then have [measurable]:"A \ sets \x" "A \ sets X" by(auto simp: sets_marginal_measure) show "(density \x (\x. \i. RN_deriv \x (marginal_measure_on X Y \ (F i)) x)) A = (marginal_measure_on X Y \ (\ (range F))) A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+x\A. (\i. RN_deriv \x (marginal_measure_on X Y \ (F i)) x)\\x)" by(auto intro!: emeasure_density) also have "... = (\\<^sup>+x. (\i. RN_deriv \x (marginal_measure_on X Y \ (F i)) x * indicator A x)\\x)" by simp also have "... = (\i. (\\<^sup>+x\A. RN_deriv \x (marginal_measure_on X Y \ (F i)) x \\x))" by(rule nn_integral_suminf) auto also have "... = (\i. \ (A \ F i))" using kernel_RN_deriv[of A "F _"] assms by(auto intro!: suminf_cong simp: \'_def) also have "... = \ (\i. A \ F i)" using assms nu_sets by(fastforce intro!: suminf_emeasure simp: disjoint_family_on_def) also have "... = \ (A \ (\i. F i))" proof - have "(\i. A \ F i) = (A \ (\i. F i))" by blast thus ?thesis by simp qed also have "... = ?rhs" using nu_sets assms by(auto intro!: emeasure_marginal_measure_on[symmetric]) finally show ?thesis . qed qed(simp add: sets_marginal_measure) qed auto lemma kernel_finite_sum_AE: assumes "disjoint_family_on F S" "finite S" and "\i. i \ S \ F i \ sets Y" shows "AE x in \x. (\i\S. \' x (F i)) = \' x (\i\S. F i)" proof - consider "S = {}" | "S \ {}" by auto then show ?thesis proof cases case 1 then show ?thesis by(simp add: kernel_empty0_AE) next case S:2 define F' where "F' \ (\n. if n < card S then F (from_nat_into S n) else {})" have F'[simp]:"\i. F' i \ sets Y" using assms(3) by (metis F'_def bot.extremum_strict bot_nat_def card.empty from_nat_into sets.empty_sets) have F'_disj: "disjoint_family F'" unfolding disjoint_family_on_def proof safe fix m n x assume h:"m \ n" "x \ F' m" "x \ F' n" consider "n < card S" "m < card S" | "n \ card S" | "m \ card S" by arith then show "x \ {}" proof cases case 1 then have "S \ {}" by auto with 1 have "from_nat_into S n \ S" "from_nat_into S m \ S" using from_nat_into[of S ] by blast+ moreover have "from_nat_into S n \ from_nat_into S m" by (metis "1"(1) "1"(2) assms(2) bij_betw_def h(1) lessThan_iff to_nat_on_finite to_nat_on_from_nat_into) ultimately show ?thesis using h assms(1) 1 by(auto simp: disjoint_family_on_def F'_def) qed(use h F'_def in simp_all) qed have 1:"(\i\S. \' x (F i)) = (\i' x (F' i))" for x unfolding F'_def by auto (metis (no_types, lifting) sum.card_from_nat_into sum.cong) have 2: "(\ (range F')) = (\i\S. F i)" proof safe fix x n assume h:"x \ F' n" then have "S \ {}" "n < card S" by(auto simp: F'_def) (meson empty_iff) with h show "x \ \ (F ` S)" by(auto intro!: exI[where x="from_nat_into S n"] simp: F'_def from_nat_into \S \ {}\) next fix x s assume "s \ S" "x \ F s" with bij_betwE[OF to_nat_on_finite[OF assms(2)]] show "x \ \ (range F')" by(auto intro!: exI[where x="to_nat_on S s"] simp: F'_def from_nat_into_to_nat_on[OF countable_finite[OF assms(2)]]) qed have "AE x in \x. (\i' x (F' i)) = (\i. \' x (F' i))" proof - have "AE x in \x. \i\card S. \' x (F' i) = 0" using kernel_empty0_AE by(auto simp: F'_def) hence "AE x in \x. (\i. \' x (F' i)) = (\i' x (F' i))" proof show "AE x in \x. (\i\card S. \' x (F' i) = 0) \ (\i. \' x (F' i)) = (\i' x (F' i))" proof - { fix x assume "\i\card S. \' x (F' i) = 0" then have "(\i. \' x (F' i)) = (\i' x (F' i))" using suminf_offset[of "\i. \' x (F' i)" "card S"] by(auto simp: F'_def) } thus ?thesis by auto qed qed thus ?thesis by auto qed moreover have "AE x in \x. (\i. \' x (F' i)) = \' x (\ (range F'))" using kernel_suminf_AE[OF F'_disj] by simp ultimately show ?thesis by(auto simp: 1 2) qed qed lemma kernel_disjoint_sum_AE: assumes "B \ sets Y" "C \ sets Y" and "B \ C = {}" shows "AE x in \x. \' x (B \ C) = \' x B + \' x C" proof - define F where "F \ \b. if b then B else C" have [simp]:"disjoint_family F" "\i. F i \ sets Y" "\x. (\i\UNIV. \' x (F i)) = \' x B + \' x C" "\ (range F) = B \ C" using assms by(auto simp: F_def disjoint_family_on_def comm_monoid_add_class.sum.Int_Diff[of UNIV _ "{True}"]) show ?thesis using kernel_finite_sum_AE[of F UNIV] by auto qed lemma kernel_mono_AE: assumes "B \ sets Y" "C \ sets Y" and "B \ C" shows "AE x in \x. \' x B \ \' x C" proof - have 1: "B \ (C - B) = C" using assms(3) by auto have "AE x in \x. \' x C = \' x B + \' x (C - B)" using assms by(auto intro!: kernel_disjoint_sum_AE[of B "C - B",simplified 1]) thus ?thesis by auto qed lemma kernel_incseq_AE: assumes "range B \ sets Y" "incseq B" shows "AE x in \x. incseq (\n. \' x (B n))" using assms(1) by(auto simp: incseq_Suc_iff AE_all_countable intro!: kernel_mono_AE[OF _ _ incseq_SucD[OF assms(2)]]) lemma kernel_decseq_AE: assumes "range B \ sets Y" "decseq B" shows "AE x in \x. decseq (\n. \' x (B n))" using assms(1) by(auto simp: decseq_Suc_iff AE_all_countable intro!: kernel_mono_AE[OF _ _ decseq_SucD[OF assms(2)]]) corollary kernel_01_AE: assumes "B \ sets Y" shows "AE x in \x. 0 \ \' x B \ \' x B \ 1" proof - have "{} \ B" "B \ space Y" using assms sets.sets_into_space by auto from kernel_empty0_AE kernel_Y1_AE kernel_mono_AE[OF _ _ this(1)] kernel_mono_AE[OF _ _ this(2)] assms show ?thesis by auto qed lemma kernel_get_0: "0 \ \' x B" by simp lemma kernel_le1_AE: assumes "B \ sets Y" shows "AE x in \x. \' x B \ 1" using kernel_01_AE[OF assms] by auto corollary kernel_n_infty: assumes "B \ sets Y" shows "AE x in \x. \' x B \ \" by(rule AE_mp[OF kernel_le1_AE[OF assms]],standard) (auto simp: neq_top_trans[OF ennreal_one_neq_top]) corollary kernel_le_infty: assumes "B \ sets Y" shows "AE x in \x. \' x B < \" using kernel_n_infty[OF assms] by (simp add: top.not_eq_extremum) lemma kernel_SUP_incseq: assumes "range B \ sets Y" "incseq B" shows "AE x in \x. \' x (\ (range B)) = (\n. \' x (B n))" proof - define Bn where "Bn \ (\n. if n = 0 then {} else B (n - 1))" have "incseq Bn" using assms(2) by(auto simp: Bn_def incseq_def) define Cn where "Cn \ (\n. Bn (Suc n) - Bn n)" have Cn_simp: "Cn 0 = B 0" "Cn (Suc n) = B (Suc n) - B n" for n by(simp_all add: Cn_def Bn_def) have Cn_sets:"Cn n \ sets Y" for n using assms(1) by(induction n) (auto simp: Cn_simp) have Cn_disj: "disjoint_family Cn" by(auto intro!: disjoint_family_Suc[OF ] incseq_SucD[OF \incseq Bn\] simp: Cn_def) have Cn_un: "(\kx. \n. (\i' x (Cn i)) = \' x (B n)" unfolding AE_all_countable using kernel_finite_sum_AE[OF disjoint_family_on_mono[OF _ Cn_disj],of "{.. (range B)) = (\ (range Cn))" (is "?lhs = ?rhs") proof safe fix n x assume "x \ B n" with Cn_un[of n] show "x \ \ (range Cn)" by blast next fix n x assume "x \ Cn n" then show "x \ \ (range B)" by(cases n,auto simp: Cn_simp) qed hence "AE x in \x. \' x (\ (range B)) = \' x (\ (range Cn))" by simp moreover have "AE x in \x. \' x (\ (range Cn)) = (\n. \' x (Cn n))" by(rule AE_symmetric[OF kernel_suminf_AE[OF Cn_disj]]) (use Cn_def Bn_def assms(1) in auto) moreover have "AE x in \x. (\n. \' x (Cn n)) = (\n. \i' x (Cn i))" by(auto simp: suminf_eq_SUP) moreover have "AE x in \x. (\n. \i' x (Cn i)) = (\n. \i' x (Cn i))" proof(intro AE_I2 antisym) fix x show "(\n. \i' x (Cn i)) \ (\n. \i' x (Cn i))" by(rule complete_lattice_class.Sup_mono, auto, use le_iff_add in blast) next fix x show "(\n. \i' x (Cn i)) \ (\n. \i' x (Cn i))" by(rule complete_lattice_class.Sup_mono) blast qed moreover have "AE x in \x. (\n. \i' x (Cn i)) = (\n. \' x (B n))" by(rule AE_mp[OF Cn_sum_Bn]) (standard+, auto) ultimately show ?thesis by auto qed lemma kernel_lim_incseq: assumes "range B \ sets Y" "incseq B" shows "AE x in \x. (\n. \' x (B n)) \ \' x (\ (range B))" by(rule AE_mp[OF AE_conjI[OF kernel_SUP_incseq[OF assms] kernel_incseq_AE[OF assms]]],auto simp: LIMSEQ_SUP) lemma kernel_INF_decseq: assumes "range B \ sets Y" "decseq B" shows "AE x in \x. \' x (\ (range B)) = (\n. \' x (B n))" proof - define C where "C \ (\k. space Y - B k)" have C:"range C \ sets Y" "incseq C" using assms by(auto simp: C_def decseq_def incseq_def) have eq1: "AE x in \x. 1 - \' x (\ (range B)) = \' x (\ (range C))" proof - have "AE x in \x. \' x (\ (range C)) + \' x (\ (range B)) - \' x (\ (range B)) = \' x (\ (range C))" using assms(1) kernel_n_infty[of "\ (range B)"] by auto moreover have "AE x in \x. \' x (\ (range C)) + \' x (\ (range B)) = 1" proof - have [simp]:"(\ (range C)) \ (\ (range B)) = space Y" "(\ (range C)) \ (\ (range B)) = {}" by(auto simp: C_def) (meson assms(1) range_subsetD sets.sets_into_space subsetD) from kernel_disjoint_sum_AE[OF _ _ this(2)] C(1) assms(1) kernel_Y1_AE show ?thesis by auto qed ultimately show ?thesis by auto qed have eq2: "AE x in \x. \' x (\ (range C)) = (\n. \' x (C n))" using kernel_SUP_incseq[OF C] by auto have eq3: "AE x in \x. (\n. \' x (C n)) = (\n. 1 - \' x (B n))" proof - have "AE x in \x. \n. \' x (C n) = 1 - \' x (B n)" unfolding AE_all_countable proof safe fix n have "AE x in \x. \' x (C n) + \' x (B n) - \' x (B n) = \' x (C n)" using assms(1) kernel_n_infty[of "B n"] by auto moreover have "AE x in \x. \' x (C n) + \' x (B n) = 1" proof - have [simp]: "C n \ B n = space Y" "C n \ B n = {}" by(auto simp: C_def) (meson assms(1) range_subsetD sets.sets_into_space subsetD) thus ?thesis using kernel_disjoint_sum_AE[of "C n" "B n"] C(1) assms(1) kernel_Y1_AE by fastforce qed ultimately show "AE x in \x. \' x (C n) = 1 - \' x (B n)" by auto qed thus ?thesis by auto qed have [simp]: "(\n. 1 - \' x (B n)) = 1 - (\n. \' x (B n))" for x by(auto simp: ennreal_INF_const_minus) have eq: "AE x in \x. 1 - \' x (\ (range B)) = 1 - (\n. \' x (B n))" using eq1 eq2 eq3 by auto have le1: "AE x in \x. (\n. \' x (B n)) \ 1" proof - have "AE x in \x. \n. \' x (B n) \ 1" using assms(1) by(auto intro!: kernel_le1_AE simp: AE_all_countable) thus ?thesis by (auto simp: INF_lower2) qed show ?thesis by(rule AE_mp[OF AE_conjI[OF AE_conjI[OF eq le1] kernel_le1_AE[of "\ (range B)"]]]) (insert assms(1),auto simp: ennreal_minus_cancel[OF ennreal_one_neq_top]) qed lemma kernel_lim_decseq: assumes "range B \ sets Y" "decseq B" shows "AE x in \x. (\n. \' x (B n)) \ \' x (\ (range B))" by(rule AE_mp[OF AE_conjI[OF kernel_INF_decseq[OF assms] kernel_decseq_AE[OF assms]]],standard,auto simp: LIMSEQ_INF) end lemma qlim_eq_lim_mono_at_bot: fixes g :: "rat \ 'a :: linorder_topology" assumes "mono f" "(g \ a) at_bot" "\r::rat. f (real_of_rat r) = g r" shows "(f \ a) at_bot" proof - have "mono g" by(metis assms(1,3) mono_def of_rat_less_eq) have ga:"\r. g r \ a" proof(rule ccontr) fix r assume "\ a \ g r" then have "g r < a" by simp from order_topology_class.order_tendstoD(1)[OF assms(2) this] obtain Q :: rat where q: "\q. q \ Q \ g r < g q" by(auto simp: eventually_at_bot_linorder) define q where "q \ min r Q" show False using q[of q] \mono g\ by(auto simp: q_def mono_def) (meson linorder_not_less min.cobounded1) qed show ?thesis proof(rule decreasing_tendsto) show "\\<^sub>F n in at_bot. a \ f n" unfolding eventually_at_bot_linorder by(rule exI[where x=undefined],auto) (metis Ratreal_def assms(1,3) dual_order.trans ga less_eq_real_def lt_ex monoD of_rat_dense) (*metis assms(1) assms(3) ga less_eq_real_def lfp.leq_trans lt_ex monoD of_rat_dense*) next fix x assume "a < x" with topological_space_class.topological_tendstoD[OF assms(2),of "{..q. q \ Q \ g q < x" by(auto simp: eventually_at_bot_linorder) show "\\<^sub>F n in at_bot. f n < x" using q assms(1,3) by(auto intro!: exI[where x="real_of_rat Q"] simp: eventually_at_bot_linorder) (metis dual_order.refl monoD order_le_less_trans) qed qed lemma qlim_eq_lim_mono_at_top: fixes g :: "rat \ 'a :: linorder_topology" assumes "mono f" "(g \ a) at_top" "\r::rat. f (real_of_rat r) = g r" shows "(f \ a) at_top" proof - have "mono g" by(metis assms(1,3) mono_def of_rat_less_eq) have ga:"\r. g r \ a" proof(rule ccontr) fix r assume "\ g r \ a" then have "a < g r" by simp from order_topology_class.order_tendstoD(2)[OF assms(2) this] obtain Q :: rat where q: "\q. Q \ q \ g q < g r" by(auto simp: eventually_at_top_linorder) define q where "q \ max r Q" show False using q[of q] \mono g\ by(auto simp: q_def mono_def leD) qed show ?thesis proof(rule increasing_tendsto) show "\\<^sub>F n in at_top. f n \ a" unfolding eventually_at_top_linorder by(rule exI[where x=undefined],auto) (metis (no_types, opaque_lifting) assms(1) assms(3) dual_order.trans ga gt_ex monoD of_rat_dense order_le_less) next fix x assume "x < a" with topological_space_class.topological_tendstoD[OF assms(2),of "{x<..}"] obtain Q :: rat where q: "\q. Q \ q \ x < g q" by(auto simp: eventually_at_top_linorder) show "\\<^sub>F n in at_top. x < f n" using q assms(1,3) by(auto simp: eventually_at_top_linorder intro!: exI[where x="real_of_rat Q"]) (metis dual_order.refl monoD order_less_le_trans) qed qed subsection \ Theorem 14.D.10. (Measure Disintegration Theorem) \ locale projection_sigma_finite_standard = projection_sigma_finite + standard_borel_ne Y begin theorem measure_disintegration: "\\. prob_kernel X Y \ \ measure_kernel.disintegration X Y \ \ \x \ (\\''. prob_kernel X Y \'' \ measure_kernel.disintegration X Y \'' \ \x \ (AE x in \x. \ x = \'' x))" proof - have *:"\\. prob_kernel X (borel :: real measure) \ \ measure_kernel.disintegration X borel \ \ (marginal_measure X borel \) \ (\\''. prob_kernel X borel \'' \ measure_kernel.disintegration X borel \'' \ (marginal_measure X borel \) \ (AE x in (marginal_measure X borel \). \ x = \'' x))" if nu_sets': "sets \ = sets (X \\<^sub>M borel)" and marginal_sigma_finite': "sigma_finite_measure (marginal_measure X borel \)" for X :: "'a measure" and \ proof - interpret r: projection_sigma_finite X borel \ using that by(auto simp: projection_sigma_finite_def) define \ :: "'a \ rat \ real" where "\ \ (\x r. enn2real (r.\' x {..real_of_rat r}))" have as1: "AE x in r.\x. \r s. r \ s \ \ x r \ \ x s" unfolding AE_all_countable proof(safe intro!: AE_impI) fix r s :: rat assume "r \ s" have "AE x in r.\x. r.\' x {..real_of_rat k} < top" for k using atMost_borel r.kernel_le_infty by blast from this[of s] r.kernel_mono_AE[of "{..real_of_rat r}" "{..real_of_rat s}"] \r \ s\ show "AE x in r.\x. \ x r \ \ x s" by(auto simp: \_def of_rat_less_eq enn2real_mono) qed have as2: "AE x in r.\x. \r. (\n. \ x (r + 1 / rat_of_nat (Suc n))) \ \ x r" unfolding AE_all_countable proof safe fix r have 1:"(\n. {..real_of_rat (r + 1 / rat_of_nat (Suc n))}) = {..real_of_rat r}" proof safe fix x assume h:" x \ (\n. {..real_of_rat (r + 1 / rat_of_nat (Suc n))})" show "x \ real_of_rat r" proof(rule ccontr) assume "\ x \ real_of_rat r" then have "0 < x - real_of_rat r" by simp then obtain n where "(1 / (real (Suc n))) < x - real_of_rat r" using nat_approx_posE by blast hence "real_of_rat (r + 1 / (1 + rat_of_nat n)) < x" by (simp add: of_rat_add of_rat_divide) with h show False using linorder_not_le by fastforce qed next fix x n assume "x \ real_of_rat r" then show " x \ real_of_rat (r + 1 / rat_of_nat (Suc n))" by (metis le_add_same_cancel1 of_nat_0_le_iff of_rat_less_eq order_trans zero_le_divide_1_iff) qed have "AE x in r.\x. (\n. r.\' x {..real_of_rat (r + 1 / rat_of_nat (Suc n))}) \ r.\' x {..real_of_rat r}" unfolding 1[symmetric] by(rule r.kernel_lim_decseq) (auto simp: decseq_Suc_iff of_rat_less_eq frac_le) from AE_conjI[OF r.kernel_le_infty[of "{..real_of_rat r}",simplified] this] show "AE x in r.\x. (\n. \ x (r + 1 / (rat_of_nat (Suc n)))) \ \ x r" unfolding \_def by eventually_elim (rule tendsto_enn2real, auto) qed have as3: "AE x in r.\x. (\ x \ 0) at_bot" proof - have 0: "range (\n. {..- real n}) \ sets borel" "decseq (\n. {..- real n})" by(auto simp: decseq_def) show ?thesis proof(safe intro!: AE_I2[THEN AE_mp[OF AE_conjI[OF r.kernel_empty0_AE AE_conjI[OF r.kernel_lim_decseq[OF 0] as1]]]]) fix x assume h: "r.\' x {} = 0" " (\n. r.\' x {..- real n}) \ r.\' x (\n. {..- real n})" "\r s. r \ s \ \ x r \ \ x s" have [simp]: "(\n. {..- real n}) = {}" by auto (meson le_minus_iff linorder_not_less reals_Archimedean2) show "(\ x \ 0) at_bot" proof(rule decreasing_tendsto) fix r :: real assume "0 < r" with h(2) eventually_sequentially obtain N where N:"\n. n \ N \ r.\' x {..- real n} < r" by(fastforce simp: order_tendsto_iff h(1)) show "\\<^sub>F q in at_bot. \ x q < r" unfolding eventually_at_bot_linorder proof(safe intro!: exI[where x="- rat_of_nat N"]) fix q assume "q \ - rat_of_nat N" with h(3) have "\ x q \ \ x (- rat_of_nat N)" by simp also have "... < r" by(auto simp: \_def) (metis N[OF order_refl] \0 < r\ enn2real_less_iff enn2real_top of_rat_minus of_rat_of_nat_eq top.not_eq_extremum) finally show "\ x q < r" . qed qed(simp add: \_def) qed qed have as4: "AE x in r.\x. (\ x \ 1) at_top" proof - have 0: "range (\n. {..real n}) \ sets borel" "incseq (\n. {..real n})" by(auto simp: incseq_def) have [simp]: "(\n. {..real n}) = UNIV" by (auto simp: real_arch_simple) have 1: "AE x in r.\x. \n. r.\' x {..real n} \ 1" "AE x in r.\x. \q. r.\' x {..real_of_rat q} \ 1" by(auto simp: AE_all_countable intro!: r.kernel_le1_AE) show ?thesis proof(safe intro!: AE_I2[THEN AE_mp[OF AE_conjI[OF AE_conjI[OF 1] AE_conjI[OF r.kernel_Y1_AE AE_conjI[OF r.kernel_lim_incseq[OF 0] as1]]],simplified]]) fix x assume h: "\q. r.\' x {..real_of_rat q} \ 1" "\n. r.\' x {..real n} \ 1" "(\n. r.\' x {..real n}) \ r.\' x UNIV" "\r s. r \ s \ \ x r \ \ x s" "r.\' x UNIV = 1" then have h3: "(\n. r.\' x {..real n}) \ 1" by auto show "(\ x \ 1) at_top" proof(rule increasing_tendsto) fix r :: real assume "r < 1" with h3 eventually_sequentially obtain N where N: "\n. n \ N \ r < r.\' x {..real n}" by(fastforce simp: order_tendsto_iff) show "\\<^sub>F n in at_top. r < \ x n" unfolding eventually_at_top_linorder proof(safe intro!: exI[where x="rat_of_nat N"]) fix q assume "rat_of_nat N \ q" have "r < \ x (rat_of_nat N)" by(auto simp: \_def) (metis N[OF order_refl] h(2) enn2real_1 enn2real_ennreal enn2real_positive_iff ennreal_cases ennreal_leI linorder_not_less zero_less_one) also have "... \ \ x q" using h(4) \rat_of_nat N \ q\ by simp finally show "r < \ x q" . qed qed(use h(1) enn2real_leI \_def in auto) qed qed from AE_E3[OF AE_conjI[OF as1 AE_conjI[OF as2 AE_conjI[OF as3 as4]]],simplified space_marginal_measure] obtain N where N: "N \ null_sets r.\x" "\x r s. x \ space X - N \ r \ s \ \ x r \ \ x s" "\x r. x \ space X - N \ (\n. \ x (r + 1 / rat_of_nat (Suc n))) \ \ x r" "\x. x \ space X - N \ (\ x \ 0) at_bot" "\x. x \ space X - N \ (\ x \ 1) at_top" by metis define F where "F \ (\x y. indicat_real (space X - N) x * Inf {\ x r |r. y \ real_of_rat r} + indicat_real N x * indicat_real {0..} y)" have [simp]: "{\ x r |r. y \ real_of_rat r} \ {}" for x y by auto (meson gt_ex less_eq_real_def of_rat_dense) have [simp]: "bdd_below {\ x r |r. y \ real_of_rat r}" if "x \ space X - N" for x y proof - obtain r' where "real_of_rat r' \ y" by (metis less_eq_real_def lt_ex of_rat_dense) from order_trans[OF this] of_rat_less_eq show ?thesis by(auto intro!: bdd_belowI[of _ "\ x r'"] N(2)[OF that]) qed have Feq: "F x (real_of_rat r) = \ x r" if "x \ space X - N" for x r using that N(2)[OF that] by(auto intro!: cInf_eq_minimum simp: of_rat_less_eq F_def) have Fmono: "mono (F x)" if "x \ space X" for x by(auto simp: F_def mono_def indicator_def intro!: cInf_superset_mono) (meson gt_ex less_eq_real_def of_rat_dense) have F1: "(F x \ 0) at_bot" if "x \ space X" for x proof(cases "x \ N") case True with that show ?thesis by(auto simp: F_def tendsto_iff eventually_at_bot_dense indicator_def intro!: exI[where x=0]) next case False with qlim_eq_lim_mono_at_bot[OF Fmono[OF that] N(4)] Feq that show ?thesis by auto qed have F2: "(F x \ 1) at_top" if "x \ space X" for x proof(cases "x \ N") case True with that show ?thesis by(auto simp: F_def tendsto_iff eventually_at_top_dense indicator_def intro!: exI[where x=0]) next case False with qlim_eq_lim_mono_at_top[OF Fmono[OF that] N(5)] Feq that show ?thesis by auto qed have F3: "continuous (at_right a) (F x)" if "x \ space X" for x a proof(cases "x \ N") case x:True { fix e :: real assume e:"0 < e" consider "a \ 0" | "a < 0" by fastforce then have "\d>0. indicat_real {0..} (a + d) - indicat_real {0..} a < e" proof cases case 1 with e show ?thesis by(auto intro!: exI[where x=1]) next case 2 then obtain b where "b > 0" "a + b < 0" by (metis add_less_same_cancel2 of_rat_dense real_add_less_0_iff) with e 2 show ?thesis by(auto intro!: exI[where x=b]) qed } with x show ?thesis unfolding continuous_at_right_real_increasing[of "F x",OF monoD[OF Fmono[OF that]],simplified] by(auto simp: F_def) next case x:False { fix e :: real assume e: "e > 0" have "\k. a \ real_of_rat k \ \ {\ x r |r. a \ real_of_rat r} + e / 2 \ \ x k" proof(rule ccontr) assume "\k. a \ real_of_rat k \ \ x k \ \ {\ x r |r. a \ real_of_rat r} + e / 2" then have cont: "\k. a \ real_of_rat k \ \ x k - e / 2 > \ {\ x r |r. a \ real_of_rat r}" by auto hence "a \ real_of_rat k \ \r. a \ real_of_rat r \ \ x r < \ x k - e / 2" for k using cont \x \ space X\ x cInf_less_iff [of "{\ x r |r. a \ real_of_rat r}" "\ x k - e / 2"] by auto then obtain r where r:"\k. a \ real_of_rat k \ a \ real_of_rat (r k)" "\k. a \ real_of_rat k \ \ x (r k) < \ x k - e / 2" by metis obtain k where k:"a \ real_of_rat k" by (meson gt_ex less_eq_real_def of_rat_dense) define f where "f \ rec_nat k (\n fn. r fn)" have f_simp: "f 0 = k" "f (Suc n) = r (f n)" for n by(auto simp: f_def) have f1: "a \ real_of_rat (f n)" for n using r(1) k by(induction n) (auto simp: f_simp) have f2: "n \ 1 \ \ x (f n) < \ x k - real n * e / 2" for n proof(induction n) case ih:(Suc n) consider "n = 0" | "n \ 1" by fastforce then show ?case proof cases case 1 with r k show ?thesis by(simp add: f_simp) next case 2 show ?thesis using less_trans[OF r(2)[OF f1[of n]] diff_strict_right_mono[OF ih(1)[OF 2],of "e / 2"]] by(auto simp: f_simp ring_distribs(2) add_divide_distrib) qed qed simp have "\ bdd_below {\ x r |r. a \ real_of_rat r}" unfolding bdd_below_def proof safe fix M obtain n where "\ x k - M < real n * e / 2" using f2 e reals_Archimedean3 by fastforce then have "\ x k - M < real (Suc n) * e / 2" using divide_strict_right_mono pos_divide_less_eq e by fastforce thus "Ball {\ x r |r. a \ real_of_rat r} ((\) M) \ False" using f2[of "Suc n"] f1[of "Suc n"] by(auto intro!: exI[where x="\ x (f (Suc n))"]) qed with that x show False by simp qed then obtain k where k: "a \ real_of_rat k" "\ {\ x r |r. a \ real_of_rat r} + e / 2 \ \ x k" by auto obtain no where no:"\n. n\no \ (\ x (k + 1 / rat_of_nat (Suc n))) - (\ x k) < e / 2" using \x \ space X\ x metric_LIMSEQ_D[OF N(3)[of x k],of "e/2"] e N(2)[of x k "k + 1 / rat_of_nat (Suc _)"] by(auto simp: dist_real_def) have "\d>0. \ {\ x r |r. a + d \ real_of_rat r} - \ {\ x r |r. a \ real_of_rat r} < e" proof(safe intro!: exI[where x="real_of_rat (1 / rat_of_nat (Suc no))"]) have "\ x (k + 1 / rat_of_nat (Suc no)) - e < \ x k - e / 2" using no[OF order_refl] by simp also have "... \ \ {\ x r |r. a \ real_of_rat r}" using k by simp finally have "\ x (k + 1 / rat_of_nat (Suc no)) - \ {\ x r |r. a \ real_of_rat r} < e" by simp moreover have "\ {\ x r |r. a + real_of_rat (1 / (1 + rat_of_nat no)) \ real_of_rat r} \ \ x (k + 1 / rat_of_nat (Suc no))" using k that x by(auto intro!: cInf_lower simp: of_rat_add) ultimately show "\ {\ x r |r. a + real_of_rat (1 / (rat_of_nat (Suc no))) \ real_of_rat r} - \ {\ x r |r. a \ real_of_rat r} < e" by simp qed simp } with that x show ?thesis unfolding continuous_at_right_real_increasing[of "F x",OF monoD[OF Fmono[OF that]],simplified] by(auto simp: F_def) qed define \ where "\ \ (\x. interval_measure (F x))" have \: "\x. x \ space X \ \ x UNIV = 1" "\x r. x \ space X \ \ x {..r} = ennreal (F x r)" and[simp]: "\x. sets (\ x) = sets borel" "\x. space (\ x) = UNIV" using emeasure_interval_measure_Iic[OF _ F3 F1] interval_measure_UNIV[OF _ F3 F1 F2] Fmono by(auto simp: mono_def \_def) interpret \: prob_kernel X borel \ unfolding prob_kernel_def' proof(rule measurable_prob_algebra_generated[OF _ atMostq_Int_stable,of _ UNIV]) show "\a. a \ space X \ prob_space (\ a)" by(auto intro!: prob_spaceI \(1)) next fix A assume "A \ {{..r} |r::real. r \ \}" then obtain r where r: "A = {..real_of_rat r}" using Rats_cases by blast have "(\x. ennreal (indicat_real (space X - N) x * \ x r + indicat_real N x * indicat_real {0..} (real_of_rat r))) \ borel_measurable X" proof - have "N \ sets X" using null_setsD2[OF N(1)] by(auto simp: sets_marginal_measure) thus ?thesis by(auto simp: \_def) qed moreover have "indicat_real (space X - N) x * \ x r + indicat_real N x * indicat_real {0..} (real_of_rat r) = emeasure (\ x) A" if "x \ space X" for x using Feq[of x r] \(2)[OF that,of "real_of_rat r"] by(cases "x \ N") (auto simp: r indicator_def F_def) ultimately show " (\x. emeasure (\ x) A) \ borel_measurable X" using measurable_cong[of _ "\x. emeasure (\ x) A" "\x. ennreal (indicat_real (space X - N) x * \ x r + indicat_real N x * indicat_real {0..} (real_of_rat r))"] by simp qed(auto simp: rborel_eq_atMostq) have \_AE:"AE x in r.\x. \ x {..real_of_rat r} = r.\' x {..real_of_rat r}" for r proof - have "AE x in r.\x. \ x {..real_of_rat r} = ennreal (F x (real_of_rat r))" by(auto simp: space_marginal_measure \(2)) moreover have "AE x in r.\x. ennreal (F x (real_of_rat r))= ennreal (\ x r)" using Feq[of _ r] by(auto simp add: space_marginal_measure intro!: AE_I''[OF N(1)]) moreover have "AE x in r.\x. ennreal (\ x r) = ennreal (enn2real (r.\' x {..real_of_rat r}))" by(simp add: \_def) moreover have "AE x in r.\x. ennreal (enn2real (r.\' x {..real_of_rat r})) = r.\' x {..real_of_rat r}" using r.kernel_le_infty[of "{..real_of_rat r}",simplified] by(auto simp: ennreal_enn2real_if) ultimately show ?thesis by auto qed have \_dis: "\.disintegration \ r.\x" proof - interpret D: Dynkin_system UNIV "{B \ sets borel. \A\ sets X. \ (A \ B) = (\\<^sup>+x\A. (\ x) B\r.\x)}" proof { fix A assume h:"A\ sets X" then have "\ (A \ UNIV) = (\\<^sup>+x\A. 1 \r.\x)" using emeasure_marginal_measure[OF nu_sets' h] sets_marginal_measure[of X borel \ "space borel"] by auto also have "... = (\\<^sup>+x\A. (\ x) UNIV\r.\x)" by(auto intro!: nn_integral_cong simp: \ space_marginal_measure) finally have "\ (A \ UNIV) = (\\<^sup>+x\A. emeasure (\ x) UNIV\r.\x)" . } - thus "UNIV \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = \\<^sup>+x\A. emeasure (\ x) B\r.\x}" + thus "UNIV \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B\r.\x)}" by auto hence univ:"\A. A \ sets X \ \ (A \ UNIV) = (\\<^sup>+x\A. emeasure (\ x) UNIV\r.\x)" by auto - show "\B. B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = \\<^sup>+x\A. emeasure (\ x) B\r.\x} \ UNIV - B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = \\<^sup>+x\A. emeasure (\ x) B\r.\x}" + show "\B. B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B\r.\x)} + \ UNIV - B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B\r.\x)}" proof(rule r.\x.sigma_finite_disjoint) fix B and J :: "nat \ _" assume "B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B \r.\x)}" "range J \ sets r.\x" "\ (range J) = space r.\x" "(\i. emeasure r.\x (J i) \ \)" "disjoint_family J" then have B: "B \ sets borel" " \A\sets X. \ (A \ B) = (\\<^sup>+x\A. (\ x) B\r.\x)" and J: "range J \ sets X" "\ (range J) = space X" "\i. emeasure r.\x (J i) \ \" "disjoint_family J" by (auto simp: sets_marginal_measure space_marginal_measure) { fix A assume A: "A \ sets X" have "\ (A \ (UNIV - B)) = (\\<^sup>+x\A. (\ x) (UNIV - B)\r.\x)" (is "?lhs = ?rhs") proof - have AJi1: "disjoint_family (\i. (A \ J i) \ (UNIV - B))" using B(1) J(4) by(fastforce simp: disjoint_family_on_def) have AJi2[simp]: "(\i. ((A \ J i) \ (UNIV - B))) = A \ (UNIV - B)" using J(2) sets.sets_into_space[OF A] by blast have AJi3: "(range (\i. (A \ J i) \ (UNIV - B))) \ sets \" using B(1) J(1) A by(auto simp: nu_sets') have "?lhs = (\i. \ ((A \ J i) \ (UNIV - B)))" by(simp add: suminf_emeasure[OF AJi3 AJi1]) also have "... = (\i. (\\<^sup>+x\ A \ J i. (\ x) (UNIV - B) \r.\x))" proof(safe intro!: suminf_cong) fix n have Jn: "J n \ sets X" using J by auto have fin: "\ ((A \ J n) \ C) \ \" for C proof(cases "(A \ J n) \ C \ sets \") case True then have "\ ((A \ J n) \ C) \ \ ((A \ J n) \ UNIV)" using Jn nu_sets' A by(intro emeasure_mono) auto also have "\ ((A \ J n) \ UNIV) \ \ (J n \ UNIV)" using Jn nu_sets' by(intro emeasure_mono) auto also have "... = r.\x (J n)" using emeasure_marginal_measure[OF nu_sets' Jn] by simp finally show ?thesis by (metis J(3)[of n] infinity_ennreal_def neq_top_trans) qed(simp add: emeasure_notin_sets) show "\ ((A \ J n) \ (UNIV - B)) = (\\<^sup>+x\ A \ J n. (\ x) (UNIV - B) \r.\x)" (is "?lhs = ?rhs") proof - have "?lhs = \ ((A \ J n) \ UNIV) - \ ((A \ J n) \ B)" proof - have [simp]: "?lhs + \ ((A \ J n) \ B) = \ ((A \ J n) \ UNIV)" proof - have [simp]:"((A \ J n) \ (UNIV - B)) \ ((A \ J n) \ B) = ((A \ J n) \ UNIV)" by blast show ?thesis using B(1) A Jn nu_sets' by(intro plus_emeasure[of "(A \ J n) \ (UNIV - B)" _ "(A \ J n) \ B",simplified]) auto qed have "?lhs = ?lhs + \ ((A \ J n) \ B) - \ ((A \ J n) \ B)" by(simp only: ennreal_add_diff_cancel[OF fin[of B]]) also have "... = \ ((A \ J n) \ UNIV) - \ ((A \ J n) \ B)" by simp finally show ?thesis . qed also have "... = (\\<^sup>+x\ A \ J n. (\ x) UNIV \r.\x) - (\\<^sup>+x\ A \ J n. (\ x) B \r.\x)" using B(2) A Jn univ by auto also have "... = (\\<^sup>+x. ((\ x) UNIV * indicator (A \ J n) x - (\ x) B * indicator (A \ J n) x) \r.\x)" proof(rule nn_integral_diff[symmetric]) show "(\x. (\ x) UNIV * indicator (A \ J n) x) \ borel_measurable r.\x" "(\x. (\ x) B * indicator (A \ J n) x) \ borel_measurable r.\x" using sets_marginal_measure[of X borel \ "space borel"] \.emeasure_measurable[OF B(1)] \.emeasure_measurable[of UNIV] A Jn by(auto simp del: space_borel) next - show "\\<^sup>+x\A \ J n. (\ x) B\r.\x \ \" + show "(\\<^sup>+x\A \ J n. (\ x) B\r.\x) \ \" using B(2) A Jn univ fin[of B] by auto next show "AE x in r.\x. (\ x) B * indicator (A \ J n) x \ (\ x) UNIV * indicator (A \ J n) x" by(standard, auto simp: space_marginal_measure indicator_def intro!: emeasure_mono) qed also have "... = (\\<^sup>+x\ A \ J n. ((\ x) UNIV - (\ x) B) \r.\x)" by(auto intro!: nn_integral_cong simp: indicator_def) also have "... = ?rhs" proof(safe intro!: nn_integral_cong) fix x assume "x \ space r.\x" then have "x \ space X" by(simp add: space_marginal_measure) show "((\ x) UNIV - (\ x) B) * indicator (A \ J n) x = (\ x) (UNIV - B) * indicator (A \ J n) x" by(auto intro!: emeasure_compl[of B "\ x",simplified,symmetric] simp: B \.prob_spaces \x \ space X\ prob_space_imp_subprob_space subprob_space.emeasure_subprob_space_less_top indicator_def) qed finally show ?thesis . qed qed also have "... = (\\<^sup>+x. (\i. (\ x) (UNIV - B) * indicator (A \ J i) x)\r.\x)" using \.emeasure_measurable[of "UNIV - B"] B(1) sets_marginal_measure[of X borel \ "space borel"] A J by(intro nn_integral_suminf[symmetric]) (auto simp del: space_borel) also have "... = (\\<^sup>+x. (\ x) (UNIV - B) * indicator A x * (\i. indicator (A \ J i) x) \r.\x)" by(auto simp: indicator_def intro!: nn_integral_cong) also have "... = (\\<^sup>+x. (\ x) (UNIV - B) * indicator A x * (indicator (\i. A \ J i) x) \r.\x)" proof - have "(\i. indicator (A \ J i) x) = (indicator (\i. A \ J i) x :: ennreal)" for x using J(4) by(intro suminf_indicator) (auto simp: disjoint_family_on_def) thus ?thesis by(auto intro!: nn_integral_cong) qed also have "... = ?rhs" using J(2) by(auto simp: indicator_def space_marginal_measure intro!: nn_integral_cong) finally show ?thesis . qed } thus "UNIV - B \ {B \ sets borel. \A\sets X. emeasure \ (A \ B) = (\\<^sup>+x\A. emeasure (\ x) B \r.\x)}" using B by auto qed next fix J :: "nat \ _" - assume J1: "disjoint_family J" "range J \ {B \ sets borel. \A\sets X. \ (A \ B) = \\<^sup>+x\A. (\ x) B\r.\x}" + assume J1: "disjoint_family J" "range J \ {B \ sets borel. \A\sets X. \ (A \ B) = (\\<^sup>+x\A. (\ x) B\r.\x)}" then have J2: "range J \ sets borel" "\ (range J) \ sets borel" "\n A. A \ sets X \ \ (A \ (J n)) = (\\<^sup>+x\A. (\ x) (J n) \r.\x)" by auto - show "\ (range J) \ {B \ sets borel. \A\sets X. \ (A \ B) = \\<^sup>+x\A. (\ x) B\r.\x}" + show "\ (range J) \ {B \ sets borel. \A\sets X. \ (A \ B) = (\\<^sup>+x\A. (\ x) B\r.\x)}" proof - { fix A assume A:"A \ sets X" have "\ (A \ \ (range J)) = (\\<^sup>+x\A. (\ x) (\ (range J)) \r.\x)" (is "?lhs = ?rhs") proof - have "?lhs = \ (\n. A \ J n)" proof - have "(A \ \ (range J)) = (\n. A \ J n)" by blast thus ?thesis by simp qed also have "... = (\n. \ (A \ J n))" using J1(1) J2(1) A nu_sets' by(fastforce intro!: suminf_emeasure[symmetric] simp: disjoint_family_on_def) also have "... = (\n. (\\<^sup>+x\A. (\ x) (J n) \r.\x))" by(simp add: J2(3)[OF A]) also have "... = (\\<^sup>+x. (\n. (\ x) (J n) * indicator A x) \r.\x)" using \.emeasure_measurable J2(1) A sets_marginal_measure[of X borel \ "space borel"] by(intro nn_integral_suminf[symmetric]) auto also have "... = (\\<^sup>+x\A. (\n. (\ x) (J n)) \r.\x)" by auto also have "... = (\\<^sup>+x\A. (\ x) (\ (range J)) \r.\x)" using J1 J2 by(auto intro!: nn_integral_cong suminf_emeasure simp: space_marginal_measure indicator_def) finally show ?thesis . qed } thus ?thesis using J2(2) by auto qed qed auto have "{ {..r} | r::real. r \ \} \ {B \ sets borel. \A\ sets X. \ (A \ B) = (\\<^sup>+x\A. (\ x) B\r.\x)}" proof - { fix r ::real and A assume h: "r \ \" "A \ sets X" then obtain r' where r':"r = real_of_rat r'" using Rats_cases by blast have "\ (A \ {..r}) = (\\<^sup>+x\A. (\ x) {..r} \r.\x)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+x\A. r.\' x {..r}\r.\x)" using h by(simp add: r.kernel_RN_deriv) also have "... = ?rhs" using \_AE[of r'] by(auto intro!: nn_integral_cong_AE simp: r' simp del: space_borel) finally show ?thesis . qed } thus ?thesis by auto qed from D.Dynkin_subset[OF this] rborel_eq_atMostq[symmetric] show ?thesis by(auto simp: \.disintegration_def sets_marginal_measure nu_sets' sigma_eq_Dynkin[OF _ atMostq_Int_stable,of UNIV,simplified,symmetric] rborel_eq_atMostq_sets simp del: space_borel) qed show ?thesis proof(intro exI conjI strip) fix \'' assume "prob_kernel X (borel :: real measure) \''" interpret \'': prob_kernel X borel \'' by fact assume disi: "\''.disintegration \ r.\x" have eq_atMostr_AE:"AE x in r.\x. \r. \ x {..real_of_rat r} = \'' x {..real_of_rat r}" unfolding AE_all_countable proof safe fix r have "AE x in r.\x. (\'' x) {..real_of_rat r} = r.\' x {..real_of_rat r}" proof(safe intro!: r.\x.RN_deriv_unique[of "\x. \'' x {..real_of_rat r}" "marginal_measure_on X borel \ {..real_of_rat r}",simplified r.\'_def[of _ "{..real_of_rat r}",symmetric]]) show 1:"(\x. emeasure (\'' x) {..real_of_rat r}) \ borel_measurable r.\x" using \''.emeasure_measurable[of "{..real_of_rat r}"] sets_marginal_measure[of X borel \ "space borel"] by simp show "density r.\x (\x. emeasure (\'' x) {..real_of_rat r}) = marginal_measure_on X borel \ {..real_of_rat r}" proof(rule measure_eqI) fix A assume "A \ sets (density r.\x (\x. (\'' x) {..real_of_rat r}))" then have A [measurable]:"A \ sets X" by(simp add: sets_marginal_measure) show "emeasure (density r.\x (\x. emeasure (\'' x) {..real_of_rat r})) A = emeasure (marginal_measure_on X borel \ {..real_of_rat r}) A" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+x\A. (\'' x) {..real_of_rat r} \r.\x)" using emeasure_density[OF 1,of A] A by(simp add: sets_marginal_measure) also have "... = \ (A \ {..real_of_rat r})" using disi A by(auto simp: \''.disintegration_def) also have "... = ?rhs" by(simp add: emeasure_marginal_measure_on[OF nu_sets' _ A]) finally show ?thesis . qed qed(simp add: sets_marginal_measure) qed with \_AE[of r] show "AE x in r.\x. \ x {..real_of_rat r} = \'' x {..real_of_rat r}" by auto qed { fix x assume h:"x \ space r.\x" " \r. (\ x) {..real_of_rat r} = (\'' x) {..real_of_rat r}" then have x: "x \ space X" by(simp add: space_marginal_measure) have "\ x = \'' x" proof(rule measure_eqI_generator_eq[OF atMostq_Int_stable,of UNIV _ _ "\n. {..real n}"]) show "\A. A \ {{..r} |r. r \ \} \ (\ x) A = (\'' x) A" using h(2) Rats_cases by auto next show "(\n. {..real n}) = UNIV" by (simp add: real_arch_simple subsetI subset_antisym) next fix n have "(\ x) {..real n} \ \ x UNIV" by(auto intro!: emeasure_mono) also have "... = 1" by(rule \(1)[OF x]) finally show "(\ x) {..real n} \ \" using linorder_not_le by fastforce next show "range (\n. {..real n}) \ {{..r} |r. r \ \}" using Rats_of_nat by blast qed(auto simp: \.kernel_sets[OF x] \''.kernel_sets[OF x] rborel_eq_atMostq_sets) } then show "AE x in r.\x. \ x = \'' x" using eq_atMostr_AE by fastforce qed(auto simp del: space_borel simp add: \_dis \.prob_kernel_axioms) qed show ?thesis proof - define \' where "\' = distr \ (X \\<^sub>M borel) (\(x,y). (x, to_real y))" have \_distr:"\ = distr \' (X \\<^sub>M Y) (\(x,y). (x, from_real y))" using nu_sets sets_eq_imp_space_eq[OF nu_sets] from_real_to_real by(auto simp: \'_def distr_distr space_pair_measure intro!: distr_id'[symmetric]) have \x_eq:"(marginal_measure X borel \') = \x" using emeasure_marginal_measure[of \' X borel] emeasure_marginal_measure[OF nu_sets] sets_eq_imp_space_eq[OF nu_sets] by(auto intro!: measure_eqI simp: sets_marginal_measure \'_def emeasure_distr map_prod_vimage[of id to_real,simplified map_prod_def id_def] space_pair_measure Times_Int_Times) interpret \' : projection_sigma_finite X borel \' by(auto simp: projection_sigma_finite_def \x_eq \x.sigma_finite_measure_axioms simp del: space_borel,auto simp add: \'_def) obtain \' where \': "prob_kernel X borel \'" "measure_kernel.disintegration X borel \' \' \'.\x" "\\''. prob_kernel X borel \'' \ measure_kernel.disintegration X borel \'' \' \'.\x \ (AE x in \'.\x. \' x = \'' x)" using *[of \' X] \'.nu_sets \'.\x.sigma_finite_measure_axioms by blast interpret \': prob_kernel X borel \' by fact define \ where "\ \ (\x. distr (\' x) Y from_real)" interpret \: prob_kernel X Y \ by(auto simp: prob_kernel_def' \_def) have disi: "\.disintegration \ \x" proof(rule \.disintegrationI) fix A B assume A[measurable]:"A \ sets X" and B[measurable]: "B \ sets Y" have [measurable]: "from_real -` B \ sets borel" by(simp add: measurable_sets_borel[OF _ B]) show "\ (A \ B) = (\\<^sup>+x\A. \ x B\\x)" (is "?lhs = ?rhs") proof - have "?lhs = \' (A \ (from_real -` B))" by(auto simp: \_distr emeasure_distr map_prod_vimage[of id from_real,simplified map_prod_def id_def]) also have "... = (\\<^sup>+x\A. \' x (from_real -` B) \\x)" using \'.disintegrationD[OF \'(2),of A "from_real -` B"] by(auto simp add: \x_eq simp del: space_borel) also have "... = ?rhs" by(auto intro!: nn_integral_cong simp: space_marginal_measure \_def emeasure_distr) finally show ?thesis . qed qed(simp_all add: sets_marginal_measure nu_sets) show ?thesis proof(safe intro!: exI[where x=\]) fix \'' assume h:"prob_kernel X Y \''" "measure_kernel.disintegration X Y \'' \ \x" interpret \'': prob_kernel X Y \'' by fact show "AE x in \x. \ x = \'' x" proof - define \''' where "\''' \ (\x. distr (\'' x) borel to_real)" interpret \''': prob_kernel X borel \''' by(auto simp: prob_kernel_def' \'''_def) have \''_def: "\'' x = distr (\''' x) Y from_real" if "x \ space X" for x using distr_distr[of from_real borel Y to_real "\'' x",simplified measurable_cong_sets[OF \''.kernel_sets[OF that] refl,of borel]] by(auto simp: \'''_def comp_def \''.kernel_sets[OF that] measurable_cong_sets[OF \''.kernel_sets[OF that] \''.kernel_sets[OF that]] sets_eq_imp_space_eq[OF \''.kernel_sets[OF that]] intro!: distr_id'[symmetric]) have \'''_disi: "\'''.disintegration \' \'.\x" proof(rule \'''.disintegrationI) fix A and B :: "real set" assume A[measurable]:"A \ sets X" and B[measurable]:"B \ sets borel" show "\' (A \ B) = (\\<^sup>+x\A. (\''' x) B \\'.\x)" (is "?lhs = ?rhs") proof - have "?lhs = \ (A \ (to_real -` B \ space Y))" by(auto simp: \'_def emeasure_distr map_prod_vimage[of id to_real,simplified map_prod_def id_def] sets_eq_imp_space_eq[OF nu_sets] space_pair_measure Times_Int_Times) also have "... = (\\<^sup>+x\A. (\'' x) (to_real -` B \ space Y) \\x)" using \''.disintegrationD[OF h(2) A,of "to_real -` B \ space Y"] by auto also have "... = ?rhs" by(auto simp: \x_eq[symmetric] space_marginal_measure \'''_def emeasure_distr sets_eq_imp_space_eq[OF \''.kernel_sets] intro!: nn_integral_cong) finally show ?thesis . qed qed(auto simp: \'_def sets_marginal_measure) show ?thesis by(rule AE_mp[OF \'(3)[OF \'''.prob_kernel_axioms \'''_disi,simplified \x_eq]],standard) (auto simp: space_marginal_measure \''_def \_def) qed qed(simp_all add: disi \.prob_kernel_axioms) qed qed end subsection \ Lemma 14.D.12. \ lemma ex_finite_density_measure: fixes A :: "nat \ _" assumes A: "range A \ sets M" "\ (range A) = space M" "\i. emeasure M (A i) \ \" "disjoint_family A" defines "h \ (\x. (\n. (1/2)^(Suc n) * (1 / (1 + M (A n))) * indicator (A n) x))" shows "h \ borel_measurable M" "\x. x \ space M \ 0 < h x" "\x. x \ space M \ h x < 1" "finite_measure (density M h)" proof - have less1:"0 < 1 / (1 + M (A n))" "1 / (1 + M (A n)) \ 1" for n using A(3)[of n] ennreal_zero_less_divide[of 1 "1 + M (A n)"] by (auto intro!: divide_le_posI_ennreal simp: add_pos_nonneg) show [measurable]:"h \ borel_measurable M" using A by(simp add: h_def) { fix x assume x:"x \ space M" then obtain i where i: "x \ A i" using A(2) by auto show "0 < h x" using A(3)[of i] less1[of i] by(auto simp: h_def suminf_pos_iff i ennreal_divide_times ennreal_zero_less_divide power_divide_distrib_ennreal power_less_top_ennreal intro!: exI[where x=i]) have "h x = (\n. (1/2)^(Suc n + 2) * (1 / (1 + M (A (n + 2)))) * indicator (A (n + 2)) x) + (1/2) * (1 / (1 + M (A 0))) * indicator (A 0) x + (1/2)^2 * (1 / (1 + M (A 1))) * indicator (A 1) x" by(auto simp: h_def suminf_split_head suminf_offset[of "\n. (1/2)^(Suc n) * (1 / (1 + M (A n))) * indicator (A n) x" 2] simp del: power_Suc sum_mult_indicator) (auto simp: numeral_2_eq_2) also have "... \ 1/4 + (1/2) * (1 / (1 + M (A 0))) * indicator (A 0) x + (1/2)^2 * (1 / (1 + M (A 1))) * indicator (A 1) x" proof - have "(\n. (1/2)^(Suc n + 2) * (1 / (1 + M (A (n + 2)))) * indicator (A (n + 2)) x) \ (\n. (1/2)^(Suc n + 2))" using less1(2)[of "Suc (Suc _)"] by(intro suminf_le,auto simp: indicator_def) (metis mult.right_neutral mult_left_mono zero_le) also have "... = (\n. ennreal ((1 / 2) ^ (Suc n + 2)))" by(simp only: ennreal_power[of "1/2",symmetric]) (metis divide_ennreal ennreal_1 ennreal_numeral linorder_not_le not_one_less_zero zero_less_numeral) also have "... = ennreal (\n. (1 / 2) ^ (Suc n + 2))" by(rule suminf_ennreal2) auto also have "... = ennreal (1/4)" using nsum_of_r'[of "1/2" "Suc (Suc (Suc 0))" 1] by auto also have "... = 1 / 4" by (metis ennreal_divide_numeral ennreal_numeral numeral_One zero_less_one_class.zero_le_one) finally show ?thesis by simp qed also have "... < 1" (is "?lhs < _") proof(cases "x \ A 0") case True then have "x \ A 1" using A(4) by (auto simp: disjoint_family_on_def) hence "?lhs = 1 / 4 + 1 / 2 * (1 / (1 + emeasure M (A 0)))" by(simp add: True) also have "... \ 1 / 4 + 1 / 2" using less1(2)[of 0] by (simp add: divide_right_mono_ennreal ennreal_divide_times) also have "... = 1 / 4 + 2 / 4" using divide_mult_eq[of 2 1 2] by simp also have "... = 3 / 4" by(simp add: add_divide_distrib_ennreal[symmetric]) also have "... < 1" by(simp add: divide_less_ennreal) finally show ?thesis . next case False then have "?lhs = 1 / 4 + (1 / 2)\<^sup>2 * (1 / (1 + emeasure M (A 1))) * indicator (A 1) x" by simp also have "... \ 1 / 4 + (1 / 2)\<^sup>2" by (metis less1(2)[of 1] add_left_mono indicator_eq_0_iff indicator_eq_1_iff mult.right_neutral mult_eq_0_iff mult_left_mono zero_le) also have "... = 2 / 4" by(simp add: power_divide_distrib_ennreal add_divide_distrib_ennreal[symmetric]) also have "... < 1" by(simp add: divide_less_ennreal) finally show ?thesis . qed finally show "h x < 1" . } show "finite_measure (density M h)" proof show "emeasure (density M h) (space (density M h)) \ \" proof - have "integral\<^sup>N M h \ \" (is "?lhs \ _") proof - have "?lhs = (\n. (\\<^sup>+ x\A n. ((1/2)^(Suc n) * (1 / (1 + M (A n)))) \M))" using A by(simp add: h_def nn_integral_suminf) also have "... = (\n. (1/2)^(Suc n) * (1 / (1 + M (A n))) * M (A n))" by(rule suminf_cong,rule nn_integral_cmult_indicator) (use A in auto) also have "... = (\n. (1/2)^(Suc n) * ((1 / (1 + M (A n))) * M (A n)))" by (simp add: mult.assoc) also have "... \ (\n. (1/2)^(Suc n))" proof - have "(1 / (1 + M (A n))) * M (A n) \ 1" for n using A(3)[of n] by (simp add: add_pos_nonneg divide_le_posI_ennreal ennreal_divide_times) thus ?thesis by(intro suminf_le) (metis mult.right_neutral mult_left_mono zero_le,auto) qed also have "... = (\n. ennreal ((1/2)^(Suc n)))" by(simp only: ennreal_power[of "1/2",symmetric]) (metis divide_ennreal ennreal_1 ennreal_numeral linorder_not_le not_one_less_zero zero_less_numeral) also have "... = ennreal (\n. (1/2)^(Suc n))" by(rule suminf_ennreal2) auto also have "... = 1" using nsum_of_r'[of "1/2" 1 1] by auto finally show ?thesis using nle_le by fastforce qed thus ?thesis by(simp add: emeasure_density) qed qed qed lemma(in sigma_finite_measure) finite_density_measure: obtains h where "h \ borel_measurable M" "\x. x \ space M \ 0 < h x" "\x. x \ space M \ h x < 1" "finite_measure (density M h)" by (metis (no_types, lifting) sigma_finite_disjoint ex_finite_density_measure) subsection \ Lemma 14.D.13. \ lemma (in measure_kernel) assumes "disintegration \ \" defines "\x \ marginal_measure X Y \" shows disintegration_absolutely_continuous: "absolutely_continuous \ \x" and disintegration_density: "\x = density \ (\x. \ x (space Y))" and disintegration_absolutely_continuous_iff: "absolutely_continuous \x \ \ (AE x in \. \ x (space Y) > 0)" proof - note sets_eq[measurable_cong] = disintegration_sets_eq[OF assms(1)] note [measurable] = emeasure_measurable[OF sets.top] have \x_eq: "\x A = (\\<^sup>+x\A. (\ x (space Y)) \\)" if A:"A \ sets X" for A by(simp add: disintegrationD[OF assms(1) A sets.top] emeasure_marginal_measure[OF sets_eq(1) A] \x_def) thus 1:"\x = density \ (\x. \ x (space Y))" by(auto intro!: measure_eqI simp: sets_marginal_measure \x_def sets_eq emeasure_density) hence sets_\x:"sets \x = sets X" using sets_eq by simp show "absolutely_continuous \ \x" unfolding absolutely_continuous_def proof safe fix A assume A: "A \ null_sets \" have "0 = (\\<^sup>+x\A. (\ x (space Y)) \\)" by(simp add: A nn_integral_null_set) also have "... = \x A" using A \x_eq[of A,simplified sets_eq(2)[symmetric]] by auto finally show "A \ null_sets \x" using A by(auto simp: null_sets_def \x_def sets_marginal_measure sets_eq) qed show "absolutely_continuous \x \ \ (AE x in \. \ x (space Y) > 0)" proof assume h:"absolutely_continuous \x \" define N where "N = {x \ space \. (\ x) (space Y) = 0}" have "N \ null_sets \" proof - have "\x N = (\\<^sup>+x\N. (\ x (space Y)) \\)" using \x_eq[of N] by(simp add: N_def sets_eq_imp_space_eq[OF sets_eq(2)]) also have "... = (\\<^sup>+x\N. 0 \\)" by(rule nn_integral_cong) (auto simp: N_def indicator_def) also have "... = 0" by simp finally have "N \ null_sets \x" by(auto simp: null_sets_def 1 N_def) thus ?thesis using h by(auto simp: absolutely_continuous_def) qed then show "AE x in \. 0 < (\ x) (space Y)" by(auto intro!: AE_I'[OF _ subset_refl] simp: N_def) next assume "AE x in \. 0 < (\ x) (space Y)" then show "absolutely_continuous \x \" using \x_eq by(auto simp: absolutely_continuous_def intro!: null_if_pos_func_has_zero_nn_int[where f="\x. emeasure (\ x) (space Y)"]) (auto simp: null_sets_def sets_\x) qed qed subsection \ Theorem 14.D.14. \ locale sigma_finite_measure_on_pair_standard = sigma_finite_measure_on_pair + standard_borel_ne Y sublocale projection_sigma_finite_standard \ sigma_finite_measure_on_pair_standard by (simp add: sigma_finite_measure_on_pair_axioms sigma_finite_measure_on_pair_standard_def standard_borel_ne_axioms) context sigma_finite_measure_on_pair_standard begin lemma measure_disintegration_extension: "\\ \. finite_measure \ \ measure_kernel X Y \ \ measure_kernel.disintegration X Y \ \ \ \ (\x\space X. sigma_finite_measure (\ x)) \ (\x\space X. \ x (space Y) > 0) \ \ ~\<^sub>M \x" (is "?goal") proof(rule sigma_finite_measure.sigma_finite_disjoint[OF sigma_finite]) fix A :: "nat \ _" assume A:"range A \ sets \" "\ (range A) = space \" "\i. emeasure \ (A i) \ \" "disjoint_family A" define h where "h \ (\x. \n. (1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))) * indicator (A n) x)" have h: "h \ borel_measurable \" "\x y. x \ space X \ y \ space Y \ 0 < h (x,y)" "\x y. x \ space X \ y \ space Y \ h (x,y) < 1" "finite_measure (density \ h)" using ex_finite_density_measure[OF A] by(auto simp: sets_eq_imp_space_eq[OF nu_sets] h_def space_pair_measure) interpret psfs_\x: finite_measure "marginal_measure X Y (density \ h)" by(rule finite_measure_marginal_measure_finite[OF h(4),simplified,OF nu_sets]) interpret psfs: projection_sigma_finite_standard X Y "density \ h" by(auto simp: projection_sigma_finite_standard_def projection_sigma_finite_def standard_borel_ne_axioms nu_sets finite_measure.sigma_finite_measure[OF finite_measure_marginal_measure_finite[OF h(4),simplified,OF nu_sets]]) from psfs.measure_disintegration obtain \' where \': "prob_kernel X Y \'" "measure_kernel.disintegration X Y \' (density \ h) psfs.\x" by auto interpret pk: prob_kernel X Y \' by fact define \ where "\ \ (\x. density (\' x) (\y. 1 / h (x,y)))" have \B: "\ x B = (\\<^sup>+y\B. (1 / h (x, y))\\' x)" if "x \ space X" and [measurable]:"B \ sets Y" for x B using nu_sets pk.kernel_sets[OF that(1)] that h(1) by(auto simp: \_def emeasure_density) interpret mk: measure_kernel X Y \ proof fix B assume [measurable]:"B \ sets Y" have 1:"(\x. \\<^sup>+y\B. (1 / h (x, y))\\' x) \ borel_measurable X" using h(1) nu_sets by(auto intro!: pk.nn_integral_measurable_f'[of "\z. (1 / h z) * indicator B (snd z)",simplified]) show "(\x. (\ x) B) \ borel_measurable X" by(rule measurable_cong[THEN iffD1,OF _ 1],simp add: \B) qed(simp_all add: \_def pk.kernel_sets space_ne) have disi: "mk.disintegration \ psfs.\x" proof(rule mk.disintegrationI) fix A B assume A[measurable]:"A \ sets X" and B[measurable]:"B \ sets Y" show "\ (A \ B) = (\\<^sup>+x\A. (\ x) B\psfs.\x)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+z\A \ B. 1 \\)" by auto also have "... = (\\<^sup>+z\A \ B. (1 / h z * h z) \\)" proof - have 1: "a * (1 / a) = 1" if "0 < a" "a < 1" for a :: ennreal proof - have "a * (1 / a) = ennreal (enn2real a * 1 / (enn2real a))" by (simp add: divide_eq_1_ennreal enn2real_eq_0_iff ennreal_times_divide) also have "... = ennreal 1" using enn2real_eq_0_iff that by fastforce finally show ?thesis using ennreal_1 by simp qed show ?thesis by(rule nn_integral_cong,auto simp add: sets_eq_imp_space_eq[OF nu_sets] space_pair_measure ennreal_divide_times indicator_def 1[OF h(2,3)]) qed also have "... = (\\<^sup>+z. h z * ((1 / h z) * indicator (A \ B) z) \\)" by(auto intro!: nn_integral_cong simp: indicator_def mult.commute) also have "... = (\\<^sup>+z\A \ B. (1 / h z) \(density \ h))" using h(1) by(simp add: nn_integral_density) also have "... = (\\<^sup>+ x. \\<^sup>+ y. (1 / h (x,y) * indicator (A \ B) (x,y)) \\' x \psfs.\x)" using h(1) by(simp add: pk.nn_integral_fst_finite'[OF _ \'(2) psfs_\x.finite_measure_axioms]) also have "... = (\\<^sup>+ x\A. (\\<^sup>+ y\B. (1 / h (x,y)) \\' x) \psfs.\x)" by(auto intro!: nn_integral_cong simp: indicator_def) also have "... = ?rhs" by(auto intro!: nn_integral_cong simp: \B[OF _ B] space_marginal_measure) finally show ?thesis . qed qed(simp_all add: nu_sets sets_marginal_measure) have geq0: "0 < (\ x) (space Y)" if "x \ space X" for x proof - have "0 = (\\<^sup>+ y. 0 \\' x)" by simp also have "... < (\\<^sup>+ y. (1 / h (x,y)) \\' x)" proof(rule nn_integral_less) show "\ (AE y in \' x. 1 / h (x, y) \ 0)" proof assume "AE y in \' x. 1 / h (x, y) \ 0" moreover have "h (x,y) \ \" if "y \ space (\' x)" for y using h(3)[OF \x \ space X\ that[simplified sets_eq_imp_space_eq[OF pk.kernel_sets[OF \x \ space X\]]]] top.not_eq_extremum by fastforce ultimately show False using prob_space.AE_False[OF pk.prob_spaces[OF that]] by simp qed qed(use h(1) pk.kernel_sets[OF that] that in auto) also have "... = (\ x) (space Y)" by(simp add: \B[OF that sets.top]) (simp add: sets_eq_imp_space_eq[OF pk.kernel_sets[OF that],symmetric]) finally show ?thesis . qed show ?goal proof(safe intro!: exI[where x=psfs.\x] exI[where x=\] disi) show "absolutely_continuous \x psfs.\x" unfolding mk.disintegration_absolutely_continuous_iff[OF disi] by standard (simp add: space_marginal_measure geq0) next fix x assume x:"x \ space X" define C where "C \ range (\n. Pair x -` (A n) \ space Y)" have 1:"countable C" "C \ sets Y" using A(1,2) x by (auto simp: nu_sets sets_eq_imp_space_eq[OF nu_sets] space_pair_measure C_def) have 2: "\ C = space Y" using A(1,2) by(auto simp: sets_eq_imp_space_eq[OF nu_sets] space_pair_measure C_def) (use x in auto) show "sigma_finite_measure (\ x)" unfolding sigma_finite_measure_def proof(safe intro!: exI[where x=C]) fix c assume "c \ C" "(\ x) c = \" then obtain n where c:"c = Pair x -` (A n) \ space Y" by(auto simp: C_def) have "(\ x) c = (\\<^sup>+y\c. (1 / h (x, y))\\' x)" using \B[OF x,of c] 1 \c \ C\ by auto also have "... = (\\<^sup>+y\Pair x -` (A n). (1 / h (x, y))\\' x)" by(auto intro!: nn_integral_cong simp: c indicator_def sets_eq_imp_space_eq[OF pk.kernel_sets[OF x]]) also have "... = (\\<^sup>+y\Pair x -` (A n). (1 / ((1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))))) \\' x)" proof - { fix y assume xy:"(x, y) \ A n" have "1 / h (x, y) = 1 / ((1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))))" proof - have "h (x, y) = (1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n)))" (is "?lhs = ?rhs") proof - have "?lhs = (\m. (1 / 2) ^ Suc m * (1 / (1 + emeasure \ (A m))) * indicator (A m) (x,y))" by(simp add: h_def) also have "... = (\m. if m = n then (1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))) else 0)" using xy A(4) by(fastforce intro!: suminf_cong simp: disjoint_family_on_def indicator_def) also have "... = (\j. if j + Suc n = n then (1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))) else 0) + (\j (A n))) else 0)" by(auto simp: suminf_offset[of "\m. if m = n then (1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))) else 0" "Suc n"] simp del: power_Suc) also have "... = ?rhs" by simp finally show ?thesis . qed thus ?thesis by simp qed } thus ?thesis by(intro nn_integral_cong) (auto simp: sets_eq_imp_space_eq[OF pk.kernel_sets[OF x]] indicator_def simp del: power_Suc) qed also have "... \ (\\<^sup>+y. (1 / ((1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n))))) \\' x)" by(rule nn_integral_mono) (auto simp: indicator_def) also have "... = (1 / ((1 / 2) ^ Suc n * (1 / (1 + emeasure \ (A n)))))" by(simp add: prob_space.emeasure_space_1[OF pk.prob_spaces[OF x]]) also have "... < \" by (metis A(3) ennreal_add_eq_top ennreal_divide_eq_0_iff ennreal_divide_eq_top_iff ennreal_top_neq_one infinity_ennreal_def mult_eq_0_iff power_eq_0_iff top.not_eq_extremum top_neq_numeral) finally show False using \(\ x) c = \\ by simp qed(insert 1 2, auto simp: mk.kernel_sets[OF x] sets_eq_imp_space_eq[OF mk.kernel_sets[OF x]]) qed(auto simp: psfs_\x.finite_measure_axioms geq0 mk.measure_kernel_axioms mk.disintegration_absolutely_continuous[OF disi]) qed end (* TODO: \AE x in \. \B \ sets Y. ... = ...\ for a standard Borel Y. *) lemma(in sigma_finite_measure_on_pair) measure_disintegration_extension_AE_unique: assumes "sigma_finite_measure \" "sigma_finite_measure \'" "measure_kernel X Y \" "measure_kernel X Y \'" "measure_kernel.disintegration X Y \ \ \" "measure_kernel.disintegration X Y \' \ \'" and "absolutely_continuous \ \'" "B \ sets Y" shows "AE x in \. \' x B * RN_deriv \ \' x = \ x B" proof - interpret s1: sigma_finite_measure \ by fact interpret s2: sigma_finite_measure \' by fact interpret mk1: measure_kernel X Y \ by fact interpret mk2: measure_kernel X Y \' by fact have sets[measurable_cong]:"sets \ = sets X" "sets \' = sets X" using assms(5,6) by(auto dest: mk1.disintegration_sets_eq mk2.disintegration_sets_eq) have 1:"AE x in \. \ x B = RN_deriv \ (marginal_measure_on X Y \ B) x" using sets mk1.emeasure_measurable[OF assms(8)] mk1.disintegrationD[OF assms(5) _ assms(8)] by(auto intro!: measure_eqI s1.RN_deriv_unique simp: emeasure_density emeasure_marginal_measure_on[OF nu_sets assms(8)] sets sets_marginal_measure) have 2:"AE x in \. \' x B * RN_deriv \ \' x = RN_deriv \ (marginal_measure_on X Y \ B) x" proof - { fix A assume A: "A \ sets X" have "(\\<^sup>+x\A. ((\' x) B * RN_deriv \ \' x) \\) = (\\<^sup>+x. RN_deriv \ \' x * (\' x B * indicator A x)\\)" by(auto intro!: nn_integral_cong simp: indicator_def mult.commute) also have "... = (\\<^sup>+x\A. \' x B \\')" using mk2.emeasure_measurable[OF assms(8)] sets A by(auto intro!: s1.RN_deriv_nn_integral[OF assms(7),symmetric]) also have "... = \ (A \ B)" by(simp add: mk2.disintegrationD[OF assms(6) A assms(8)]) finally have "(\\<^sup>+x\A. ((\' x) B * RN_deriv \ \' x) \\) = \ (A \ B)" . } thus ?thesis using sets mk2.emeasure_measurable[OF assms(8)] by(auto intro!: measure_eqI s1.RN_deriv_unique simp: emeasure_density emeasure_marginal_measure_on[OF nu_sets assms(8)]sets sets_marginal_measure) qed show ?thesis using 1 2 by auto qed end \ No newline at end of file diff --git a/thys/Hidden_Markov_Models/Hidden_Markov_Model.thy b/thys/Hidden_Markov_Models/Hidden_Markov_Model.thy --- a/thys/Hidden_Markov_Models/Hidden_Markov_Model.thy +++ b/thys/Hidden_Markov_Models/Hidden_Markov_Model.thy @@ -1,626 +1,625 @@ section \Hidden Markov Models\ theory Hidden_Markov_Model imports Markov_Models.Discrete_Time_Markov_Chain Auxiliary "HOL-Library.IArray" begin subsection \Definitions\ text \Definition of Markov Kernels that are closed w.r.t. to a set of states.\ locale Closed_Kernel = fixes K :: "'s \ 't pmf" and S :: "'t set" assumes finite: "finite S" and wellformed: "S \ {}" and closed: "\ s. K s \ S" text \ An HMM is parameterized by a Markov kernel for the transition probabilites between internal states, a Markov kernel for the output probabilities of observations, and a fixed set of observations. \ locale HMM_defs = fixes \ :: "'s \ 's pmf" and \ :: "'s \ 't pmf" and \\<^sub>s :: "'t set" locale HMM = HMM_defs + O: Closed_Kernel \ \\<^sub>s begin lemma observations_finite: "finite \\<^sub>s" and observations_wellformed: "\\<^sub>s \ {}" and observations_closed: "\ s. \ s \ \\<^sub>s" using O.finite O.wellformed O.closed by - end text \Fixed set of internal states.\ locale HMM2_defs = HMM_defs \ \ for \ :: "'s \ 's pmf" and \ :: "'s \ 't pmf" + fixes \ :: "'s set" locale HMM2 = HMM2_defs + HMM + K: Closed_Kernel \ \ begin lemma states_finite: "finite \" and states_wellformed: "\ \ {}" and states_closed: "\ s. \ s \ \" using K.finite K.wellformed K.closed by - end text \ The set of internal states is now given as a list to iterate over. This is needed for the computations on HMMs. \ locale HMM3_defs = HMM2_defs \\<^sub>s \ for \\<^sub>s :: "'t set" and \ :: "'s \ 's pmf" + fixes state_list :: "'s list" locale HMM3 = HMM3_defs _ _ \\<^sub>s \ + HMM2 \\<^sub>s \ for \\<^sub>s :: "'t set" and \ :: "'s \ 's pmf" + assumes state_list_\: "set state_list = \" context HMM_defs begin no_notation (ASCII) comp (infixl "o" 55) text \The ``default'' observation.\ definition "obs \ SOME x. x \ \\<^sub>s" lemma (in HMM) obs: "obs \ \\<^sub>s" unfolding obs_def using observations_wellformed by (auto intro: someI_ex) text \ The HMM is encoded as a Markov chain over pairs of states and observations. This is the Markov chain's defining Markov kernel. \ definition "K \ \ (s\<^sub>1, o\<^sub>1 :: 't). bind_pmf (\ s\<^sub>1) (\ s\<^sub>2. map_pmf (\ o\<^sub>2. (s\<^sub>2, o\<^sub>2)) (\ s\<^sub>2))" sublocale MC_syntax K . text \ Uniform distribution of the pairs \(s, o)\ for a fixed state \s\. \ definition "I (s :: 's) = map_pmf (\ x. (s, x)) (pmf_of_set \\<^sub>s)" text \ The likelihood of an observation sequence given a starting state \s\ is defined in terms of the trace space of the Markov kernel given the uniform distribution of pairs for \s\. \ definition "likelihood s os = T' (I s) {\ \ space S. \ o\<^sub>0 xs \'. \ = (s, o\<^sub>0) ## xs @- \' \ map snd xs = os}" abbreviation (input) "L os \ \ \ xs \'. \ = xs @- \' \ map snd xs = os" lemma likelihood_alt_def: "likelihood s os = T' (I s) {(s, o) ## xs @- \' |o xs \'. map snd xs = os}" unfolding likelihood_def by (simp add: in_S) subsection \Iteration Rule For Likelihood\ lemma L_Nil: "L [] \ = True" by simp lemma emeasure_T_observation_Nil: "T (s, o\<^sub>0) {\ \ space S. L [] \} = 1" by simp lemma L_Cons: "L (o # os) \ \ snd (shd \) = o \ L os (stl \)" apply (cases \; cases "shd \"; safe; clarsimp) apply force subgoal for x xs \' by (force intro: exI[where x = "(x, o) # xs"]) done lemma L_measurable[measurable]: "Measurable.pred S (L os)" apply (induction os) apply (simp; fail) subgoal premises that for o os by(subst L_Cons) (intro Measurable.pred_intros_logic measurable_compose[OF measurable_shd] measurable_compose[OF measurable_stl that]; measurable) done lemma init_measurable[measurable]: "Measurable.pred S (\x. \o\<^sub>0 xs \'. x = (s, o\<^sub>0) ## xs @- \' \ map snd xs = os)" (is "Measurable.pred S ?f") proof - have *: "?f \ \ fst (shd \) = s \ L os (stl \)" for \ by (cases \) auto show ?thesis by (subst *) (intro Measurable.pred_intros_logic measurable_compose[OF measurable_shd]; measurable) qed lemma T_init_observation_eq: "T (s, o) {\ \ space S. L os \} = T (s, o') {\ \ space S. L os \}" apply (subst emeasure_Collect_T[unfolded space_T], (measurable; fail)) apply (subst (2) emeasure_Collect_T[unfolded space_T], (measurable; fail)) apply (simp add: K_def) done text \ Shows that it is equivalent to define likelihood in terms of the trace space starting at a single pair of an internal state \s\ and the default observation @{term obs}. \ lemma (in HMM) likelihood_init: "likelihood s os = T (s, obs) {\ \ space S. L os \}" proof - have *: "(\o\\\<^sub>s. emeasure (T (s, o)) {\ \ space S. L os \}) = of_nat (card \\<^sub>s) * emeasure (T (s, obs)) {\ \ space S. L os \}" by (subst sum_constant[symmetric]) (fastforce intro: sum.cong T_init_observation_eq[simplified]) show ?thesis unfolding likelihood_def apply (subst emeasure_T') subgoal by measurable using * apply (simp add: I_def in_S observations_finite observations_wellformed nn_integral_pmf_of_set) apply (subst mult.commute) apply (simp add: observations_finite observations_wellformed mult_divide_eq_ennreal) done qed lemma emeasure_T_observation_Cons: "T (s, o\<^sub>0) {\ \ space S. L (o\<^sub>1 # os) \} = (\\<^sup>+ t. ennreal (pmf (\ t) o\<^sub>1) * T (t, o\<^sub>1) {\ \ space S. L os \} \(\ s))" (is "?l = ?r") proof - have *: "\\<^sup>+ y. T (s', y) {x \ space S. \xs. (\\'. (s', y) ## x = xs @- \') \ map snd xs = o\<^sub>1 # os} \measure_pmf (\ s') = ennreal (pmf (\ s') o\<^sub>1) * T (s', o\<^sub>1) {\ \ space S. \xs. (\\'. \ = xs @- \') \ map snd xs = os}" (is "?L = ?R") for s' proof - have "?L = \\<^sup>+ x. ennreal (pmf (\ s') x) * T (s', x) {\ \ space S. \xs. (\\'. (s', x) ## \ = xs @- \') \ map snd xs = o\<^sub>1 # os} \count_space UNIV" by (rule nn_integral_measure_pmf) also have "\ = \\<^sup>+ o\<^sub>2. (if o\<^sub>2 = o\<^sub>1 then ennreal (pmf (\ s') o\<^sub>1) * T (s', o\<^sub>1) {\ \ space S. L os \} else 0) \count_space UNIV" apply (rule nn_integral_cong_AE [where v = "\ o\<^sub>2. if o\<^sub>2 = o\<^sub>1 then ennreal (pmf (\ s') o\<^sub>1) * T (s', o\<^sub>1) {\ \ space S. L os \} else 0"] ) apply (rule AE_I2) apply (split if_split, safe) subgoal by (auto intro!: arg_cong2[where f = times, OF HOL.refl] arg_cong2[where f = emeasure]; metis list.simps(9) shift.simps(2) snd_conv ) subgoal by (subst arg_cong2[where f = emeasure and d = "{}", OF HOL.refl]) auto done - also have "\ = \\<^sup>+o\<^sub>2\{o\<^sub>1}. - (ennreal (pmf (\ s') o\<^sub>1) * T (s', o\<^sub>1) {\ \ space S. L os \}) - \count_space UNIV" + also have "\ = (\\<^sup>+o\<^sub>2\{o\<^sub>1}. (ennreal (pmf (\ s') o\<^sub>1) * T (s', o\<^sub>1) {\ \ space S. L os \}) + \count_space UNIV)" by (rule nn_integral_cong_AE) auto also have "\ = ?R" by simp finally show ?thesis . qed have "?l = \\<^sup>+ t. T t {x \ space S. \xs \'. t ## x = xs @- \' \ map snd xs = o\<^sub>1 # os} \ (K (s, o\<^sub>0))" by (subst emeasure_Collect_T[unfolded space_T], measurable) also have "\ = ?r" using * by (simp add: K_def) finally show ?thesis . qed subsection \Computation of Likelihood\ fun backward where "backward s [] = 1" | "backward s (o # os) = (\\<^sup>+ t. ennreal (pmf (\ t) o) * backward t os \measure_pmf (\ s))" lemma emeasure_T_observation_backward: "emeasure (T (s, o)) {\ \ space S. L os \} = backward s os" using emeasure_T_observation_Cons by (induction os arbitrary: s o; simp) lemma (in HMM) likelihood_backward: "likelihood s os = backward s os" unfolding likelihood_init emeasure_T_observation_backward .. end (* HMM Defs *) context HMM2 begin fun (in HMM2_defs) forward where "forward s t_end [] = indicator {t_end} s" | "forward s t_end (o # os) = (\t \ \. ennreal (pmf (\ t) o) * ennreal (pmf (\ s) t) * forward t t_end os)" lemma forward_split: "forward s t (os1 @ os2) = (\t' \ \. forward s t' os1 * forward t' t os2)" if "s \ \" using that apply (induction os1 arbitrary: s) subgoal for s apply (simp add: sum_indicator_mult[OF states_finite]) apply (subst sum.cong[where B = "{s}"]) by auto subgoal for a os1 s apply simp apply (subst sum_distrib_right) apply (subst sum.swap) apply (simp add: sum_distrib_left algebra_simps) done done lemma (in -) "(\t \ S. f t) = f t" if "finite S" "t \ S" "\ s \ S - {t}. f s = 0" thm sum.empty sum.insert sum.mono_neutral_right[of S "{t}"] apply (subst sum.mono_neutral_right[of S "{t}"]) using that apply auto done (* oops by (metis add.right_neutral empty_iff finite.intros(1) insert_iff subsetI sum.empty sum.insert sum.mono_neutral_right that) using that apply auto *) lemma forward_backward: "(\t \ \. forward s t os) = backward s os" if "s \ \" using \s \ \\ apply (induction os arbitrary: s) subgoal for s by (subst sum.mono_neutral_right[of \ "{s}", OF states_finite]) (auto split: if_split_asm simp: indicator_def) subgoal for a os s apply (simp add: sum.swap sum_distrib_left[symmetric]) apply (subst nn_integral_measure_pmf_support[where A = \]) using states_finite states_closed by (auto simp: algebra_simps) done theorem likelihood_forward: "likelihood s os = (\t \ \. forward s t os)" if \s \ \\ unfolding likelihood_backward forward_backward[symmetric, OF \s \ \\] .. subsection \Definition of Maximum Probabilities\ abbreviation (input) "V os as \ \ (\ \'. \ = zip as os @- \')" definition "max_prob s os = Max {T' (I s) {\ \ space S. \o \'. \ = (s, o) ## zip as os @- \'} | as. length as = length os \ set as \ \}" fun viterbi_prob where "viterbi_prob s t_end [] = indicator {t_end} s" | "viterbi_prob s t_end (o # os) = (MAX t \ \. ennreal (pmf (\ t) o * pmf (\ s) t) * viterbi_prob t t_end os)" definition "is_decoding s os as \ T' (I s) {\ \ space S. \o \'. \ = (s, o) ## zip as os @- \'} = max_prob s os \ length as = length os \ set as \ \" subsection \Iteration Rule For Maximum Probabilities\ lemma emeasure_T_state_Nil: "T (s, o\<^sub>0) {\ \ space S. V [] as \} = 1" by simp lemma max_prob_T_state_Nil: "Max {T (s, o) {\ \ space S. V [] as \} | as. length as = length [] \ set as \ \} = 1" by (simp add: emeasure_T_state_Nil) lemma V_Cons: "V (o # os) (a # as) \ \ fst (shd \) = a \ snd (shd \) = o \ V os as (stl \)" by (cases \) auto lemma measurable_V[measurable]: "Measurable.pred S (\\. V os as \)" proof (induction os as rule: list_induct2') case (4 x xs y ys) then show ?case by (subst V_Cons) (intro Measurable.pred_intros_logic measurable_compose[OF measurable_shd] measurable_compose[OF measurable_stl]; measurable) qed simp+ lemma init_V_measurable[measurable]: "Measurable.pred S (\x. \o \'. x = (s, o) ## zip as os @- \')" (is "Measurable.pred S ?f") proof - have *: "?f \ \ fst (shd \) = s \ V os as (stl \)" for \ by (cases \) auto show ?thesis by (subst *) (intro Measurable.pred_intros_logic measurable_compose[OF measurable_shd]; measurable) qed lemma max_prob_Cons': "Max {T (s, o\<^sub>1) {\ \ space S. V (o # os) as \} | as. length as = length (o # os) \ set as \ \} = ( MAX t \ \. ennreal (pmf (\ t) o * pmf (\ s) t) * (MAX as \ {as. length as = length os \ set as \ \}. T (t, o) {\ \ space S. V os as \}) )" (is "?l = ?r") and T_V_Cons: "T (s, o\<^sub>1) {\ \ space S. V (o # os) (t # as) \} = ennreal (pmf (\ t) o * pmf (\ s) t) * T (t, o) {\ \ space S. V os as \}" (is "?l' = ?r'") if "length as = length os" proof - let ?S = "\ os. {as. length as = length os \ set as \ \}" have S_finite: "finite (?S os)" for os :: "'t list" using finite_lists_length_eq[OF states_finite] by (rule finite_subset[rotated]) auto have S_nonempty: "?S os \ {}" for os :: "'t list" proof - let ?a = "SOME a. a \ \" let ?as = "replicate (length os) ?a" from states_wellformed have "?a \ \" by (auto intro: someI_ex) then have "?as \ ?S os" by auto then show ?thesis by force qed let ?f = "\t as os. T t {\ \ space S. V os as (t ## \)}" let ?g = "\t as os. T t {\ \ space S. V os as \}" have *: "?f t as (o # os) = ?g t (tl as) os * indicator {(hd as, o)} t" if "length as = Suc n" for t as n unfolding indicator_def using that by (cases as) auto have **: "K (s, o\<^sub>1) {(t, o)} = pmf (\ t) o * pmf (\ s) t" for t unfolding K_def apply (simp add: vimage_def) apply (subst arg_cong2[where f = nn_integral and d = "\ x. \ x {xa. xa = o \ x = t} * indicator {t} x", OF HOL.refl]) subgoal by (auto simp: indicator_def) by (simp add: emeasure_pmf_single ennreal_mult') have "?l = (MAX as \ ?S (o # os). \\<^sup>+ t. ?f t as (o # os) \K (s, o\<^sub>1))" by (subst Max_to_image2; subst emeasure_Collect_T[unfolded space_T]; rule measurable_V HOL.refl) also have "\ = (MAX as \ ?S (o # os). \\<^sup>+ t. ?g t (tl as) os * indicator {(hd as,o)} t \K (s,o\<^sub>1))" by (simp cong: Max_image_cong_simp add: *) also have "\ = (MAX(t, as)\ \ \ ?S os. ennreal (pmf (\ t) o * pmf (\ s) t) * ?g (t, o) as os)" proof ((rule Max_eq_image_if; clarsimp?), goal_cases) case 1 from S_finite[of "o # os"] show ?case by simp next case 2 from states_finite show ?case by (blast intro: S_finite) next case (3 as) then show ?case by - (rule bexI[where x = "hd as"]; cases as; auto simp: algebra_simps **) next case (4 x as) then show ?case by - (rule exI[where x = "x # as"], simp add: algebra_simps **) qed also have "\ = ?r" by (subst Max_image_left_mult[symmetric], fact+) (rule sym, rule Max_image_pair, rule states_finite, fact+) finally show "?l = ?r" . have "?l' = \\<^sup>+ t'. ?f t' (t # as) (o # os) \K (s, o\<^sub>1)" by (rule emeasure_Collect_T[unfolded space_T]; rule measurable_V) also from that have "\ = \\<^sup>+ t'. ?g t' as os * indicator {(t,o)} t' \K (s,o\<^sub>1)" by (subst *[of _ "length as"]; simp) also have "\ = ?r'" by (simp add: **, simp only: algebra_simps) finally show "?l' = ?r'" . qed lemmas max_prob_Cons = max_prob_Cons'[OF length_replicate] subsection \Computation of Maximum Probabilities\ lemma T_init_V_eq: "T (s, o) {\ \ space S. V os as \} = T (s, o') {\ \ space S. V os as \}" apply (subst emeasure_Collect_T[unfolded space_T], (measurable; fail)) apply (subst (2) emeasure_Collect_T[unfolded space_T], (measurable; fail)) apply (simp add: K_def) done lemma T'_I_T: "T' (I s) {\ \ space S. \o \'. \ = (s, o) ## zip as os @- \'} = T (s,o) {\ \ space S. V os as \}" proof - have "(\o\\\<^sub>s. T (s, o) {\ \ space S. V os as \}) = of_nat (card \\<^sub>s) * T (s, o) {\ \ space S. V os as \}" for as by (subst sum_constant[symmetric]) (fastforce intro: sum.cong T_init_V_eq[simplified]) then show ?thesis unfolding max_prob_def apply (subst emeasure_T') subgoal by measurable apply (simp add: I_def in_S observations_finite observations_wellformed nn_integral_pmf_of_set) apply (subst mult.commute) apply (simp add: observations_finite observations_wellformed mult_divide_eq_ennreal) done qed lemma max_prob_init: "max_prob s os = Max {T (s,o) {\ \ space S. V os as \} | as. length as = length os \ set as \ \}" unfolding max_prob_def by (simp add: T'_I_T[symmetric]) lemma max_prob_Nil[simp]: "max_prob s [] = 1" unfolding max_prob_init[where o = obs] by auto lemma Max_start: "(MAX t\\. (indicator {t} s :: ennreal)) = 1" if "s \ \" using states_finite that by (auto simp: indicator_def intro: Max_eqI) lemma Max_V_viterbi: "(MAX t \ \. viterbi_prob s t os) = Max {T (s, o) {\ \ space S. V os as \} | as. length as = length os \ set as \ \}" if "s \ \" using that states_finite states_wellformed by (induction os arbitrary: s o; simp add: Max_start max_prob_Cons[simplified] Max_image_commute Max_image_left_mult Max_to_image2 cong: Max_image_cong ) lemma max_prob_viterbi: "(MAX t \ \. viterbi_prob s t os) = max_prob s os" if "s \ \" using max_prob_init[of s os] Max_V_viterbi[OF \s \ \\, symmetric] by simp end subsection \Decoding the Most Probable Hidden State Sequence\ context HMM3 begin fun viterbi where "viterbi s t_end [] = ([], indicator {t_end} s)" | "viterbi s t_end (o # os) = fst ( argmax snd (map (\t. let (xs, v) = viterbi t t_end os in (t # xs, ennreal (pmf (\ t) o * pmf (\ s) t) * v)) state_list))" lemma state_list_nonempty: "state_list \ []" using state_list_\ states_wellformed by auto lemma viterbi_viterbi_prob: "snd (viterbi s t_end os) = viterbi_prob s t_end os" proof (induction os arbitrary: s) case Nil then show ?case by simp next case (Cons o os) let ?f = "\t. let (xs, v) = viterbi t t_end os in (t # xs, ennreal (pmf (\ t) o * pmf (\ s) t) * v)" let ?xs = "map ?f state_list" from state_list_nonempty have "map ?f state_list \ []" by simp from argmax(2,3)[OF this, of snd] have *: "snd (fst (argmax snd ?xs)) = snd (argmax snd ?xs)" "snd (argmax snd ?xs) = (MAX x \ set ?xs. snd x)" . then show ?case apply (simp add: state_list_\) apply (rule Max_eq_image_if) apply (intro finite_imageI states_finite; fail) apply (intro finite_imageI states_finite; fail) subgoal apply clarsimp subgoal for x using Cons.IH[of x] by (auto split: prod.splits) done apply clarsimp subgoal for x using Cons.IH[of x] by (force split: prod.splits) done qed context begin private fun val_of where "val_of s [] [] = 1" | "val_of s (t # xs) (o # os) = ennreal (pmf (\ t) o * pmf (\ s) t) * val_of t xs os" lemma val_of_T: "val_of s as os = T (s, o\<^sub>1) {\ \ space S. V os as \}" if "length as = length os" using that by (induction arbitrary: o\<^sub>1 rule: val_of.induct; (subst T_V_Cons)?; simp) lemma viterbi_sequence: "snd (viterbi s t_end os) = val_of s (fst (viterbi s t_end os)) os" if "snd (viterbi s t_end os) > 0" using that proof (induction os arbitrary: s) case Nil then show ?case by (simp add: indicator_def split: if_split_asm split_of_bool_asm) next case (Cons o os s) let ?xs = "map (\t. let (xs, v) = viterbi t t_end os in (t # xs, ennreal (pmf (\ t) o * pmf (\ s) t) * v)) state_list" from state_list_nonempty have "?xs \ []" by simp from argmax(1)[OF this, of snd] obtain t where "t \ set state_list" "fst (argmax snd ?xs) = (t # fst (viterbi t t_end os), ennreal (pmf (\ t) o * pmf (\ s) t) * snd (viterbi t t_end os))" by (auto split: prod.splits) with Cons show ?case by (auto simp: ennreal_zero_less_mult_iff) qed lemma viterbi_valid_path: "length as = length os \ set as \ \" if "viterbi s t_end os = (as, v)" using that proof (induction os arbitrary: s as v) case Nil then show ?case by simp next case (Cons o os s as v) let ?xs = "map (\t. let (xs, v) = viterbi t t_end os in (t # xs, ennreal (pmf (\ t) o * pmf (\ s) t) * v)) state_list" from state_list_nonempty have "?xs \ []" by simp from argmax(1)[OF this, of snd] obtain t where "t \ \" "fst (argmax snd ?xs) = (t # fst (viterbi t t_end os), ennreal (pmf (\ t) o * pmf (\ s) t) * snd (viterbi t t_end os))" by (auto simp: state_list_\ split: prod.splits) with Cons.prems show ?case by (cases "viterbi t t_end os"; simp add: Cons.IH) qed definition "viterbi_final s os = fst (argmax snd (map (\ t. viterbi s t os) state_list))" lemma viterbi_finalE: obtains t where "t \ \" "viterbi_final s os = viterbi s t os" "snd (viterbi s t os) = Max ((\t. snd (viterbi s t os)) ` \)" proof - from state_list_nonempty have "map (\ t. viterbi s t os) state_list \ []" by simp from argmax[OF this, of snd] show ?thesis by (auto simp: state_list_\ image_comp comp_def viterbi_final_def intro: that) qed theorem viterbi_final_max_prob: assumes "viterbi_final s os = (as, v)" "s \ \" shows "v = max_prob s os" proof - obtain t where "t \ \" "viterbi_final s os = viterbi s t os" "snd (viterbi s t os) = Max ((\t. snd (viterbi s t os)) ` \)" by (rule viterbi_finalE) with assms show ?thesis by (simp add: viterbi_viterbi_prob max_prob_viterbi) qed theorem viterbi_final_is_decoding: assumes "viterbi_final s os = (as, v)" "v > 0" "s \ \" shows "is_decoding s os as" proof - from viterbi_valid_path[of s _ os as v] assms have as: "length as = length os" "set as \ \" by - (rule viterbi_finalE[of s os]; simp)+ obtain t where "t \ \" "viterbi_final s os = viterbi s t os" by (rule viterbi_finalE) with assms viterbi_sequence[of s t os] have "val_of s as os = v" by (cases "viterbi s t os") (auto simp: snd_def split!: prod.splits) with val_of_T as have "max_prob s os = T (s, obs) {\ \ space S. V os as \}" by (simp add: viterbi_final_max_prob[OF assms(1,3)]) with as show ?thesis unfolding is_decoding_def by (simp only: T'_I_T) qed end (* Anonymous context *) end (* HMM 3 *) end (* Theory *) \ No newline at end of file