diff --git a/thys/Actuarial_Mathematics/Life_Table.thy b/thys/Actuarial_Mathematics/Life_Table.thy --- a/thys/Actuarial_Mathematics/Life_Table.thy +++ b/thys/Actuarial_Mathematics/Life_Table.thy @@ -1,746 +1,746 @@ theory Life_Table imports Survival_Model begin section \Life Table\ text \Define a life table axiomatically.\ locale life_table = fixes l :: "real \ real" ("$l'__" [101] 200) assumes l_0_pos: "0 < l 0" and l_neg_nil: "\x. x \ 0 \ l x = l 0" and l_PInfty_0: "(l \ 0) at_top " and l_antimono: "antimono l" and l_right_continuous: "\x. continuous (at_right x) l" begin subsection \Basic Properties of Life Table\ definition death :: "real \ real \ real" ("$d'_{_&_}" [0,0] 200) where "$d_{t&x} \ $l_x - $l_(x+t)" \ \the number of deaths between ages x and x+t\ \ \The parameter t is usually nonnegative, but theoretically it can be negative.\ abbreviation death1 :: "real \ real" ("$d'__" [101] 200) where "$d_x \ $d_{1&x}" lemma l_0_neq_0[simp]: "$l_0 \ 0" using l_0_pos by simp lemma l_nonneg[simp]: "$l_x \ 0" for x::real using l_antimono l_PInfty_0 by (rule antimono_at_top_le) lemma l_bounded[simp]: "$l_x \ $l_0" for x::real using l_neg_nil l_antimono by (smt (verit) antimonoD) lemma l_measurable[measurable, simp]: "l \ borel_measurable borel" by (rule borel_measurable_antimono, rule l_antimono) lemma l_integrable_Icc: "set_integrable lborel {a..b} l" for a b :: real unfolding set_integrable_def apply (rule integrableI_bounded_set[where A="{a..b}" and B="$l_0"], simp_all) using emeasure_compact_finite by auto corollary l_integrable_on_Icc: "l integrable_on {a..b}" for a b :: real using l_integrable_Icc by (rewrite integrable_on_iff_set_integrable_nonneg[THEN sym]; simp) lemma d_0_0: "$d_{0&x} = 0" for x::real unfolding death_def by simp lemma d_nonneg[simp]: "$d_{t&x} \ 0" if "t \ 0" for t x :: real using l_antimono that unfolding death_def antimono_def by simp lemma dx_l: "$d_x = $l_x - $l_(x+1)" for x::real unfolding death_def by simp lemma add_d: "$d_{t&x} + $d_{t' & x+t} = $d_{t+t' & x}" for t t' :: real unfolding death_def by (smt (verit)) lemma l_normal_antimono: "antimono (\x. $l_x / $l_0)" using divide_le_cancel l_0_pos l_antimono unfolding antimono_def by fastforce lemma compl_l_normal_right_continuous: "continuous (at_right x) (\x. 1 - $l_x / $l_0)" for x::real using l_0_pos l_right_continuous by (intro continuous_intros; simp) lemma compl_l_normal_NInfty_0: "((\x. 1 - $l_x / $l_0) \ 0) at_bot" apply (rewrite tendsto_cong[where g="\_. 0"], simp_all) by (smt (verit) div_self eventually_at_bot_linorder l_0_pos l_neg_nil) lemma compl_l_normal_PInfty_1: "((\x. 1 - $l_x / $l_0) \ 1) at_top" using l_0_pos l_PInfty_0 by (intro tendsto_eq_intros) simp_all+ lemma compl_l_real_distribution: "real_distribution (interval_measure (\x. 1 - $l_x / $l_0))" using l_normal_antimono compl_l_normal_right_continuous compl_l_normal_NInfty_0 compl_l_normal_PInfty_1 by (intro real_distribution_interval_measure; simp add: antimono_def) subsection \Construction of Survival Model from Life Table\ definition life_table_measure :: "real measure" ("\") where "\ \ interval_measure (\x. 1 - $l_x / $l_0)" lemma prob_space_actuary_MM: "prob_space_actuary \" unfolding life_table_measure_def using compl_l_real_distribution real_distribution_def by (intro prob_space_actuary.intro) force definition survival_model_X :: "real \ real" ("X") where "X \ \x. x" lemma survival_model_MM_X: "survival_model \ X" proof - let ?F = "\x. 1 - $l_x / $l_0" show "survival_model \ X" unfolding life_table_measure_def survival_model_X_def proof (rule survival_model.intro) show "prob_space_actuary (interval_measure ?F)" using prob_space_actuary_MM unfolding life_table_measure_def by simp show "survival_model_axioms (interval_measure ?F) (\x. x)" proof (rule survival_model_axioms.intro, simp_all) have [simp]: "{\::real. \ \ 0} = {..0}" by blast have "measure (interval_measure (\x. 1 - $l_x / $l_0)) {..0} = 0" using l_normal_antimono compl_l_normal_right_continuous compl_l_normal_NInfty_0 by (rewrite measure_interval_measure_Iic, simp_all add: antimono_def) hence "emeasure (interval_measure (\x. 1 - $l_x / $l_0)) {..0} = ennreal 0" apply (rewrite finite_measure.emeasure_eq_measure, simp_all) using compl_l_real_distribution prob_space_def real_distribution_def by blast thus "almost_everywhere (interval_measure (\x. 1 - $l_x / $l_0)) (\x. 0 < x)" apply (rewrite AE_iff_null, simp) by (rewrite not_less) auto qed qed qed end sublocale life_table \ survival_model \ X by (rule survival_model_MM_X) context life_table begin interpretation distrX_RD: real_distribution "distr \ borel X" using MM_PS.real_distribution_distr by simp subsubsection \Relations between Life Table and Survival Function for X\ lemma ccdfX_l_normal: "ccdf (distr \ borel X) = (\x. $l_x / $l_0)" proof (rule ext) let ?F = "\x. 1 - $l_x / $l_0" interpret F_FBM: finite_borel_measure "interval_measure ?F" using compl_l_real_distribution real_distribution.finite_borel_measure_M by blast show "\x. ccdf (distr \ borel X) x = $l_x / $l_0" unfolding ccdf_def life_table_measure_def survival_model_X_def apply (rewrite measure_distr, simp_all) using l_normal_antimono compl_l_normal_right_continuous compl_l_normal_NInfty_0 compl_l_normal_PInfty_1 by (rewrite F_FBM.measure_interval_measure_Ioi; simp add: antimono_def) qed corollary deriv_ccdfX_l: "deriv (ccdf (distr \ borel X)) x = deriv l x / $l_0" if "l differentiable at x" for x::real using differentiable_eq_field_differentiable_real that by (rewrite ccdfX_l_normal, rewrite deriv_cdivide_right; simp) notation death_pt ("$\") lemma l_0_equiv: "$l_x = 0 \ x \ $\" for x::real using ccdfX_l_normal ccdfX_0_equiv by simp lemma d_old_0: "$d_{t&x} = 0" if "x \ $\" "t \ 0" for x t :: real unfolding death_def using l_0_equiv that by (smt (verit) le_ereal_le) lemma d_l_equiv: "$d_{t&x} = $l_x \ x+t \ $\" for x t :: real unfolding death_def using l_0_equiv by simp lemma continuous_ccdfX_l: "continuous F (ccdf (distr \ borel X)) \ continuous F l" for F :: "real filter" proof - have "continuous F (ccdf (distr \ borel X)) \ continuous F (\x. $l_x / $l_0)" using ccdfX_l_normal by simp also have "\ \ continuous F l" using continuous_cdivide_iff l_0_neq_0 by blast finally show ?thesis . qed lemma has_real_derivative_ccdfX_l: "(ccdf (distr \ borel X) has_real_derivative D) (at x) \ (l has_real_derivative $l_0 * D) (at x)" for D x :: real proof - have "(ccdf (distr \ borel X) has_real_derivative D) (at x) \ ((\x. $l_x / $l_0) has_real_derivative D) (at x)" by (rule has_field_derivative_cong_eventually; simp add: ccdfX_l_normal) also have "\ \ ((\x. $l_0 * ($l_x / $l_0)) has_real_derivative $l_0 * D) (at x)" by (rule DERIV_cmult_iff, simp) also have "\ \ (l has_real_derivative $l_0 * D) (at x)" by simp finally show ?thesis . qed corollary differentiable_ccdfX_l: "ccdf (distr \ borel X) differentiable (at x) \ l differentiable (at x)" for D x :: real using has_real_derivative_ccdfX_l by (metis l_0_neq_0 mult.commute nonzero_divide_eq_eq real_differentiable_def) lemma PX_l_normal: "\

(\ in \. X \ > x) = $l_x / $l_0" for x::real using MM_PS.ccdf_distr_P ccdfX_l_normal X_RV by (metis (mono_tags, lifting) Collect_cong) lemma set_integrable_ccdfX_l: "set_integrable lborel A (ccdf (distr \ borel X)) \ set_integrable lborel A l" if "A \ sets lborel" for A :: "real set" proof - have "set_integrable lborel A (ccdf (distr \ borel X)) \ set_integrable lborel A (\x. $l_x / $l_0)" by (rule set_integrable_cong; simp add: ccdfX_l_normal) also have "\ \ set_integrable lborel A l" by simp finally show ?thesis . qed corollary integrable_ccdfX_l: "integrable lborel (ccdf (distr \ borel X)) \ integrable lborel l" using set_integrable_ccdfX_l[where A=UNIV] by (simp add: set_integrable_def) lemma integrable_on_ccdfX_l: "ccdf (distr \ borel X) integrable_on A \ l integrable_on A" for A :: "real set" proof - have "ccdf (distr \ borel X) integrable_on A \ (\x. $l_x / $l_0) integrable_on A" by (rule integrable_cong) (simp add: ccdfX_l_normal) also have "\ \ l integrable_on A" using integrable_on_cdivide_iff[of "$l_0" l] by simp finally show ?thesis . qed subsubsection \Relations between Life Table and Cumulative Distributive Function for X\ lemma cdfX_l_normal: "cdf (distr \ borel X) = (\x. 1 - $l_x / $l_0)" for x::real using ccdfX_l_normal distrX_RD.cdf_ccdf distrX_RD.prob_space by presburger lemma deriv_cdfX_l: "deriv (cdf (distr \ borel X)) x = - deriv l x / $l_0" if "l differentiable at x" for x::real using distrX_RD.cdf_ccdf differentiable_eq_field_differentiable_real that differentiable_ccdfX_l deriv_diff deriv_ccdfX_l that by simp lemma continuous_cdfX_l: "continuous F (cdf (distr \ borel X)) \ continuous F l" for F :: "real filter" using distrX_RD.continuous_cdf_ccdf continuous_ccdfX_l by simp lemma has_real_derivative_cdfX_l: "(cdf (distr \ borel X) has_real_derivative D) (at x) \ (l has_real_derivative - ($l_0 * D)) (at x)" for D x :: real using distrX_RD.has_real_derivative_cdf_ccdf has_real_derivative_ccdfX_l by simp lemma differentiable_cdfX_l: "cdf (distr \ borel X) differentiable (at x) \ l differentiable (at x)" for D x :: real using differentiable_eq_field_differentiable_real distrX_RD.differentiable_cdf_ccdf differentiable_ccdfX_l by simp lemma PX_compl_l_normal: "\

(\ in \. X \ \ x) = 1 - $l_x / $l_0" for x::real using PX_l_normal by (metis MM_PS.prob_compl X_compl_gt_le X_gt_event) subsubsection \Relations between Life Table and Survival Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin notation futr_life ("T") interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma lx_neq_0[simp]: "$l_x \ 0" using l_0_equiv x_lt_psi linorder_not_less by blast corollary lx_pos[simp]: "$l_x > 0" using lx_neq_0 l_nonneg by (smt (verit)) lemma ccdfTx_l_normal: "ccdf (distr (\ \ alive x) borel (T x)) t = $l_(x+t) / $l_x" if "t \ 0" for t::real using ccdfTx_PX PX_l_normal l_0_neq_0 that by simp lemma deriv_ccdfTx_l: "deriv (ccdf (distr (\ \ alive x) borel (T x))) t = deriv (\t. $l_(x+t) / $l_x) t" if "t > 0" "l differentiable at (x+t)" for t::real proof - have "\\<^sub>F s in nhds t. ccdf (distr (\ \ alive x) borel (T x)) s = $l_(x+s) / $l_x" apply (rewrite eventually_nhds_metric) using that ccdfTx_l_normal dist_real_def by (intro exI[of _ t]) auto thus ?thesis by (rule deriv_cong_ev) simp qed lemma continuous_at_within_ccdfTx_l: "continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x))) \ continuous (at (x+t) within {x..}) l" if "t \ 0" for t::real using continuous_ccdfX_ccdfTx that continuous_ccdfX_l by force lemma isCont_ccdfTx_l: "isCont (ccdf (distr (\ \ alive x) borel (T x))) t \ isCont l (x+t)" if "t > 0" for t::real using that continuous_ccdfX_l isCont_ccdfX_ccdfTx by force lemma has_real_derivative_ccdfTx_l: "(ccdf (distr (\ \ alive x) borel (T x)) has_real_derivative D) (at t) \ (l has_real_derivative $l_x * D) (at (x+t))" if "t > 0" for t D :: real proof - have "(ccdf (distr (\ \ alive x) borel (T x)) has_real_derivative D) (at t) \ (ccdf (distr (\ \ alive x) borel (T x)) has_real_derivative ($l_x / $l_0 * D / \

(\ in \. X \ > x))) (at t)" using PX_l_normal by force also have "\ = (ccdf (distr \ borel X) has_real_derivative ($l_x / $l_0 * D)) (at (x+t))" using has_real_derivative_ccdfX_ccdfTx that by simp also have "\ = (l has_real_derivative ($l_x * D)) (at (x+t))" using has_real_derivative_ccdfX_l by simp finally show ?thesis . qed lemma differentiable_ccdfTx_l: "ccdf (distr (\ \ alive x) borel (T x)) differentiable at t \ l differentiable (at (x+t))" if "t > 0" for t::real using differentiable_ccdfX_ccdfTx differentiable_ccdfX_l that by force lemma PTx_l_normal: "\

(\ in \. T x \ > t \ T x \ > 0) = $l_(x+t) / $l_x" if "t \ 0" for t::real using ccdfTx_l_normal that by (simp add: ccdfTx_cond_prob) subsubsection \Relations between Life Table and Cumulative Distributive Function for T(x)\ lemma cdfTx_compl_l_normal: "cdf (distr (\ \ alive x) borel (T x)) t = 1 - $l_(x+t) / $l_x" if "t \ 0" for t::real using distrTx_RD.cdf_ccdf ccdfTx_l_normal that distrTx_RD.prob_space by auto lemma deriv_cdfTx_l: "deriv (cdf (distr (\ \ alive x) borel (T x))) t = - deriv (\t. $l_(x+t) / $l_x) t" if "t > 0" "l differentiable at (x+t)" for t::real using deriv_ccdfTx_l differentiable_cdfX_cdfTx differentiable_cdfX_l distrTx_RD.deriv_cdf_ccdf that by fastforce lemma continuous_at_within_cdfTx_l: "continuous (at t within {0..}) (cdf (distr (\ \ alive x) borel (T x))) \ continuous (at (x+t) within {x..}) l" if "t \ 0" for t::real using that continuous_cdfX_l continuous_cdfX_cdfTx by force lemma isCont_cdfTx_l: "isCont (cdf (distr (\ \ alive x) borel (T x))) t \ isCont l (x+t)" if "t > 0" for t::real using that continuous_cdfX_l isCont_cdfX_cdfTx by force lemma has_real_derivative_cdfTx_l: "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative D) (at t) \ (l has_real_derivative - $l_x * D) (at (x+t))" if "t > 0" for t D :: real using has_real_derivative_ccdfTx_l that distrTx_RD.has_real_derivative_cdf_ccdf by auto lemma differentiable_cdfTx_l: "cdf (distr (\ \ alive x) borel (T x)) differentiable at t \ l differentiable (at (x+t))" if "t > 0" for t::real using differentiable_cdfX_l that differentiable_cdfX_cdfTx by auto lemma PTx_compl_l_normal: "\

(\ in \. T x \ \ t \ T x \ > 0) = 1- $l_(x+t) / $l_x" if "t \ 0" for t::real using cdfTx_compl_l_normal that by (simp add: cdfTx_cond_prob) subsubsection \Life Table and Actuarial Notations\ notation survive ("$p'_{_&_}" [0,0] 200) notation survive_1 ("$p'__" [101] 200) notation die ("$q'_{_&_}" [0,0] 200) notation die_1 ("$q'__" [101] 200) notation die_defer ("$q'_{_\_&_}" [0,0,0] 200) notation die_defer_1 ("$q'_{_\&_}" [0,0] 200) notation life_expect ("$e`\'__" [101] 200) notation temp_life_expect ("$e`\'_{_:_}" [0,0] 200) notation curt_life_expect ("$e'__" [101] 200) notation temp_curt_life_expect ("$e'_{_:_}" [0,0] 200) lemma p_l: "$p_{t&x} = $l_(x+t) / $l_x" if "t \ 0" for t::real unfolding survive_def using ccdfTx_l_normal that by simp corollary p_1_l: "$p_x = $l_(x+1) / $l_x" using p_l by simp lemma isCont_p_l: "isCont (\s. $p_{s&x}) t \ isCont l (x+t)" if "t > 0" for t::real proof - have "\\<^sub>F s in nhds t. $p_{s&x} = $l_(x+s) / $l_x" apply (rewrite eventually_nhds_metric) apply (rule exI[of _ t], auto simp add: that) by (rewrite p_l; simp add: dist_real_def) hence "isCont (\s. $p_{s&x}) t \ isCont (\s. $l_(x+s) / $l_x) t" by (rule isCont_cong) also have "\ \ isCont (\s. $l_(x+s)) t" using continuous_cdivide_iff lx_neq_0 by metis also have "\ \ isCont l (x+t)" using isCont_shift by (force simp add: add.commute) finally show ?thesis . qed lemma p_PTx_ge_l_isCont: "$p_{t&x} = \

(\ in \. T x \ \ t \ T x \ > 0)" if "isCont l (x+t)" "t > 0" for t::real using p_PTx_ge_ccdf_isCont that continuous_ccdfX_l by force lemma q_defer_l: "$q_{f\t&x} = ($l_(x+f) - $l_(x+f+t)) / $l_x" if "f \ 0" "t \ 0" for f t :: real apply (rewrite q_defer_p, simp_all add: that) using that by (rewrite p_l, simp)+ (smt (verit) diff_divide_distrib) corollary q_defer_d_l: "$q_{f\t&x} = $d_{t & x+f} / $l_x" if "f \ 0" "t \ 0" for f t :: real using q_defer_l that death_def by simp corollary q_defer_1_d_l: "$q_{f\&x} = $d_(x+f) / $l_x" if "f \ 0" for f::real using q_defer_d_l that by simp corollary q_d_l: "$q_{t&x} = $d_{t&x} / $l_x" if "t \ 0" for t::real using q_defer_d_l[of 0] that by simp corollary q_1_d_l: "$q_x = $d_x / $l_x" using q_d_l by simp -lemma LBINT_p_l: "LBINT t:A. $p_{t&x} = (LBINT t:A. $l_(x+t)) / $l_x" +lemma LBINT_p_l: "(LBINT t:A. $p_{t&x}) = (LBINT t:A. $l_(x+t)) / $l_x" if "A \ {0..}" "A \ sets lborel" for A :: "real set" \ \Note that 0 = 0 holds when the integral diverges.\ proof - have [simp]: "\t. t \ A \ $p_{t&x} = $l_(x+t) / $l_x" using p_l that by blast - hence "LBINT t:A. $p_{t&x} = LBINT t:A. $l_(x+t) / $l_x" + hence "(LBINT t:A. $p_{t&x}) = (LBINT t:A. $l_(x+t) / $l_x)" using that by (rewrite set_lebesgue_integral_cong[where g="\t. $l_(x+t) / $l_x"]; simp) also have "\ = (LBINT t:A. $l_(x+t)) / $l_x" by (rewrite set_integral_divide_zero) simp finally show ?thesis . qed corollary e_LBINT_l: "$e`\_x = (LBINT t:{0..}. $l_(x+t)) / $l_x" \ \Note that 0 = 0 holds when the integral diverges.\ by (simp add: e_LBINT_p LBINT_p_l) corollary e_LBINT_l_Icc: "$e`\_x = (LBINT t:{0..n}. $l_(x+t)) / $l_x" if "x+n \ $\" for n::real using e_LBINT_p_Icc by (rewrite LBINT_p_l[THEN sym]; simp add: that) lemma temp_e_LBINT_l: "$e`\_{x:n} = (LBINT t:{0..n}. $l_(x+t)) / $l_x" if "n \ 0" for n::real using temp_e_LBINT_p by (rewrite LBINT_p_l[THEN sym]; simp add: that) lemma integral_p_l: "integral A (\t. $p_{t&x}) = (integral A (\t. $l_(x+t))) / $l_x" if "A \ {0..}" "A \ sets lborel" for A :: "real set" \ \Note that 0 = 0 holds when the integral diverges.\ using that apply (rewrite set_borel_integral_eq_integral_nonneg[THEN sym], simp_all) apply (simp add: survive_def) apply (rewrite set_borel_integral_eq_integral_nonneg[THEN sym], simp_all) by (rule LBINT_p_l; simp) corollary e_integral_l: "$e`\_x = integral {0..} (\t. $l_(x+t)) / $l_x" if "(\t. $l_(x+t)) integrable_on {0..}" by (simp add: e_integral_p integral_p_l) corollary e_integral_l_Icc: "$e`\_x = integral {0..n} (\t. $l_(x+t)) / $l_x" if "x+n \ $\" for n::real using e_integral_p_Icc by (rewrite integral_p_l[THEN sym]; simp add: that) lemma temp_e_integral_l: "$e`\_{x:n} = integral {0..n} (\t. $l_(x+t)) / $l_x" if "n \ 0" for n::real using temp_e_integral_p by (rewrite integral_p_l[THEN sym]; simp add: that) lemma curt_e_sum_l: "$e_x = (\k. $l_(x+k+1)) / $l_x" if "summable (\k. $l_(x+k+1))" "\k::nat. isCont l (x+k+1)" proof - have "summable (\k. $p_{k+1&x})" apply (rewrite p_l, simp) apply (rule summable_divide) by (rewrite add.assoc[THEN sym], simp add: that) moreover have "\k::nat. isCont (\t. $p_{t&x}) (k+1)" using isCont_p_l that by (simp add: add.assoc) ultimately show ?thesis apply (rewrite curt_e_sum_p, simp_all) using p_l by (smt (verit) of_nat_0_le_iff suminf_cong suminf_divide that(1)) qed lemma curt_e_sum_l_finite: "$e_x = (\kk::nat. k < n \ isCont l (x+k+1)" "x+n+1 > $\" for n::nat apply (rewrite curt_e_sum_p_finite[of x n], simp_all add: that) using isCont_p_l that apply (simp add: add.assoc) apply (rewrite sum_divide_distrib, rule sum.cong, simp) using p_l by (smt (verit) of_nat_0_le_iff) lemma temp_curt_e_sum_p: "$e_{x:n} = (\kk::nat. k < n \ isCont l (x+k+1)" for n::nat apply (rewrite temp_curt_e_sum_p[of x n], simp_all add: that) using isCont_p_l that apply (simp add: add.assoc) apply (rewrite sum_divide_distrib, rule sum.cong, simp) using p_l by (smt (verit) of_nat_0_le_iff) end lemma l_p: "$l_x / $l_0 = $p_{x&0}" for x::real using ccdfX_l_normal ccdfX_p by force end subsection \Piecewise Differentiable Life Table\ locale smooth_life_table = life_table + assumes l_piecewise_differentiable[simp]: "l piecewise_differentiable_on UNIV" begin lemma smooth_survival_function_MM_X: "smooth_survival_function \ X" proof (rule smooth_survival_function.intro) show "survival_model \ X" by (rule survival_model_axioms) show "smooth_survival_function_axioms \ X" proof show "ccdf (distr \ borel X) piecewise_differentiable_on UNIV" apply (rewrite ccdfX_l_normal) apply (rewrite divide_inverse, rewrite mult.commute) using l_piecewise_differentiable piecewise_differentiable_scaleR[of l] by simp qed qed end sublocale smooth_life_table \ smooth_survival_function \ X by (rule smooth_survival_function_MM_X) context smooth_life_table begin notation force_mortal ("$\'__" [101] 200) lemma l_continuous[simp]: "continuous_on UNIV l" using l_piecewise_differentiable piecewise_differentiable_on_imp_continuous_on by fastforce lemma l_nondifferentiable_finite_set[simp]: "finite {x. \ l differentiable at x}" using differentiable_ccdfX_l ccdfX_nondifferentiable_finite_set by simp lemma l_differentiable_borel_set[measurable, simp]: "{x. l differentiable at x} \ sets borel" using differentiable_ccdfX_l ccdfX_differentiable_borel_set by simp lemma l_differentiable_AE: "AE x in lborel. l differentiable at x" using differentiable_ccdfX_l ccdfX_differentiable_AE by simp lemma deriv_l_measurable[measurable]: "deriv l \ borel_measurable borel" proof - let ?S = "{x. \ l differentiable at x}" have "\x. x \ ?S \ $l_0 * deriv (ccdf (distr \ borel X)) x = deriv l x" using deriv_ccdfX_l by simp thus ?thesis apply - by (rule measurable_discrete_difference [where X="?S" and f="\x. $l_0 * deriv (ccdf (distr \ borel X)) x"]) (simp_all add: countable_finite) qed lemma pdfX_l_normal: "pdfX x = (if l differentiable at x then - deriv l x / $l_0 else 0)" for x::real unfolding pdfX_def using differentiable_eq_field_differentiable_real differentiable_cdfX_l deriv_cdfX_l by simp lemma mu_deriv_l: "$\_x = - deriv l x / $l_x" if "l differentiable at x" for x::real using mu_pdfX that ccdfX_l_normal that pdfX_l_normal by (simp add: differentiable_cdfX_l) lemma mu_nonneg_differentiable_l: "$\_x \ 0" if "l differentiable at x" for x::real using differentiable_cdfX_l mu_nonneg_differentiable that by simp lemma mu_deriv_ln_l: "$\_x = - deriv (\x. ln ($l_x)) x" if "l differentiable at x" "x < $\" for x::real proof - have "\\<^sub>F x in nhds x. ln ($l_x / $l_0) = ln ($l_x) - ln ($l_0)" proof (cases \$\ < \\) case True thus ?thesis apply (rewrite eventually_nhds_metric) apply (intro exI[of _ "real_of_ereal $\ - x"], auto) using that True not_inftyI apply fastforce apply (rewrite ln_div, simp_all) using lx_pos dist_real_def not_inftyI that(2) by fastforce next case False hence "\x. $l_x > 0" using l_0_equiv by force thus ?thesis by (intro always_eventually, rewrite ln_div; simp) qed hence "deriv (\x. ln ($l_x / $l_0)) x = deriv (\x. ln ($l_x)) x" apply (rewrite deriv_cong_ev[of _ "\x. ln ($l_x) - ln ($l_0)"], simp_all) apply (rewrite deriv_diff, simp_all) - unfolding field_differentiable_def using has_derivative_ln that - by (metis DERIV_deriv_iff_real_differentiable differentiable_def lx_pos) + unfolding field_differentiable_def + by (metis that DERIV_ln_divide_chain lx_pos real_differentiableE) thus ?thesis using ccdfX_l_normal mu_deriv_ln that differentiable_ccdfX_l by force qed lemma deriv_l_shift: "deriv l (x+t) = deriv (\t. $l_(x+t)) t" if "l differentiable at (x+t)" for x t :: real using deriv_shift differentiable_eq_field_differentiable_real that by simp context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma p_mu_l: "$p_{t&x} * $\_(x+t) = - deriv l (x+t) / $l_x" if "l differentiable at (x+t)" "t > 0" "x+t < $\" for t::real using p_l mu_deriv_l that by simp lemma p_mu_l_AE: "AE s in lborel. 0 < s \ x+s < $\ \ $p_{s&x} * $\_(x+s) = - deriv l (x+s) / $l_x" proof - have "AE s in lborel. l differentiable at (x+s)" apply (rule AE_borel_affine[of 1 "\u. l differentiable at u" x, simplified]) unfolding pred_def using l_differentiable_AE by simp_all moreover have "AE s in lborel. l differentiable at (x+s) \ 0 < s \ x+s < $\ \ $p_{s&x} * $\_(x+s) = - deriv l (x+s) / $l_x" using p_mu_l by (intro AE_I2) simp ultimately show ?thesis by (rule AE_mp) qed lemma LBINT_l_mu_q: "(LBINT s:{f<..f+t}. $l_(x+s) * $\_(x+s)) / $l_x = $q_{f\t&x}" if "t \ 0" "f \ 0" proof - have "\s. s\{f<..f+t} \ $p_{s&x} = $l_(x+s) / $l_x" using p_l that by simp - hence "$q_{f\t&x} = LBINT s:{f<..f+t}. $l_(x+s) / $l_x * $\_(x+s)" + hence "$q_{f\t&x} = (LBINT s:{f<..f+t}. $l_(x+s) / $l_x * $\_(x+s))" using LBINT_p_mu_q by (smt (verit) greaterThanAtMost_borel set_lebesgue_integral_cong sets_lborel that x_lt_psi) also have "\ = (LBINT s:{f<..f+t}. $l_(x+s) * $\_(x+s)) / $l_x" using set_integral_divide_zero by simp finally show ?thesis by simp qed lemma set_integrable_l_mu: "set_integrable lborel {f<..f+t} (\s. $l_(x+s) * $\_(x+s))" if "t \ 0" "f \ 0" proof - have "set_integrable lborel {f<..f+t} (\s. $l_(x+s) * $\_(x+s) / $l_x)" using p_l set_integrable_p_mu that by (rewrite set_integrable_cong[where f'="\s. $p_{s&x} * $\_(x+s)"]) simp_all+ thus ?thesis by simp qed lemma l_mu_has_integral_q_defer: "((\s. $l_(x+s) * $\_(x+s) / $l_x) has_integral $q_{f\t&x}) {f..f+t}" if "t \ 0" "f \ 0" using p_l that p_mu_has_integral_q_defer_Icc by (rewrite has_integral_cong[of _ _ "\s. $p_{s&x} * $\_(x+s)"]; simp) corollary l_mu_has_integral_q: "((\s. $l_(x+s) * $\_(x+s) / $l_x) has_integral $q_{t&x}) {0..t}" if "t \ 0" using l_mu_has_integral_q_defer[where f=0] that by simp lemma l_mu_has_integral_d: "((\s. $l_(x+s) * $\_(x+s)) has_integral $d_{t & x+f}) {f..f+t}" if "t \ 0" "f \ 0" proof - have "((\s. $l_x * ($p_{s&x} * $\_(x+s))) has_integral $l_x * $q_{f\t&x}) {f..f+t}" apply (rule has_integral_mult_right) by (rule p_mu_has_integral_q_defer_Icc; simp add: that) thus ?thesis using that apply (rewrite in asm q_defer_d_l, simp_all) apply (rewrite has_integral_cong[where g="\s. $l_x * ($p_{s&x} * $\_(x + s))"]) by (rewrite p_l; simp) qed corollary l_mu_has_integral_d_1: "((\s. $l_(x+s) * $\_(x+s)) has_integral $d_(x+f)) {f..f+1}" if "t \ 0" "f \ 0" using l_mu_has_integral_d[where t=1] that by simp lemma e_LBINT_l: "$e`\_x = (LBINT s:{0..}. $l_(x+s) * $\_(x+s) * s) / $l_x" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - have "\s. s\{0..} \ $p_{s&x} = $l_(x+s) / $l_x" using p_l by simp - hence "$e`\_x = LBINT s:{0..}. $l_(x+s) / $l_x * $\_(x+s) * s" + hence "$e`\_x = (LBINT s:{0..}. $l_(x+s) / $l_x * $\_(x+s) * s)" using e_LBINT_p_mu by (smt (verit) atLeast_borel set_lebesgue_integral_cong sets_lborel x_lt_psi) also have "\ = (LBINT s:{0..}. $l_(x+s) * $\_(x+s) * s) / $l_x" using set_integral_divide_zero by simp finally show ?thesis . qed lemma e_integral_l: "$e`\_x = integral {0..} (\s. $l_(x+s) * $\_(x+s) * s) / $l_x" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - have "AE s in lborel. $\_(x+s) \ 0" by (rule AE_translation, rule mu_nonneg_AE) - hence "LBINT s:{0..}. $l_(x+s) * $\_(x+s) * s = integral {0..} (\s. $l_(x+s) * $\_(x+s) * s)" + hence "(LBINT s:{0..}. $l_(x+s) * $\_(x+s) * s) = integral {0..} (\s. $l_(x+s) * $\_(x+s) * s)" by (intro set_borel_integral_eq_integral_nonneg_AE; force) thus ?thesis using e_LBINT_l by simp qed end lemma deriv_x_p_mu_l: "deriv (\y. $p_{t&y}) x = $p_{t&x} * ($\_x - $\_(x+t))" if "l differentiable at x" "l differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real using deriv_x_p_mu that differentiable_ccdfX_l by blast lemma e_has_derivative_mu_e_l: "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. set_integrable lborel {x..} l" "l differentiable at x" "x\{a<.. $\" for a b x :: real using e_has_derivative_mu_e that differentiable_ccdfX_l set_integrable_ccdfX_l by force corollary e_has_derivative_mu_e_l': "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. l integrable_on {x..}" "l differentiable at x" "x\{a<.. $\" for a b x :: real using that apply (intro e_has_derivative_mu_e_l[where a=a], simp_all) using l_nonneg by (rewrite integrable_on_iff_set_integrable_nonneg; simp) context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma curt_e_sum_l_smooth: "$e_x = (\k. $l_(x+k+1)) / $l_x" if "summable (\k. $l_(x+k+1))" proof - have "summable (\k. $p_{k+1&x})" using that by (rewrite p_l; simp add: add.assoc) hence "$e_x = (\k. $p_{k+1&x})" using curt_e_sum_p_smooth by simp also have "\ = (\k. $l_(x+k+1) / $l_x)" by (rewrite p_l; simp add: add.assoc) also have "\ = (\k. $l_(x+k+1)) / $l_x" by (simp add: suminf_divide that) finally show ?thesis . qed lemma curt_e_sum_l_finite_smooth: "$e_x = (\k $\" for n::nat apply (rewrite curt_e_sum_p_finite_smooth[of x n], simp_all add: that) apply (rewrite p_l, simp_all) by (smt (verit) sum.cong sum_divide_distrib) lemma temp_curt_e_sum_l_smooth: "$e_{x:n} = (\kFinite Life Table\ locale finite_life_table = life_table + assumes l_finite: "\x::real. $l_x = 0" begin lemma finite_survival_function_MM_X: "finite_survival_function \ X" proof (rule finite_survival_function.intro) show "survival_model \ X" by (rule survival_model_MM_X) show "finite_survival_function_axioms \ X" unfolding finite_survival_function_axioms_def using l_finite death_pt_def by fastforce qed end sublocale finite_life_table \ finite_survival_function \ X by (rule finite_survival_function_MM_X) context finite_life_table begin notation ult_age ("$\") lemma l_omega_0: "$l_$\ = 0" using ccdfX_l_normal ccdfX_omega_0 by simp lemma l_0_equiv_nat: "$l_x = 0 \ x \ $\" for x::nat using ccdfX_l_normal ccdfX_0_equiv_nat by simp lemma d_l_equiv_nat: "$d_{t&x} = $l_x \ x+t \ $\" for x t :: nat by (metis d_l_equiv of_nat_add psi_le_iff_omega_le) corollary d_1_omega_l: "$d_($\-1) = $l_($\-1)" using d_l_equiv_nat[of 1 "$\-1"] omega_pos by simp context fixes x::nat assumes x_lt_omega[simp]: "x < $\" begin lemma curt_e_sum_l_finite_nat: "$e_x = (\kk::nat. k < n \ isCont l (x+k+1)" "x+n \ $\" for n::nat apply (rewrite curt_e_sum_l_finite[of x n], simp) using that le_ereal_less psi_le_omega apply (metis of_nat_1 of_nat_add, force) by (simp add: add.commute) end end end diff --git a/thys/Actuarial_Mathematics/Preliminaries.thy b/thys/Actuarial_Mathematics/Preliminaries.thy --- a/thys/Actuarial_Mathematics/Preliminaries.thy +++ b/thys/Actuarial_Mathematics/Preliminaries.thy @@ -1,2571 +1,2567 @@ theory Preliminaries imports "HOL-Library.Rewrite" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Extended_Real" "HOL-Probability.Probability" begin notation powr (infixr ".^" 80) section \Preliminary Lemmas\ lemma Collect_conj_eq2: "{x \ A. P x \ Q x} = {x \ A. P x} \ {x \ A. Q x}" by blast lemma vimage_compl_atMost: fixes f :: "'a \ 'b::linorder" shows "-(f -` {..y}) = f -` {y<..}" by fastforce lemma Ico_nat_disjoint: "disjoint_family (\n::nat. {real n ..< real n + 1})" proof - { fix m n :: nat assume "{real m ..< real m + 1} \ {real n ..< real n + 1} \ {}" then obtain x::real where "x \ {real m ..< real m + 1} \ {real n ..< real n + 1}" by force hence "m = n" by simp } thus ?thesis unfolding disjoint_family_on_def by blast qed lemma Ico_nat_union: "(\n::nat. {real n ..< real n + 1}) = {0..}" proof show "(\n::nat. {real n ..< real n + 1}) \ {0..}" by auto next show "{0..} \ (\n::nat. {real n ..< real n + 1})" proof fix x::real assume "x \ {0..}" hence "nat \x\ \ x \ x < nat \x\ + 1" by (metis add_le_same_cancel1 atLeast_iff le_nat_floor less_one linorder_not_le of_nat_floor) thus "x \ (\n::nat. {real n ..< real n + 1})" unfolding atLeastLessThan_def by force qed qed lemma Ico_nat_union_finite: "(\(n::nat)(n::nat) {0.. (\(n::nat): "x \ {0..x\ < m" using of_nat_floor by fastforce moreover with \ have "nat \x\ \ x \ x < nat \x\ + 1" by (metis Nat.add_0_right atLeastLessThan_iff le_nat_floor less_one linorder_not_le nat_add_left_cancel_le of_nat_floor) ultimately show "x \ (\(n::nat) 0" defines "A \ \i::nat. {i*m ..< (i+1)*m}" shows "\i j. i \ j \ A i \ A j = {}" and "(\i j \ A i \ A j = {}" proof (erule contrapos_np) assume "A i \ A j \ {}" then obtain k where "k \ A i \ A j" by blast hence "i*m < (j+1)*m \ j*m < (i+1)*m" unfolding A_def by force hence "i < j+1 \ j < i+1" using mult_less_cancel2 by blast thus "i = j" by force qed } thus "\i j. i \ j \ A i \ A j = {}" by blast next show "(\ii {..< n*m}" proof fix x::nat assume "x \ (\i n" by linarith hence "x < n*m" by (meson less_le_trans mult_le_cancel2 i_x) thus "x \ {..< n*m}" using diff_mult_distrib mult_1 i_n by auto qed next show "{..< n*m} \ (\i {..< n*m}" hence "?i < n" by (simp add: less_mult_imp_div_less) moreover have "?i*m \ x \ x < (?i+1)*m" using assms div_times_less_eq_dividend dividend_less_div_times by auto ultimately show "x \ (\i b" for a b :: real unfolding frontier_def using that by force lemma(in field) divide_mult_cancel[simp]: fixes a b assumes "b \ 0" shows "a / b * b = a" by (simp add: assms) lemma inverse_powr: "(1/a).^b = a.^-b" if "a > 0" for a b :: real by (smt that powr_divide powr_minus_divide powr_one_eq_one) lemma powr_eq_one_iff_gen[simp]: "a.^x = 1 \ x = 0" if "a > 0" "a \ 1" for a x :: real by (metis powr_eq_0_iff powr_inj powr_zero_eq_one that) lemma powr_less_cancel2: "0 < a \ 0 < x \ 0 < y \ x.^a < y.^a \ x < y" for a x y ::real proof - assume a_pos: "0 < a" and x_pos: "0 < x" and y_pos: "0 < y" show "x.^a < y.^a \ x < y" proof (erule contrapos_pp) assume "\ x < y" hence "x \ y" by fastforce hence "x.^a \ y.^a" proof (cases "x = y") case True thus ?thesis by simp next case False hence "x.^a > y.^a" using \x \ y\ powr_less_mono2 a_pos y_pos by auto thus ?thesis by auto qed thus "\ x.^a < y.^a" by fastforce qed qed lemma geometric_increasing_sum_aux: "(1-r)^2 * (\kk 1" for n::nat and r::real by (subst geometric_increasing_sum_aux[THEN sym], simp add: that) lemma Reals_UNIV[simp]: "\ = {x::real. True}" unfolding Reals_def by auto lemma Lim_cong: assumes "\\<^sub>F x in F. f x = g x" shows "Lim F f = Lim F g" unfolding t2_space_class.Lim_def using tendsto_cong assms by fastforce lemma antimono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ antimono_on A f" by (rule monotone_onI) lemma antimono_onD: "\antimono_on A f; r \ A; s \ A; r \ s\ \ f r \ f s" by (rule monotone_onD) lemma antimono_imp_mono_on: "antimono f \ antimono_on A f" by (simp add: antimonoD antimono_onI) lemma antimono_on_subset: "antimono_on A f \ B \ A \ antimono_on B f" by (rule monotone_on_subset) lemma mono_on_antimono_on: fixes f :: "'a::order \ 'b::ordered_ab_group_add" shows "mono_on A f \ antimono_on A (\r. - f r)" by (simp add: monotone_on_def) corollary mono_antimono: fixes f :: "'a::order \ 'b::ordered_ab_group_add" shows "mono f \ antimono (\r. - f r)" by (rule mono_on_antimono_on) lemma mono_on_at_top_le: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_mono: "mono_on {a..} f" and f_to_l: "(f \ l) at_top" shows "\x. x \ {a..} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "f b - l" have lim_top: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_top" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_top" using fb_l by (intro lim_top; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_top_linorder by metis let ?n = "max b N" have "f ?n < f b" using fn_in by force moreover have "f ?n \ f b" using f_mono b_a by (simp add: le_max_iff_disj mono_on_def) ultimately have False by simp } thus "\x\{a..}. f x \ l" - apply - - by (rule notnotD, rewrite Set.ball_simps) auto + by blast qed corollary mono_at_top_le: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "mono f" and "(f \ b) at_top" shows "\x. f x \ b" using mono_on_at_top_le assms by (metis atLeast_iff mono_imp_mono_on nle_le) lemma mono_on_at_bot_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_mono: "mono_on {..a} f" and f_to_l: "(f \ l) at_bot" shows "\x. x \ {..a} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "l - f b" have lim_bot: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_bot" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_bot" using fb_l by (intro lim_bot; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_bot_linorder by metis let ?n = "min b N" have "f ?n > f b" using fn_in by force moreover have "f ?n \ f b" using f_mono b_a by (simp add: min.coboundedI1 mono_onD) ultimately have False by simp } thus "\x\{..a}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary mono_at_bot_ge: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "mono f" and "(f \ b) at_bot" shows "\x. f x \ b" using mono_on_at_bot_ge assms by (metis atMost_iff mono_imp_mono_on nle_le) lemma antimono_on_at_top_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_antimono: "antimono_on {a..} f" and f_to_l: "(f \ l) at_top" shows "\x. x \ {a..} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "l - f b" have lim_top: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_top" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_top" using fb_l by (intro lim_top; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_top_linorder by metis let ?n = "max b N" have "f ?n > f b" using fn_in by force moreover have "f ?n \ f b" using f_antimono b_a by (simp add: antimono_onD le_max_iff_disj) ultimately have False by simp } thus "\x\{a..}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary antimono_at_top_le: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "antimono f" and "(f \ b) at_top" shows "\x. f x \ b" using antimono_on_at_top_ge assms antimono_imp_mono_on by blast lemma antimono_on_at_bot_ge: fixes a :: "'a::linorder" and b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a \ 'b" assumes f_antimono: "antimono_on {..a} f" and f_to_l: "(f \ l) at_bot" shows "\x. x \ {..a} \ f x \ l" proof (unfold atomize_ball) { fix b assume b_a: "b \ a" and fb_l: "\ f b \ l" let ?eps = "f b - l" have lim_bot: "\S. open S \ l \ S \ eventually (\x. f x \ S) at_bot" using assms tendsto_def by auto have "eventually (\x. f x \ {l - ?eps <..< l + ?eps}) at_bot" using fb_l by (intro lim_bot; force) then obtain N where fn_in: "\n. n \ N \ f n \ {l - ?eps <..< l + ?eps}" using eventually_at_bot_linorder by metis let ?n = "min b N" have "f ?n < f b" using fn_in by force moreover have "f ?n \ f b" using f_antimono b_a by (simp add: min.coboundedI1 antimono_onD) ultimately have False by simp } thus "\x\{..a}. f x \ l" apply - by (rule notnotD, rewrite Set.ball_simps) auto qed corollary antimono_at_bot_ge: fixes b :: "'b::{order_topology, linordered_ab_group_add}" and f :: "'a::linorder \ 'b" assumes "antimono f" and "(f \ b) at_bot" shows "\x. f x \ b" using antimono_on_at_bot_ge assms antimono_imp_mono_on by blast lemma continuous_cdivide: fixes c::"'a::real_normed_field" assumes "c \ 0" "continuous F f" shows "continuous F (\x. f x / c)" using assms continuous_mult_right by (rewrite field_class.field_divide_inverse) auto lemma continuous_mult_left_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. c * f x)" using continuous_mult_left continuous_cdivide assms by force lemma continuous_mult_right_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. f x * c)" using continuous_mult_right continuous_cdivide assms by force lemma continuous_cdivide_iff: fixes c::"'a::real_normed_field" assumes "c \ 0" shows "continuous F f \ continuous F (\x. f x / c)" proof assume "continuous F f" thus "continuous F (\x. f x / c)" by (intro continuous_cdivide) (simp add: assms) next assume "continuous F (\x. f x / c)" hence "continuous F (\x. f x / c * c)" using continuous_mult_right by fastforce thus "continuous F f" using assms by force qed lemma continuous_cong: assumes "eventually (\x. f x = g x) F" "f (Lim F (\x. x)) = g (Lim F (\x. x))" shows "continuous F f \ continuous F g" unfolding continuous_def using assms filterlim_cong by force lemma continuous_at_within_cong: assumes "f x = g x" "eventually (\x. f x = g x) (at x within s)" shows "continuous (at x within s) f \ continuous (at x within s) g" proof (cases \x \ closure (s - {x})\) case True thus ?thesis using assms apply (intro continuous_cong, simp) by (rewrite Lim_ident_at, simp add: at_within_eq_bot_iff)+ simp next case False hence "at x within s = bot" using not_trivial_limit_within by blast thus ?thesis by simp qed lemma DERIV_cmult_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" proof assume "(f has_field_derivative D) (at x within s)" thus "((\x. c * f x) has_field_derivative c * D) (at x within s)" using DERIV_cmult by force next assume "((\x. c * f x) has_field_derivative c * D) (at x within s)" hence "((\x. c * f x / c) has_field_derivative c * D / c) (at x within s)" using DERIV_cdivide assms by blast thus "(f has_field_derivative D) (at x within s)" by (simp add: assms field_simps) qed lemma DERIV_cmult_right_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" by (rewrite DERIV_cmult_iff[of c], simp_all add: assms mult_ac) lemma DERIV_cdivide_iff: assumes "c \ 0" shows "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" apply (rewrite field_class.field_divide_inverse)+ using DERIV_cmult_right_iff assms inverse_nonzero_iff_nonzero by blast lemma DERIV_ln_divide_chain: fixes f :: "real \ real" assumes "f x > 0" and "(f has_real_derivative D) (at x within s)" shows "((\x. ln (f x)) has_real_derivative (D / f x)) (at x within s)" proof - have "DERIV ln (f x) :> 1 / f x" using assms by (intro DERIV_ln_divide) blast thus ?thesis using DERIV_chain2 assms by fastforce qed lemma inverse_fun_has_integral_ln: fixes f :: "real \ real" and f' :: "real \ real" assumes "a \ b" and "\x. x\{a..b} \ f x > 0" and "continuous_on {a..b} f" and "\x. x\{a<.. (f has_real_derivative f' x) (at x)" shows "((\x. f' x / f x) has_integral (ln (f b) - ln (f a))) {a..b}" proof - have "continuous_on {a..b} (\x. ln (f x))" using assms by (intro continuous_intros; force) moreover have "\x. x\{a<.. ((\x. ln (f x)) has_vector_derivative f' x / f x) (at x)" apply (rewrite has_real_derivative_iff_has_vector_derivative[THEN sym]) using assms by (intro DERIV_ln_divide_chain; simp) ultimately show ?thesis using assms by (intro fundamental_theorem_of_calculus_interior; simp) qed lemma DERIV_fun_powr2: fixes a::real assumes a_pos: "a > 0" and f: "DERIV f x :> r" shows "DERIV (\x. a.^(f x)) x :> a.^(f x) * r * ln a" proof - let ?g = "(\x. a)" have g: "DERIV ?g x :> 0" by simp have pos: "?g x > 0" by (simp add: a_pos) show ?thesis using DERIV_powr[OF g pos f] a_pos by (auto simp add: field_simps) qed lemma has_real_derivative_powr2: assumes a_pos: "a > 0" shows "((\x. a.^x) has_real_derivative a.^x * ln a) (at x)" proof - let ?f = "(\x. x::real)" have f: "DERIV ?f x :> 1" by simp thus ?thesis using DERIV_fun_powr2[OF a_pos f] by simp qed lemma has_integral_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "((\x. a.^x) has_integral ((a.^c - 1) / (ln a))) {0..c}" proof - have "((\x. a.^x) has_integral ((a.^c)/(ln a) - (a.^0)/(ln a))) {0..c}" proof (rule fundamental_theorem_of_calculus[OF c_nneg]) fix x::real assume "x \ {0..c}" show "((\y. a.^y / ln a) has_vector_derivative a.^x) (at x within {0..c})" apply (insert has_real_derivative_powr2[OF a_pos, of x]) apply (drule DERIV_cdivide[where c = "ln a"], simp add: assms) apply (rule has_vector_derivative_within_subset[where S=UNIV and T="{0..c}"], auto) by (rule iffD1[OF has_real_derivative_iff_has_vector_derivative]) qed thus ?thesis using assms powr_zero_eq_one by (simp add: field_simps) qed lemma integrable_on_powr2_from_0: fixes a c :: real assumes a_pos: "a > 0" and a_neq_1: "a \ 1" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" using has_integral_powr2_from_0[OF assms] unfolding integrable_on_def by blast lemma integrable_on_powr2_from_0_general: fixes a c :: real assumes a_pos: "a > 0" and c_nneg: "c \ 0" shows "(\x. a.^x) integrable_on {0..c}" proof (cases "a = 1") case True thus ?thesis using has_integral_const_real by auto next case False thus ?thesis using has_integral_powr2_from_0 False assms by auto qed lemma has_integral_null_interval: fixes a b :: real and f::"real \ real" assumes "a \ b" shows "(f has_integral 0) {a..b}" using assms content_real_eq_0 by blast lemma has_integral_interval_reverse: fixes f :: "real \ real" and a b :: real assumes "a \ b" and "continuous_on {a..b} f" shows "((\x. f (a+b-x)) has_integral (integral {a..b} f)) {a..b}" proof - let ?g = "\x. a + b - x" let ?g' = "\x. -1" have g_C0: "continuous_on {a..b} ?g" using continuous_on_op_minus by simp have Dg_g': "\x. x\{a..b} \ (?g has_field_derivative ?g' x) (at x within {a..b})" by (auto intro!: derivative_eq_intros) show ?thesis using has_integral_substitution_general [of "{}" a b ?g a b f, simplified, OF assms g_C0 Dg_g', simplified] apply (simp add: has_integral_null_interval[OF assms(1), THEN integral_unique]) by (simp add: has_integral_neg_iff) qed section \Additional Lemmas for the "Analysis" Library\ lemma continuous_within_shift: fixes a x :: "'a :: {topological_ab_group_add, t2_space}" and s :: "'a set" and f :: "'a \ 'b::topological_space" shows "continuous (at x within s) (\x. f (x+a)) \ continuous (at (x+a) within plus a ` s) f" proof assume "continuous (at x within s) (\x. f (x+a))" moreover have "continuous (at (x+a) within plus a ` s) (plus (-a))" by (simp add: continuous_at_imp_continuous_at_within) moreover have "plus (-a) ` plus a ` s = s" by force ultimately show "continuous (at (x+a) within plus a ` s) f" using continuous_within_compose unfolding comp_def by force next assume "continuous (at (x+a) within plus a ` s) f" moreover have "continuous (at x within s) (plus a)" by (simp add: continuous_at_imp_continuous_at_within) ultimately show "continuous (at x within s) (\x. f (x+a))" using continuous_within_compose unfolding comp_def by (force simp add: add.commute) qed lemma isCont_shift: fixes a x :: "'a :: {topological_ab_group_add, t2_space}" and f :: "'a \ 'b::topological_space" shows "isCont (\x. f (x+a)) x \ isCont f (x+a)" using continuous_within_shift by force lemma differentiable_eq_field_differentiable_real: fixes f :: "real \ real" shows "f differentiable F \ f field_differentiable F" unfolding field_differentiable_def differentiable_def has_real_derivative using has_real_derivative_iff by presburger lemma differentiable_on_eq_field_differentiable_real: fixes f :: "real \ real" shows "f differentiable_on s \ (\x\s. f field_differentiable (at x within s))" unfolding differentiable_on_def using differentiable_eq_field_differentiable_real by simp lemma differentiable_on_cong : assumes "\x. x\s \ f x = g x" and "f differentiable_on s" shows "g differentiable_on s" proof - { fix x assume "x\s" hence "f differentiable at x within s" using assms unfolding differentiable_on_def by simp from this obtain D where "(f has_derivative D) (at x within s)" unfolding differentiable_def by blast hence "(g has_derivative D) (at x within s)" using has_derivative_transform assms \x\s\ by metis hence "g differentiable at x within s" unfolding differentiable_def by blast } hence "\x\s. g differentiable at x within s" by simp thus ?thesis unfolding differentiable_on_def by simp qed lemma C1_differentiable_imp_deriv_continuous_on: "f C1_differentiable_on S \ continuous_on S (deriv f)" using C1_differentiable_on_eq field_derivative_eq_vector_derivative by auto lemma deriv_shift: assumes "f field_differentiable at (x+a)" shows "deriv f (x+a) = deriv (\s. f (x+s)) a" proof - have "(f has_field_derivative deriv f (x+a)) (at (x+a))" using DERIV_deriv_iff_field_differentiable assms by force hence "((\s. f (x+s)) has_field_derivative deriv f (x+a)) (at a)" using DERIV_at_within_shift has_field_derivative_at_within by blast moreover have "((\s. f (x+s)) has_field_derivative deriv (\s. f (x+s)) a) (at a)" using DERIV_imp_deriv calculation by fastforce ultimately show ?thesis using DERIV_unique by force qed lemma piecewise_differentiable_on_cong: assumes "f piecewise_differentiable_on i" and "\x. x \ i \ f x = g x" shows "g piecewise_differentiable_on i" proof - have "continuous_on i g" using continuous_on_cong_simp assms piecewise_differentiable_on_imp_continuous_on by force moreover have "\S. finite S \ (\x \ i - S. g differentiable (at x within i))" proof - from assms piecewise_differentiable_on_def obtain S where fin: "finite S" and "\x \ i - S. f differentiable (at x within i)" by metis hence "\x. x \ i - S \ f differentiable (at x within i)" by simp hence "\x. x \ i - S \ g differentiable (at x within i)" using has_derivative_transform assms by (metis DiffD1 differentiable_def) with fin show ?thesis by blast qed ultimately show ?thesis unfolding piecewise_differentiable_on_def by simp qed lemma differentiable_piecewise: assumes "continuous_on i f" and "f differentiable_on i" shows "f piecewise_differentiable_on i" unfolding piecewise_differentiable_on_def using assms differentiable_onD by auto lemma piecewise_differentiable_scaleR: assumes "f piecewise_differentiable_on S" shows "(\x. a *\<^sub>R f x) piecewise_differentiable_on S" proof - from assms obtain T where fin: "finite T" "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast hence "\x. x \ S - T \ (\x. a *\<^sub>R f x) differentiable at x within S" using differentiable_scaleR by simp moreover have "continuous_on S (\x. a *\<^sub>R f x)" using assms continuous_on_scaleR continuous_on_const piecewise_differentiable_on_def by blast ultimately show "(\x. a *\<^sub>R f x) piecewise_differentiable_on S" using fin piecewise_differentiable_on_def by blast qed lemma differentiable_on_piecewise_compose: assumes "f piecewise_differentiable_on S" and "g differentiable_on f ` S" shows "g \ f piecewise_differentiable_on S" proof - from assms obtain T where fin: "finite T" "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast hence "\x. x \ S - T \ g \ f differentiable at x within S" by (meson DiffD1 assms differentiable_chain_within differentiable_onD image_eqI) hence "\T. finite T \ (\x\S-T. g \ f differentiable at x within S)" using fin by blast moreover have "continuous_on S (g \ f)" using assms continuous_on_compose differentiable_imp_continuous_on unfolding piecewise_differentiable_on_def by blast ultimately show ?thesis unfolding piecewise_differentiable_on_def by force qed lemma MVT_order_free: fixes r h :: real defines "I \ {r..r+h} \ {r+h..r}" assumes "continuous_on I f" and "f differentiable_on interior I" obtains t where "t \ {0<..<1}" and "f (r+h) - f r = h * deriv f (r + t*h)" proof - consider (h_pos) "h > 0" | (h_0) "h = 0" | (h_neg) "h < 0" by force thus ?thesis proof cases case h_pos hence "r < r+h" "I = {r..r+h}" unfolding I_def by simp_all moreover hence "interior I = {r<..x. \r < x; x < r+h\ \ f differentiable (at x)" using assms by (simp add: differentiable_on_eq_differentiable_at) ultimately obtain z where "r < z \ z < r+h \ f (r+h) - f r = h * deriv f z" using MVT assms by (smt (verit) DERIV_imp_deriv) moreover hence "(z-r) / h \ {0<..<1}" by simp moreover have "z = r + (z-r)/h * h" using h_pos by simp ultimately show ?thesis using that by presburger next case h_0 have "1/2 \ {0::real<..<1}" by simp moreover have "f (r+h) - f r = 0" using h_0 by simp moreover have "h * deriv f (r + (1/2)*h) = 0" using h_0 by simp ultimately show ?thesis using that by presburger next case h_neg hence "r+h < r" "I = {r+h..r}" unfolding I_def by simp_all moreover hence "interior I = {r+h<..x. \r+h < x; x < r\ \ f differentiable (at x)" using assms by (simp add: differentiable_on_eq_differentiable_at) ultimately obtain z where "r+h < z \ z < r \ f r - f (r+h) = -h * deriv f z" using MVT assms by (smt (verit) DERIV_imp_deriv) moreover hence "(z-r) / h \ {0<..<1}" by (simp add: neg_less_divide_eq) moreover have "z = r + (z-r)/h * h" using h_neg by simp ultimately show ?thesis using that mult_minus_left by fastforce qed qed lemma integral_combine2: fixes f :: "real \ 'a::banach" assumes "a \ c" "c \ b" and "f integrable_on {a..c}" "f integrable_on {c..b}" shows "integral {a..c} f + integral {c..b} f = integral {a..b} f" apply (rule integral_unique[THEN sym]) apply (rule has_integral_combine[of a c b], simp_all add: assms) using has_integral_integral assms by auto lemma FTC_real_deriv_has_integral: fixes F :: "real \ real" assumes "a \ b" and "F piecewise_differentiable_on {a<..x. x \ {a<.. F differentiable at x within {a<..x. x \ {a<.. (F has_real_derivative deriv F x) (at x)" proof - fix x assume x_in: "x \ {a<..x. x \ T - S \ g x = f x" shows "f integrable_on T \ g integrable_on T" using integrable_spike assms by force lemma set_borel_measurable_lborel: "set_borel_measurable lborel A f \ set_borel_measurable borel A f" unfolding set_borel_measurable_def by simp lemma restrict_space_whole[simp]: "restrict_space M (space M) = M" unfolding restrict_space_def by (simp add: measure_of_of_measure) lemma deriv_measurable_real: fixes f :: "real \ real" assumes "f differentiable_on S" "open S" "f \ borel_measurable borel" shows "set_borel_measurable borel S (deriv f)" proof - have "\x. x \ S \ deriv f x = lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))" proof - fix x assume x_S: "x \ S" hence "f field_differentiable (at x within S)" using differentiable_on_eq_field_differentiable_real assms by simp hence "(f has_field_derivative deriv f x) (at x)" using assms DERIV_deriv_iff_field_differentiable x_S at_within_open by force hence "(\h. (f (x+h) - f x) / h) \0\ deriv f x" using DERIV_def by auto hence "\d. (\i. d i \ UNIV-{0::real}) \ d \ 0 \ ((\h. (f (x+h) - f x) / h) \ d) \ deriv f x" using tendsto_at_iff_sequentially by blast moreover have "\i. 1 / Suc i \ UNIV - {0::real}" by simp moreover have "(\i. 1 / Suc i) \ 0" using LIMSEQ_Suc lim_const_over_n by blast ultimately have "((\h. (f (x + h) - f x) / h) \ (\i. 1 / Suc i)) \ deriv f x" by auto thus "deriv f x = lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))" unfolding comp_def by (simp add: limI) qed moreover have "(\x. indicator S x * lim (\i. (f (x + 1 / Suc i) - f x) / (1 / Suc i))) \ borel_measurable borel" using assms by (measurable, simp, measurable) ultimately show ?thesis unfolding set_borel_measurable_def measurable_cong by simp (smt (verit) indicator_simps(2) measurable_cong mult_eq_0_iff) qed lemma piecewise_differentiable_on_deriv_measurable_real: fixes f :: "real \ real" assumes "f piecewise_differentiable_on S" "open S" "f \ borel_measurable borel" shows "set_borel_measurable borel S (deriv f)" proof - from assms obtain T where fin: "finite T" and diff: "\x. x \ S - T \ f differentiable at x within S" unfolding piecewise_differentiable_on_def by blast with assms have "open (S - T)" using finite_imp_closed by blast moreover hence "f differentiable_on (S - T)" unfolding differentiable_on_def using assms by (metis Diff_iff at_within_open diff) ultimately have "set_borel_measurable borel (S - T) (deriv f)" by (intro deriv_measurable_real; simp add: assms) thus ?thesis unfolding set_borel_measurable_def apply simp apply (rule measurable_discrete_difference [where X=T and f="\x. indicat_real (S - T) x * deriv f x"], simp_all) using fin uncountable_infinite apply blast by (simp add: indicator_diff) qed lemma borel_measurable_antimono: fixes f :: "real \ real" shows "antimono f \ f \ borel_measurable borel" using borel_measurable_mono by (smt (verit, del_insts) borel_measurable_uminus_eq monotone_on_def) lemma set_borel_measurable_restrict_space_iff: fixes f :: "'a \ 'b::real_normed_vector" assumes \[measurable, simp]: "\ \ space M \ sets M" shows "f \ borel_measurable (restrict_space M \) \ set_borel_measurable M \ f" using assms borel_measurable_restrict_space_iff set_borel_measurable_def by blast lemma set_integrable_restrict_space_iff: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "A \ sets M" shows "set_integrable M A f \ integrable (restrict_space M A) f" unfolding set_integrable_def using assms by (rewrite integrable_restrict_space; simp) lemma set_lebesgue_integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "A \ sets M" shows "set_lebesgue_integral M A f = integral\<^sup>L (restrict_space M A) f" unfolding set_lebesgue_integral_def using assms integral_restrict_space by (metis (mono_tags) sets.Int_space_eq2) lemma distr_borel_lborel: "distr M borel f = distr M lborel f" by (metis distr_cong sets_lborel) lemma AE_translation: assumes "AE x in lborel. P x" shows "AE x in lborel. P (a+x)" proof - from assms obtain N where P: "\x. x \ space lborel - N \ P x" and null: "N \ null_sets lborel" using AE_E3 by blast hence "{y. a+y \ N} \ null_sets lborel" using null_sets_translation[of N "-a", simplified] by (simp add: add.commute) moreover have "\y. y \ space lborel - {y. a+y \ N} \ P (a+y)" using P by force ultimately show "AE y in lborel. P (a+y)" by (smt (verit, del_insts) Diff_iff eventually_ae_filter mem_Collect_eq subsetI) qed lemma set_AE_translation: assumes "AE x\S in lborel. P x" shows "AE x \ plus (-a) ` S in lborel. P (a+x)" proof - have "AE x in lborel. a+x \ S \ P (a+x)" using assms by (rule AE_translation) moreover have "\x. a+x \ S \ x \ plus (-a) ` S" by force ultimately show ?thesis by simp qed lemma AE_scale_measure_iff: assumes "r > 0" shows "(AE x in (scale_measure r M). P x) \ (AE x in M. P x)" unfolding ae_filter_def null_sets_def apply (rewrite space_scale_measure, simp) using assms by (smt Collect_cong not_gr_zero) lemma nn_set_integral_cong2: assumes "AE x\A in M. f x = g x" shows "(\\<^sup>+x\A. f x \M) = (\\<^sup>+x\A. g x \M)" proof - { fix x assume "x \ space M" have "(x \ A \ f x = g x) \ f x * indicator A x = g x * indicator A x" by force } hence "AE x in M. (x \ A \ f x = g x) \ f x * indicator A x = g x * indicator A x" by (rule AE_I2) hence "AE x in M. f x * indicator A x = g x * indicator A x" using assms AE_mp by auto thus ?thesis by (rule nn_integral_cong_AE) qed lemma set_lebesgue_integral_cong_AE2: assumes [measurable]: "A \ sets M" "set_borel_measurable M A f" "set_borel_measurable M A g" assumes "AE x \ A in M. f x = g x" - shows "LINT x:A|M. f x = LINT x:A|M. g x" + shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)" proof - let ?fA = "\x. indicator A x *\<^sub>R f x" and ?gA = "\x. indicator A x *\<^sub>R g x" have "?fA \ borel_measurable M" "?gA \ borel_measurable M" using assms unfolding set_borel_measurable_def by simp_all moreover have "AE x \ A in M. ?fA x = ?gA x" using assms by simp - ultimately have "LINT x:A|M. ?fA x = LINT x:A|M. ?gA x" + ultimately have "(LINT x:A|M. ?fA x) = (LINT x:A|M. ?gA x)" by (intro set_lebesgue_integral_cong_AE; simp) - moreover have "LINT x:A|M. f x = LINT x:A|M. ?fA x" "LINT x:A|M. g x = LINT x:A|M. ?gA x" + moreover have "(LINT x:A|M. f x) = (LINT x:A|M. ?fA x)" "(LINT x:A|M. g x) = (LINT x:A|M. ?gA x)" unfolding set_lebesgue_integral_def by (metis indicator_scaleR_eq_if)+ ultimately show ?thesis by simp qed proposition nn_integral_disjoint_family_on_finite: assumes [measurable]: "f \ borel_measurable M" "\(n::nat). n \ S \ B n \ sets M" and "disjoint_family_on B S" "finite S" shows "(\\<^sup>+x \ (\n\S. B n). f x \M) = (\n\S. (\\<^sup>+x \ B n. f x \M))" proof - let ?A = "\n::nat. if n \ S then B n else {}" have "\n::nat. ?A n \ sets M" by simp moreover have "disjoint_family ?A" unfolding disjoint_family_on_def proof - { fix m n :: nat assume "m \ n" hence "(if m \ S then B m else {}) \ (if n \ S then B n else {}) = {}" apply simp using assms unfolding disjoint_family_on_def by blast } thus "\m::nat\UNIV. \n::nat\UNIV. m \ n \ (if m \ S then B m else {}) \ (if n \ S then B n else {}) = {}" by blast qed ultimately have "set_nn_integral M (\ (range ?A)) f = (\n. set_nn_integral M (?A n) f)" by (rewrite nn_integral_disjoint_family; simp) moreover have "set_nn_integral M (\ (range ?A)) f = (\\<^sup>+x \ (\n\S. B n). f x \M)" proof - have "\ (range ?A) = (\n\S. B n)" by simp thus ?thesis by simp qed moreover have "(\n. set_nn_integral M (?A n) f) = (\n\S. set_nn_integral M (B n) f)" by (rewrite suminf_finite[of S]; simp add: assms) ultimately show ?thesis by simp qed lemma nn_integral_distr_set: assumes "T \ measurable M M'" and "f \ borel_measurable (distr M M' T)" and "A \ sets M'" and "\x. x \ space M \ T x \ A" shows "integral\<^sup>N (distr M M' T) f = set_nn_integral (distr M M' T) A f" proof - have "integral\<^sup>N (distr M M' T) f = (\\<^sup>+x\(space M'). f x \(distr M M' T))" by (rewrite nn_set_integral_space[THEN sym], simp) also have "\ = (\\<^sup>+x\A. f x \(distr M M' T))" proof - have [simp]: "sym_diff (space M') A = space M' - A" using assms by (metis Diff_mono sets.sets_into_space sup.orderE) show ?thesis apply (rule nn_integral_null_delta; simp add: assms) unfolding null_sets_def using assms apply (simp, rewrite emeasure_distr; simp) unfolding vimage_def using emeasure_empty by (smt (z3) Collect_empty_eq Diff_iff Int_def mem_Collect_eq) qed finally show ?thesis . qed (* Analogue for "measure_eqI_lessThan" in the "Lebesgue_Measure" Theory *) lemma measure_eqI_Ioc: fixes M N :: "real measure" assumes sets: "sets M = sets borel" "sets N = borel" assumes fin: "\a b. a \ b \ emeasure M {a<..b} < \" assumes eq: "\a b. a \ b \ emeasure M {a<..b} = emeasure N {a<..b}" shows "M = N" proof (rule measure_eqI_generator_eq_countable) let ?Ioc = "\(a::real,b::real). {a<..b}" let ?E = "range ?Ioc" show "Int_stable ?E" using Int_stable_def Int_greaterThanAtMost by force show "?E \ Pow UNIV" "sets M = sigma_sets UNIV ?E" "sets N = sigma_sets UNIV ?E" unfolding sets by (auto simp add: borel_sigma_sets_Ioc) show "\I. I \ ?E \ emeasure M I = emeasure N I" proof - fix I assume "I \ ?E" then obtain a b where "I = {a<..b}" by auto thus "emeasure M I = emeasure N I" by (smt (verit, best) eq greaterThanAtMost_empty) qed show "?Ioc ` (Rats \ Rats) \ ?E" "(\i\(Rats\Rats). ?Ioc i) = UNIV" using Rats_no_bot_less Rats_no_top_le by auto show "countable (?Ioc ` (Rats \ Rats))" using countable_rat by blast show "\I. I \ ?Ioc ` (Rats \ Rats) \ emeasure M I \ \" proof - fix I assume "I \ ?Ioc ` (Rats \ Rats)" then obtain a b where "(a,b) \ (Rats \ Rats)" "I = {a<..b}" by blast thus "emeasure M I \ \" by (smt (verit, best) Ioc_inj fin order.strict_implies_not_eq) qed qed lemma (in finite_measure) distributed_measure: assumes "distributed M N X f" and "\x. x \ space N \ f x \ 0" and "A \ sets N" shows "measure M (X -` A \ space M) = (\x. indicator A x * f x \N)" proof - have [simp]: "(\x. indicat_real A x * f x) \ borel_measurable N" using assms apply (measurable; simp?) using distributed_real_measurable assms by force have "emeasure M (X -` A \ space M) = (\\<^sup>+x\A. ennreal (f x) \N)" by (rule distributed_emeasure; simp add: assms) moreover have "enn2real (\\<^sup>+x\A. ennreal (f x) \N) = \x. indicator A x * f x \N" apply (rewrite enn2real_nn_integral_eq_integral [where f="\x. ennreal (indicator A x * f x)", THEN sym]; (simp add: assms)?) using distributed_emeasure assms by (smt (verit) emeasure_finite indicator_mult_ennreal mult.commute nn_integral_cong top.not_eq_extremum) ultimately show ?thesis using measure_def by metis qed (* Set Integral Version of the Lebesgue's Dominated Convergence Theorem *) lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "A \ sets M" "set_borel_measurable M A f" "\i. set_borel_measurable M A (s i)" "set_integrable M A w" assumes lim: "AE x\A in M. (\i. s i x) \ f x" assumes bound: "\i::nat. AE x\A in M. norm (s i x) \ w x" shows set_integrable_dominated_convergence: "set_integrable M A f" and set_integrable_dominated_convergence2: "\i. set_integrable M A (s i)" and set_integral_dominated_convergence: "(\i. set_lebesgue_integral M A (s i)) \ set_lebesgue_integral M A f" proof - have "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" and "\i. (\x. indicator A x *\<^sub>R s i x) \ borel_measurable M" and "integrable M (\x. indicator A x *\<^sub>R w x)" using assms unfolding set_borel_measurable_def set_integrable_def by simp_all moreover have "AE x in M. (\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof - obtain N where N_null: "N \ null_sets M" and si_f: "\x. x \ space M - N \ x \ A \ (\i. s i x) \ f x" using lim AE_E3 by (smt (verit)) hence "\x. x \ space M - N \ (\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof - fix x assume asm: "x \ space M - N" thus "(\i. indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R f x" proof (cases \x \ A\) case True with si_f asm show ?thesis by simp next case False thus ?thesis by simp qed qed thus ?thesis by (smt (verit) AE_I' DiffI N_null mem_Collect_eq subsetI) qed moreover have "\i. AE x in M. norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" proof - fix i from bound obtain N where N_null: "N \ null_sets M" and "\x. x \ space M - N \ x \ A \ norm (s i x) \ w x" using AE_E3 by (smt (verit)) hence "\x. x \ space M - N \ norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" by (simp add: indicator_scaleR_eq_if) with N_null show "AE x in M. norm (indicator A x *\<^sub>R s i x) \ indicator A x *\<^sub>R w x" by (smt (verit) DiffI eventually_ae_filter mem_Collect_eq subsetI) qed ultimately show "set_integrable M A f" "\i. set_integrable M A (s i)" "(\i. set_lebesgue_integral M A (s i)) \ set_lebesgue_integral M A f" unfolding set_integrable_def set_lebesgue_integral_def by (rule integrable_dominated_convergence, rule integrable_dominated_convergence2, rule integral_dominated_convergence) qed lemma absolutely_integrable_on_iff_set_integrable: fixes f :: "'a::euclidean_space \ real" assumes "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f absolutely_integrable_on S" unfolding set_integrable_def apply (simp, rewrite integrable_completion[THEN sym]) apply measurable using assms by simp_all corollary integrable_on_iff_set_integrable_nonneg: fixes f :: "'a::euclidean_space \ real" assumes "\x. x \ S \ f x \ 0" "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f integrable_on S" using absolutely_integrable_on_iff_set_integrable assms by (metis absolutely_integrable_on_iff_nonneg) lemma integrable_on_iff_set_integrable_nonneg_AE: fixes f :: "'a::euclidean_space \ real" assumes "AE x\S in lborel. f x \ 0" "f \ borel_measurable lborel" and "S \ sets lborel" shows "set_integrable lborel S f \ f integrable_on S" proof - from assms obtain N where nonneg: "\x. x \ S - N \ f x \ 0" and null: "N \ null_sets lborel" by (smt (verit, ccfv_threshold) AE_E3 Diff_iff UNIV_I space_borel space_lborel) let ?g = "\x. if x \ N then 0 else f x" have [simp]: "negligible N" using null negligible_iff_null_sets null_sets_completionI by blast have "N \ sets lborel" using null by auto hence [simp]: "?g \ borel_measurable borel" using assms by force have "set_integrable lborel S f \ set_integrable lborel S ?g" proof - have "AE x\S in lborel. f x = ?g x" by (rule AE_I'[of N], simp_all add: null, blast) thus ?thesis using assms by (intro set_integrable_cong_AE[of f _ ?g S]; simp) qed also have "\ \ ?g integrable_on S" using assms by (intro integrable_on_iff_set_integrable_nonneg; simp add: nonneg) also have "\ \ f integrable_on S" by (rule integrable_spike_cong[of N]; simp) finally show ?thesis . qed lemma set_borel_integral_eq_integral_nonneg: fixes f :: "'a::euclidean_space \ real" assumes "\x. x \ S \ f x \ 0" "f \ borel_measurable borel" "S \ sets borel" - shows "LINT x : S | lborel. f x = integral S f" + shows "(LINT x : S | lborel. f x) = integral S f" \ \Note that 0 = 0 holds when the integral diverges.\ proof (cases \set_integrable lborel S f\) case True thus ?thesis using set_borel_integral_eq_integral by force next case False - hence "LINT x : S | lborel. f x = 0" + hence "(LINT x : S | lborel. f x) = 0" unfolding set_lebesgue_integral_def set_integrable_def by (rewrite not_integrable_integral_eq; simp) moreover have "integral S f = 0" apply (rule not_integrable_integral) using False assms by (rewrite in asm integrable_on_iff_set_integrable_nonneg; simp) ultimately show ?thesis .. qed lemma set_borel_integral_eq_integral_nonneg_AE: fixes f :: "'a::euclidean_space \ real" assumes "AE x\S in lborel. f x \ 0" "f \ borel_measurable borel" "S \ sets borel" - shows "LINT x : S | lborel. f x = integral S f" + shows "(LINT x : S | lborel. f x) = integral S f" \ \Note that 0 = 0 holds when the integral diverges.\ proof (cases \set_integrable lborel S f\) case True thus ?thesis using set_borel_integral_eq_integral by force next case False - hence "LINT x : S | lborel. f x = 0" + hence "(LINT x : S | lborel. f x) = 0" unfolding set_lebesgue_integral_def set_integrable_def by (rewrite not_integrable_integral_eq; simp) moreover have "integral S f = 0" apply (rule not_integrable_integral) using False assms by (rewrite in asm integrable_on_iff_set_integrable_nonneg_AE; simp) ultimately show ?thesis .. qed subsection \Interchange of Differentiation and Lebesgue Integration\ definition measurize :: "'a measure \ 'b measure \ ('a \ 'b) \ 'a \ 'b" where "measurize M N f = (SOME g. g \ M \\<^sub>M N \ (\S\(null_sets M). {x \ space M. f x \ g x} \ S))" \ \The term "measurize" is what I coined in this formalization.\ lemma fixes f g assumes "g \ M \\<^sub>M N" "S \ null_sets M" "{x \ space M. f x \ g x} \ S" shows measurizeI: "AE x in M. f x = measurize M N f x" and measurizeI2: "AE x in M. g x = measurize M N f x" and measurize_measurable: "measurize M N f \ measurable M N" proof - let ?G = "\g. g \ M \\<^sub>M N" and ?S = "\g. \S\null_sets M. {x \ space M. f x \ g x} \ S" show "AE x in M. f x = measurize M N f x" unfolding measurize_def apply (rule someI2[of "\g. ?G g \ ?S g" g]) using assms apply blast using AE_I' by auto moreover have "AE x in M. g x = f x" using assms by (smt (verit, best) AE_I' Collect_cong) ultimately show "AE x in M. g x = measurize M N f x" by force show "measurize M N f \ measurable M N" unfolding measurize_def apply (rule conjE[of "?G g" "?S g"]) using assms apply auto[1] using someI_ex[of "\g. ?G g \ ?S g"] by auto qed corollary measurable_measurize_AE: fixes f assumes "f \ M \\<^sub>M N" shows "AE x in M. f x = measurize M N f x" by (rule measurizeI[where g=f and S="{}"]; simp add: assms) definition borel_measurize :: "'a measure \ ('a \ 'b::topological_space) \ 'a \ 'b" where "borel_measurize M f = measurize M borel f" lemma fixes f g assumes "g \ borel_measurable M" "S \ null_sets M" "{x \ space M. f x \ g x} \ S" shows borel_measurizeI: "AE x in M. f x = borel_measurize M f x" and borel_measurizeI2: "AE x in M. g x = borel_measurize M f x" and borel_measurize_measurable: "borel_measurize M f \ borel_measurable M" unfolding borel_measurize_def using assms apply - using measurizeI apply blast using measurizeI2 apply blast using measurize_measurable by blast corollary borel_measurable_measurize_AE: fixes f assumes "f \ borel_measurable M" shows "AE x in M. f x = borel_measurize M f x" using assms measurable_measurize_AE unfolding borel_measurize_def by auto definition set_borel_measurize :: "'a measure \ 'a set \ ('a \ 'b::topological_space) \ 'a \ 'b" where "set_borel_measurize M A f = borel_measurize (restrict_space M A) f" lemma fixes f g :: "'a \ 'b::real_normed_vector" and A assumes "A \ sets M" "set_borel_measurable M A g" "S \ null_sets M" "{x \ A. f x \ g x} \ S" shows set_borel_measurizeI: "AE x\A in M. f x = set_borel_measurize M A f x" and set_borel_measurizeI2: "AE x\A in M. g x = set_borel_measurize M A f x" and set_borel_measurize_measurable: "set_borel_measurable M A (set_borel_measurize M A f)" proof - have "g \ borel_measurable (restrict_space M A)" using assms by (rewrite set_borel_measurable_restrict_space_iff; simp) moreover have "S \ A \ null_sets (restrict_space M A)" using assms null_sets_restrict_space by (metis Int_lower2 null_set_Int2) moreover have "{x \ space (restrict_space M A). f x \ g x} \ S \ A" using assms by (rewrite space_restrict_space2; simp) ultimately show "AE x\A in M. f x = set_borel_measurize M A f x" and "AE x\A in M. g x = set_borel_measurize M A f x" and "set_borel_measurable M A (set_borel_measurize M A f)" unfolding set_borel_measurize_def using assms apply - apply (rewrite AE_restrict_space_iff[THEN sym], simp) apply (rule borel_measurizeI[of g _ "S \ A"]; simp) apply (rewrite AE_restrict_space_iff[THEN sym], simp) apply (rule borel_measurizeI2[of g _ "S \ A"]; simp) apply (rewrite set_borel_measurable_restrict_space_iff[THEN sym], simp) by (rule borel_measurize_measurable[of g _ "S \ A"]; simp) qed corollary set_borel_measurable_measurize_AE: fixes f::"'a \ 'b::real_normed_vector" and A assumes "set_borel_measurable M A f" "A \ sets M" shows "AE x\A in M. f x = set_borel_measurize M A f x" using set_borel_measurable_restrict_space_iff borel_measurable_measurize_AE AE_restrict_space_iff unfolding set_borel_measurize_def by (smt (verit) AE_cong sets.Int_space_eq2 assms) proposition interchange_deriv_LINT_general: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" assumes f_integ: "\r. r\{a<.. integrable M (f r)" and f_diff: "AE x in M. (\r. f r x) differentiable_on {a<..r\{a<..deriv (\r. f r x) r\ \ g x" "integrable M g" shows "\r. r\{a<.. ((\r. \x. f r x \M) has_real_derivative \x. borel_measurize M (\x. deriv (\r. f r x) r) x \M) (at r)" proof - text \Preparation\ have f_msr: "\r. r\{a<.. f r \ borel_measurable M" using f_integ by auto from f_diff obtain N1 where N1_null: "N1 \ null_sets M" and "\x. x \ space M - N1 \ (\s. f s x) differentiable_on {a<..x. x \ space M - N1 \ (\s. f s x) differentiable_on {a<.. null_sets M" and "\x. x \ space M - N2 \ \r\{a<..deriv (\s. f s x) r\ \ g x" by (smt (verit) AE_E3) hence Df_boundN2:"\x. x \ space M - N2 \ \r\{a<..deriv (\s. f s x) r\ \ g x" by (meson Diff_iff sets.sets_into_space subset_eq) define N where "N \ N1 \ N2" let ?CN = "space M - N" have N_null: "N \ null_sets M" and N_msr: "N \ sets M" unfolding N_def using N1_null N2_null by auto have f_diffCN: "\x. x\?CN \ (\s. f s x) differentiable_on {a<.. 'a \ real" where "Df r x \ indicator ({a<..?CN) (r,x) * deriv (\s. f s x) r" for r x have Df_boundCN: "\x. x\?CN \ \r\{a<..Df r x\ \ g x" unfolding Df_def N_def using Df_boundN2 by simp text \Main Part of the Proof\ fix r assume r_ab: "r\{a<.. 0" and ball_ab: "ball r e \ {a<..d::nat\real. \\i. d i \ UNIV-{0}; d \ 0\ \ ((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - fix d::"nat\real" assume d_neq0: "\i. d i \ UNIV-{0}" and d_to0: "d \ 0" then obtain m where "\i\m. \d i - 0\ < e" using LIMSEQ_def e_pos dist_real_def by metis hence rd_ab: "\n. r + d (n+m) \ {a<..n. (\x. (f (r + d (n+m)) x - f r x) / d (n+m)) \ borel_measurable M" using r_ab by (measurable; (intro f_msr)?; simp) hence limf_msr: "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))) \ borel_measurable M" by measurable moreover have limf_Df: "\x. x\?CN \ (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ Df r x" proof - fix x assume x_CN: "x\?CN" hence "(\s. f s x) field_differentiable (at r)" using f_diffCN r_ab by (metis at_within_open differentiable_on_eq_field_differentiable_real open_greaterThanLessThan) hence "((\h. (f (r+h) x - f r x) / h) \ Df r x) (at 0)" apply (rewrite in asm DERIV_deriv_iff_field_differentiable[THEN sym]) unfolding Df_def using r_ab x_CN by (simp add: DERIV_def) hence "(\i. (f (r + d i) x - f r x) / d i) \ Df r x" apply (rewrite in asm tendsto_at_iff_sequentially) apply (rule allE'[where x=d], simp) unfolding comp_def using d_neq0 d_to0 by simp thus "(\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ Df r x" by (rule LIMSEQ_ignore_initial_segment[where k=m]) qed ultimately have Df_eq: "\x. Df r x = indicator ?CN x * lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" proof - fix x show "Df r x = indicator ?CN x * lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" proof (cases \x\?CN\) case True hence "lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" by (intro limI, rule limf_Df) thus ?thesis using True by simp next case False thus ?thesis unfolding Df_def by simp qed qed hence Df_msr: "Df r \ borel_measurable M" apply (rewrite in "\x. \" Df_eq) apply (measurable; (rule limf_msr)?) using N_null unfolding null_sets_def by force have "((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M" proof - have "(\n. \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M" proof - \ \by Lebesgue's Dominated Convergence Theorem\ have "AE x in M. (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \ lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m))" using limf_Df Df_eq N_null by (smt (verit) DiffI AE_I' limI mem_Collect_eq subset_eq) moreover have "\n. AE x in M. norm ((f (r + d (n+m)) x - f r x) / d (n+m)) \ g x" proof - fix n { fix x assume x_CN: "x\?CN" let ?I = "{r..(r + d (n+m))} \ {(r + d (n+m))..r}" have f_diffI: "(\s. f s x) differentiable_on ?I" apply (rule differentiable_on_subset[where t="{a<..s. f s x)" "(\s. f s x) differentiable_on interior ?I" apply - using differentiable_imp_continuous_on apply blast by (metis differentiable_on_subset interior_subset) then obtain t where t_01: "t\{0<..<1}" and f_MVT: "f (r + d (n+m)) x - f r x = d (n+m) * deriv (\s. f s x) (r + t * (d (n+m)))" by (rule MVT_order_free) hence "0 < t" "t < 1" by simp_all hence rtd_ab: "r + t * (d (n+m)) \ {a<..s. f s x) (r + t * (d (n+m))) = d (n+m) * Df (r + t * (d (n+m))) x" proof - have "r + t * (d (n+m)) \ {a<..Df (r + t * (d (n+m))) x\ \ g x" using Df_boundCN x_CN rtd_ab by simp ultimately have "\(f (r + d (n+m)) x - f r x) / d (n+m)\ \ g x" by simp } thus "AE x in M. norm ((f (r + d (n+m)) x - f r x) / d (n+m)) \ g x" unfolding real_norm_def using AE_I' N_null by (smt (verit, ccfv_threshold) Diff_iff mem_Collect_eq subsetI) qed ultimately show "((\n. \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M) \ \x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M)" using limf_msr fd_msr Df_bound by (intro integral_dominated_convergence[where w=g], simp_all) qed moreover have "\n. ((\x. f (r + d (n+m)) x \M) - \x. f r x \M) / d (n+m) = \x. (f (r + d (n+m)) x - f r x) / d (n+m) \M" using d_neq0 apply simp by (rewrite Bochner_Integration.integral_diff; (rule f_integ | simp); (rule rd_ab | rule r_ab)) ultimately show ?thesis unfolding comp_def using d_neq0 apply - by (rule LIMSEQ_offset[where k=m]) simp qed moreover have "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M) = \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - have "(\x. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) \M) = \x. Df r x \M" proof - have "AE x in M. lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" proof - { fix x assume x_CN: "x\?CN" hence "lim (\n. (f (r + d (n+m)) x - f r x) / d (n+m)) = Df r x" by (simp add: Df_eq) } thus ?thesis using AE_I' N_null by (smt (verit, del_insts) DiffI mem_Collect_eq subsetI) qed thus ?thesis using limf_msr Df_msr by (intro integral_cong_AE; simp) qed also have "\ = \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" proof - have "AE x in M. Df r x = borel_measurize M (\y. deriv (\s. f s y) r) x" and "borel_measurize M (\y. deriv (\s. f s y) r) \ borel_measurable M" proof - have "{x \ space M. deriv (\s. f s x) r \ Df r x} \ N" proof - { fix x assume "x\?CN" hence "deriv (\s. f s x) r = Df r x" unfolding Df_def using r_ab by simp } thus ?thesis by blast qed thus "AE x in M. Df r x = borel_measurize M (\y. deriv (\s. f s y) r) x" and "borel_measurize M (\y. deriv (\s. f s y) r) \ borel_measurable M" - using Df_msr N_null - apply - - apply (rule borel_measurizeI2[where S=N]; simp) - by (rule borel_measurize_measurable[where g="Df r"]; simp) + using Df_msr N_null borel_measurize_measurable[where g="Df r"] + by (simp_all add: borel_measurizeI2) qed thus ?thesis using Df_msr by (intro integral_cong_AE; simp) qed finally show ?thesis . qed ultimately show "((\h. ((\x. f (r+h) x \M) - \x. f r x \M) / h) \ d) \ \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M" using tendsto_cong_limit by simp qed thus "((\s. \x. f s x \M) has_real_derivative \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) (at r)" by (rewrite DERIV_def, rewrite tendsto_at_iff_sequentially) simp qed proposition interchange_deriv_LINT: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" assumes "\r. r\{a<.. integrable M (f r)" and "AE x in M. (\r. f r x) differentiable_on {a<..r. r\{a<.. (\x. (deriv (\r. f r x) r)) \ borel_measurable M" and "AE x in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "integrable M g" shows "\r. r\{a<.. ((\r. \x. f r x \M) has_real_derivative \x. deriv (\r. f r x) r \M) (at r)" proof - fix r assume r_ab: "r\{a<..x. deriv (\s. f s x) r) \ borel_measurable M" using assms by simp have "((\s. \x. f s x \M) has_real_derivative \x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) (at r)" using assms r_ab by (intro interchange_deriv_LINT_general; simp) moreover have "(\x. borel_measurize M (\y. deriv (\s. f s y) r) x \M) = \x. deriv (\s. f s x) r \M" apply (rule integral_cong_AE) apply (rule borel_measurize_measurable[where g="\y. deriv (\s. f s y) r" and S="{}"], simp_all add: Df_msr) using borel_measurable_measurize_AE Df_msr by (smt (verit) AE_cong) ultimately show "((\r. \x. f r x \M) has_real_derivative \x. deriv (\r. f r x) r \M) (at r)" by simp qed proposition interchange_deriv_LINT_set_general: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" and A :: "'a set" assumes A_msr: "A \ sets M" and f_integ: "\r. r\{a<.. set_integrable M A (f r)" and f_diff: "AE x\A in M. (\r. f r x) differentiable_on {a<..A in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "set_integrable M A g" - shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative - \x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M) (at r)" + shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative + (\x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M)) (at r)" proof - let ?M_A = "restrict_space M A" have "\r. r\{a<.. integrable ?M_A (f r)" using A_msr f_integ set_integrable_restrict_space_iff by auto moreover have "AE x in ?M_A. (\r. f r x) differentiable_on {a<..r\{a<..deriv (\r. f r x) r\ \ g x" and "integrable ?M_A g" using A_msr Df_bound set_integrable_restrict_space_iff apply - by (simp add: AE_restrict_space_iff, auto) ultimately have "\r. r\{a<.. ((\r. integral\<^sup>L ?M_A (f r)) has_real_derivative integral\<^sup>L ?M_A (borel_measurize ?M_A (\x. deriv (\r. f r x) r))) (at r)" by (rule interchange_deriv_LINT_general[where M="restrict_space M A"]) auto thus "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative - \x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M) (at r)" + (\x\A. set_borel_measurize M A (\x. deriv (\r. f r x) r) x \M)) (at r)" unfolding set_borel_measurize_def using assms by (rewrite set_lebesgue_integral_restrict_space, simp)+ qed proposition interchange_deriv_LINT_set: fixes a b :: real and f :: "real \ 'a \ real" and g :: "'a \ real" and A :: "'a set" assumes "A \ sets M" and "\r. r\{a<.. set_integrable M A (f r)" and "AE x\A in M. (\r. f r x) differentiable_on {a<..r. r\{a<.. set_borel_measurable M A (\x. (deriv (\r. f r x) r))" and "AE x\A in M. \r\{a<..deriv (\r. f r x) r\ \ g x" "set_integrable M A g" - shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative - \x\A. deriv (\r. f r x) r \M) (at r)" + shows "\r. r\{a<.. ((\r. \x\A. f r x \M) has_real_derivative (\x\A. deriv (\r. f r x) r \M)) (at r)" proof - fix r assume r_ab: "r\{a<..x. deriv (\s. f s x) r)" using assms by simp have "((\s. \x\A. f s x \M) has_real_derivative - \x\A. set_borel_measurize M A (\y. deriv (\s. f s y) r) x \M) (at r)" + (\x\A. set_borel_measurize M A (\y. deriv (\s. f s y) r) x \M)) (at r)" using assms r_ab by (intro interchange_deriv_LINT_set_general; simp) moreover have "(\x\A. set_borel_measurize M A (\y. deriv (\s. f s y) r) x \M) = - \x\A. deriv (\s. f s x) r \M" + (\x\A. deriv (\s. f s x) r \M)" apply (rule set_lebesgue_integral_cong_AE2, simp add: assms) apply (rule set_borel_measurize_measurable[where g="\y. deriv (\s. f s y) r" and S="{}"], simp_all add: Df_msr assms) using set_borel_measurable_measurize_AE Df_msr assms by (smt (verit) AE_cong) - ultimately show "((\r. \x\A. f r x \M) has_real_derivative \x\A. deriv (\r. f r x) r \M) (at r)" + ultimately show "((\r. \x\A. f r x \M) has_real_derivative (\x\A. deriv (\r. f r x) r \M)) (at r)" by simp qed section \Additional Lemmas for the "Probability" Library\ lemma (in finite_borel_measure) fixes F :: "real \ real" assumes nondecF : "\ x y. x \ y \ F x \ F y" and right_cont_F : "\a. continuous (at_right a) F" and lim_F_at_bot : "(F \ 0) at_bot" and lim_F_at_top : "(F \ m) at_top" and m : "0 \ m" shows emeasure_interval_measure_Ioi: "emeasure (interval_measure F) {x<..} = m - F x" and measure_interval_measure_Ioi: "measure (interval_measure F) {x<..} = m - F x" proof - interpret F_FM: finite_measure "interval_measure F" using finite_borel_measure.axioms(1) finite_borel_measure_interval_measure lim_F_at_bot lim_F_at_top m nondecF right_cont_F by blast have "UNIV = {..x} \ {x<..}" by auto moreover have "{..x} \ {x<..} = {}" by auto ultimately have "emeasure (interval_measure F) UNIV = emeasure (interval_measure F) {..x} + emeasure (interval_measure F) {x<..}" by (simp add: plus_emeasure) moreover have "emeasure (interval_measure F) UNIV = m" using assms interval_measure_UNIV by presburger ultimately show \: "emeasure (interval_measure F) {x<..} = m - F x" using assms emeasure_interval_measure_Iic by (metis ennreal_add_diff_cancel_left ennreal_minus measure_interval_measure_Iic measure_nonneg top_neq_ennreal) hence "ennreal (measure (interval_measure F) {x<..}) = m - F x" using emeasure_eq_measure by (metis emeasure_eq_ennreal_measure top_neq_ennreal) moreover have "\x. F x \ m" using lim_F_at_top nondecF by (intro mono_at_top_le[where f=F]; simp add: mono_def) ultimately show "measure (interval_measure F) {x<..} = m - F x" using ennreal_inj F_FM.emeasure_eq_measure by force qed lemma (in prob_space) cond_prob_nonneg[simp]: "cond_prob M P Q \ 0" by (auto simp: cond_prob_def) lemma (in prob_space) cond_prob_whole_1: "cond_prob M P P = 1" if "prob {\ \ space M. P \} \ 0" unfolding cond_prob_def using that by simp lemma (in prob_space) cond_prob_0_null: "cond_prob M P Q = 0" if "prob {\ \ space M. Q \} = 0" unfolding cond_prob_def using that by simp lemma (in prob_space) cond_prob_AE_prob: assumes "{\ \ space M. P \} \ events" "{\ \ space M. Q \} \ events" and "AE \ in M. Q \" shows "cond_prob M P Q = prob {\ \ space M. P \}" proof - let ?setP = "{\ \ space M. P \}" let ?setQ = "{\ \ space M. Q \}" have [simp]: "prob ?setQ = 1" using assms prob_Collect_eq_1 by simp hence "cond_prob M P Q = prob (?setP \ ?setQ)" unfolding cond_prob_def by (simp add: Collect_conj_eq2) also have "\ = prob ?setP" proof (rule antisym) show "prob (?setP \ ?setQ) \ prob ?setP" using assms finite_measure_mono inf_sup_ord(1) by blast next show "prob ?setP \ prob (?setP \ ?setQ)" proof - have "prob (?setP \ ?setQ) = prob ?setP + prob ?setQ - prob (?setP \ ?setQ)" using assms by (smt (verit) finite_measure_Diff' finite_measure_Union' sup_commute) also have "\ = prob ?setP + (1 - prob (?setP \ ?setQ))" by simp also have "\ \ prob ?setP" by simp finally show ?thesis . qed qed finally show ?thesis . qed subsection \More Properties of cdf's\ context finite_borel_measure begin lemma cdf_diff_eq2: assumes "x \ y" shows "cdf M y - cdf M x = measure M {x<..y}" proof (cases \x = y\) case True thus ?thesis by force next case False hence "x < y" using assms by simp thus ?thesis by (rule cdf_diff_eq) qed end context prob_space begin lemma cdf_distr_measurable [measurable]: assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel X) \ borel_measurable borel" proof (rule borel_measurable_mono) show "mono (cdf (distr M borel X))" unfolding mono_def using finite_borel_measure.cdf_nondecreasing by (simp add: real_distribution.finite_borel_measure_M) qed lemma cdf_distr_P: assumes "random_variable borel X" shows "cdf (distr M borel X) x = \

(\ in M. X \ \ x)" unfolding cdf_def apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) lemma cdf_continuous_distr_P_lt: assumes "random_variable borel X" "isCont (cdf (distr M borel X)) x" shows "cdf (distr M borel X) x = \

(\ in M. X \ < x)" proof - have "\

(\ in M. X \ < x) = measure (distr M borel X) {.. = measure (distr M borel X) ({.. {x})" apply (rewrite finite_measure.measure_zero_union, simp_all add: assms finite_measure_distr) using finite_borel_measure.isCont_cdf real_distribution.finite_borel_measure_M assms by blast also have "\ = measure (distr M borel X) {..x}" by (metis ivl_disj_un_singleton(2)) also have "\ = cdf (distr M borel X) x" unfolding cdf_def by simp finally show ?thesis by simp qed lemma cdf_distr_diff_P: assumes "x \ y" and "random_variable borel X" shows "cdf (distr M borel X) y - cdf (distr M borel X) x = \

(\ in M. x < X \ \ X \ \ y)" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp have "cdf (distr M borel X) y - cdf (distr M borel X) x = measure (distr M borel X) {x<..y}" by (rewrite distrX_FBM.cdf_diff_eq2; simp add: assms) also have "\ = \

(\ in M. x < X \ \ X \ \ y)" apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) finally show ?thesis . qed lemma cdf_distr_max: fixes c::real assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. max (X x) c)) u = cdf (distr M borel X) u * indicator {c..} u" proof (cases \c \ u\) case True thus ?thesis unfolding cdf_def apply (rewrite measure_distr; simp?)+ by (smt (verit) Collect_cong atMost_iff vimage_def) next case False thus ?thesis unfolding cdf_def apply (rewrite measure_distr; simp?)+ by (smt (verit, best) Int_emptyI atMost_iff measure_empty vimage_eq) qed lemma cdf_distr_min: fixes c::real assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. min (X x) c)) u = 1 - (1 - cdf (distr M borel X) u) * indicator {..c \ u\) case True thus ?thesis unfolding cdf_def using real_distribution.finite_borel_measure_M real_distribution_distr apply (rewrite measure_distr; simp?) by (smt (verit, del_insts) Int_absorb1 atMost_iff prob_space subset_vimage_iff) next case False thus ?thesis unfolding cdf_def using real_distribution.finite_borel_measure_M real_distribution_distr apply (rewrite measure_distr; simp?)+ using prob_space_axioms assms by (smt (verit) Collect_cong Int_def atMost_iff prob_space prob_space.cdf_distr_P vimage_eq) qed lemma cdf_distr_floor_P: fixes X :: "'a \ real" assumes [measurable]: "random_variable borel X" shows "cdf (distr M borel (\x. \X x\)) u = \

(x in M. X x < \u\ + 1)" unfolding cdf_def apply (rewrite measure_distr; simp?) apply (rule arg_cong[where f=prob]) unfolding vimage_def using floor_le_iff le_floor_iff by blast lemma expectation_nonneg_tail: assumes [measurable]: "random_variable borel X" and X_nonneg: "\x. x \ space M \ X x \ 0" defines "F u \ cdf (distr M borel X) u" - shows "(\\<^sup>+x. ennreal (X x) \M) = \\<^sup>+u\{0..}. ennreal (1 - F u) \lborel" + shows "(\\<^sup>+x. ennreal (X x) \M) = (\\<^sup>+u\{0..}. ennreal (1 - F u) \lborel)" proof - let ?distrX = "distr M borel X" have "(\\<^sup>+x. ennreal (X x) \M) = (\\<^sup>+u. ennreal u \?distrX)" by (rewrite nn_integral_distr; simp) also have "\ = (\\<^sup>+u\{0..}. ennreal u \?distrX)" by (rule nn_integral_distr_set; simp add: X_nonneg) - also have "\ = \\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX" + also have "\ = (\\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX)" proof - - have "\u::real. u\{0..} \ ennreal u = \\<^sup>+v\{0..}. indicator {..lborel" + have "\u::real. u\{0..} \ ennreal u = (\\<^sup>+v\{0..}. indicator {..lborel)" apply (rewrite indicator_inter_arith[THEN sym]) apply (rewrite nn_integral_indicator, measurable, simp) by (metis atLeastLessThan_def diff_zero emeasure_lborel_Ico inf.commute) thus ?thesis by (metis (no_types, lifting) indicator_eq_0_iff mult_eq_0_iff) qed - also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel" + also have "\ = (\\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel)" proof - - have "\\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX = + have "(\\<^sup>+u\{0..}. (\\<^sup>+v\{0..}. indicator {..lborel) \?distrX) = \\<^sup>+u. (\\<^sup>+v. indicator {..lborel) \?distrX" by (rewrite nn_integral_multc; simp) also have "\ = \\<^sup>+v. (\\<^sup>+u. indicator {..?distrX) \lborel" apply (rewrite pair_sigma_finite.Fubini'; simp?) using pair_sigma_finite.intro assms prob_space_distr prob_space_imp_sigma_finite sigma_finite_lborel apply blast by measurable auto - also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel" + also have "\ = (\\<^sup>+v\{0..}. (\\<^sup>+u\{0..}. indicator {..?distrX) \lborel)" apply (rewrite nn_integral_multc[THEN sym]; measurable; simp?) apply (rule nn_integral_cong)+ using mult.assoc mult.commute by metis finally show ?thesis by simp qed - also have "\ = \\<^sup>+v\{0..}. (\\<^sup>+u. indicator {v<..} u \?distrX) \lborel" + also have "\ = (\\<^sup>+v\{0..}. (\\<^sup>+u. indicator {v<..} u \?distrX) \lborel)" apply (rule nn_integral_cong) apply (rewrite nn_integral_multc[THEN sym], measurable; (simp del: nn_integral_indicator)?)+ apply (rule nn_integral_cong) using lessThan_iff greaterThan_iff atLeast_iff indicator_simps by (smt (verit, del_insts) mult_1 mult_eq_0_iff) also have "\ = (\\<^sup>+v\{0..}. ennreal (1 - F v) \lborel)" apply (rule nn_integral_cong, simp) apply (rewrite emeasure_distr; simp?) apply (rewrite vimage_compl_atMost[THEN sym]) unfolding F_def cdf_def apply (rewrite measure_distr; simp?) apply (rewrite prob_compl[THEN sym], simp) by (metis (no_types, lifting) Diff_Compl Diff_Diff_Int Int_commute emeasure_eq_measure) finally show ?thesis . qed lemma expectation_nonpos_tail: assumes [measurable]: "random_variable borel X" and X_nonpos: "\x. x \ space M \ X x \ 0" defines "F u \ cdf (distr M borel X) u" - shows "(\\<^sup>+x. ennreal (- X x) \M) = \\<^sup>+u\{..0}. ennreal (F u) \lborel" + shows "(\\<^sup>+x. ennreal (- X x) \M) = (\\<^sup>+u\{..0}. ennreal (F u) \lborel)" proof - let ?distrX = "distr M borel X" have "(\\<^sup>+x. ennreal (- X x) \M) = (\\<^sup>+u. ennreal (-u) \?distrX)" by (rewrite nn_integral_distr; simp) - also have "\ = \\<^sup>+u\{..0}. ennreal (-u) \?distrX" + also have "\ = (\\<^sup>+u\{..0}. ennreal (-u) \?distrX)" proof - have [simp]: "{..0::real} \ {0<..} = UNIV" by force have "(\\<^sup>+u. ennreal (-u) \?distrX) = (\\<^sup>+u\{..0}. ennreal (-u) \?distrX) + (\\<^sup>+u\{0<..}. ennreal (-u) \?distrX)" by (rewrite nn_integral_disjoint_pair[THEN sym], simp_all, force) also have "\ = (\\<^sup>+u\{..0}. ennreal (-u) \?distrX)" apply (rewrite nn_integral_zero'[of "\u. ennreal (-u) * indicator {0<..} u"]; simp?) unfolding indicator_def using always_eventually ennreal_lt_0 by force finally show ?thesis . qed - also have "\ = \\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX" + also have "\ = (\\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX)" proof - - have "\u::real. u\{..0} \ ennreal (-u) = \\<^sup>+v\{..0}. indicator {u..} v \lborel" + have "\u::real. u\{..0} \ ennreal (-u) = (\\<^sup>+v\{..0}. indicator {u..} v \lborel)" apply (rewrite indicator_inter_arith[THEN sym]) apply (rewrite nn_integral_indicator, measurable, simp) by (metis emeasure_lborel_Icc atLeastAtMost_def diff_0) thus ?thesis by (metis (no_types, lifting) indicator_eq_0_iff mult_eq_0_iff) qed - also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel" + also have "\ = (\\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel)" proof - - have "\\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX = + have "(\\<^sup>+u\{..0}. (\\<^sup>+v\{..0}. indicator {u..} v \lborel) \?distrX) = \\<^sup>+u. (\\<^sup>+v. indicator {u..} v * indicator {..0} v * indicator {..0} u \lborel) \?distrX" by (rewrite nn_integral_multc; simp) also have "\ = \\<^sup>+v. (\\<^sup>+u. indicator {u..} v * indicator {..0} v * indicator {..0} u \?distrX) \lborel" apply (rewrite pair_sigma_finite.Fubini'; simp?) using pair_sigma_finite.intro assms prob_space_distr prob_space_imp_sigma_finite sigma_finite_lborel apply blast by measurable auto - also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel" + also have "\ = (\\<^sup>+v\{..0}. (\\<^sup>+u\{..0}. indicator {u..} v \?distrX) \lborel)" apply (rewrite nn_integral_multc[THEN sym]; measurable; simp?) apply (rule nn_integral_cong)+ using mult.assoc mult.commute by metis finally show ?thesis by simp qed - also have "\ = \\<^sup>+v\{..0}. (\\<^sup>+u. indicator {..v} u \?distrX) \lborel" + also have "\ = (\\<^sup>+v\{..0}. (\\<^sup>+u. indicator {..v} u \?distrX) \lborel)" apply (rule nn_integral_cong) apply (rewrite nn_integral_multc[THEN sym], measurable; (simp del: nn_integral_indicator)?)+ apply (rule nn_integral_cong) using atMost_iff atLeast_iff indicator_simps by (smt (verit, del_insts) mult_1 mult_eq_0_iff) - also have "\ = \\<^sup>+v\{..0}. ennreal (F v) \lborel" + also have "\ = (\\<^sup>+v\{..0}. ennreal (F v) \lborel)" apply (rule nn_integral_cong, simp) apply (rewrite emeasure_distr; simp?) unfolding F_def cdf_def by (rewrite measure_distr; simp add: emeasure_eq_measure) finally show ?thesis . qed proposition expectation_tail: assumes [measurable]: "integrable M X" defines "F u \ cdf (distr M borel X) u" shows "expectation X = (LBINT u:{0..}. 1 - F u) - (LBINT u:{..0}. F u)" proof - have "expectation X = expectation (\x. max (X x) 0) + expectation (\x. min (X x) 0)" using integrable_max integrable_min apply (rewrite Bochner_Integration.integral_add[THEN sym], measurable) by (rule Bochner_Integration.integral_cong; simp) also have "\ = expectation (\x. max (X x) 0) - expectation (\x. - min (X x) 0)" by force also have "\ = (LBINT u:{0..}. 1 - F u) - (LBINT u:{..0}. F u)" proof - - have "expectation (\x. max (X x) 0) = LBINT u:{0..}. 1 - F u" + have "expectation (\x. max (X x) 0) = (LBINT u:{0..}. 1 - F u)" proof - have "expectation (\x. max (X x) 0) = enn2real (\\<^sup>+x. ennreal (max (X x) 0) \M)" by (rule integral_eq_nn_integral; simp) also have "\ = enn2real (\\<^sup>+u\{0..}. ennreal (1 - F u) \lborel)" apply (rewrite expectation_nonneg_tail; simp?) apply (rewrite cdf_distr_max, simp) unfolding F_def by (metis (opaque_lifting) indicator_simps mult.commute mult_1 mult_eq_0_iff) also have "\ = enn2real (\\<^sup>+u. ennreal ((1 - F u) * indicator {0..} u) \lborel)" by (simp add: indicator_mult_ennreal mult.commute) - also have "\ = LBINT u:{0..}. 1 - F u" + also have "\ = (LBINT u:{0..}. 1 - F u)" apply (rewrite integral_eq_nn_integral[THEN sym], simp add: F_def) unfolding F_def using real_distribution.cdf_bounded_prob apply force unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_cong; simp) finally show ?thesis . qed - moreover have "expectation (\x. - min (X x) 0) = LBINT u:{..0}. F u" + moreover have "expectation (\x. - min (X x) 0) = (LBINT u:{..0}. F u)" proof - have "expectation (\x. - min (X x) 0) = enn2real (\\<^sup>+x. ennreal (- min (X x) 0) \M)" by (rule integral_eq_nn_integral; simp) also have "\ = enn2real (\\<^sup>+u\{..0}. ennreal (F u) \lborel)" proof - let ?distrminX = "distr M borel (\x. min (X x) 0)" have [simp]: "sym_diff {..0} {..<0} = {0::real}" by force have "enn2real (\\<^sup>+x. ennreal (- min (X x) 0) \M) = enn2real (\\<^sup>+u\{..0}. ennreal (cdf ?distrminX u) \lborel)" by (rewrite expectation_nonpos_tail; simp) also have "\ = enn2real (\\<^sup>+u\{..<0}. ennreal (cdf ?distrminX u) \lborel)" by (rewrite nn_integral_null_delta, auto) also have "\ = enn2real (\\<^sup>+u\{..<0}. ennreal (F u) \lborel)" apply (rewrite cdf_distr_min, simp) apply (rule arg_cong[where f=enn2real], rule nn_integral_cong) unfolding F_def by (smt (verit) indicator_simps mult_cancel_left1 mult_eq_0_iff) also have "\ = enn2real (\\<^sup>+u\{..0}. ennreal (F u) \lborel)" by (rewrite nn_integral_null_delta, auto simp add: sup_commute) finally show ?thesis . qed also have "\ = enn2real (\\<^sup>+u. ennreal (F u * indicator {..0} u) \lborel)" using mult.commute indicator_mult_ennreal by metis - also have "\ = LBINT u:{..0}. F u" + also have "\ = (LBINT u:{..0}. F u)" apply (rewrite integral_eq_nn_integral[THEN sym], simp add: F_def) unfolding F_def using finite_borel_measure.cdf_nonneg real_distribution.finite_borel_measure_M apply simp unfolding set_lebesgue_integral_def by (rule Bochner_Integration.integral_cong; simp) finally show ?thesis . qed ultimately show ?thesis by simp qed finally show ?thesis . qed proposition distributed_deriv_cdf: assumes [measurable]: "random_variable borel X" defines "F u \ cdf (distr M borel X) u" assumes "finite S" "\x. x \ S \ (F has_real_derivative f x) (at x)" and "continuous_on UNIV F" "f \ borel_measurable lborel" shows "distributed M lborel X f" proof - have FBM: "finite_borel_measure (distr M borel X)" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp then interpret distrX_FBM: finite_borel_measure "distr M borel X" . have FBMl: "finite_borel_measure (distr M lborel X)" using FBM distr_borel_lborel by smt then interpret distrlX_FBM: finite_borel_measure "distr M lborel X" . have [simp]: "(\x. ennreal (f x)) \ borel_measurable borel" using assms by simp moreover have "distr M lborel X = density lborel f" proof - have "\a b. a \ b \ emeasure (distr M lborel X) {a<..b} < \" using distrlX_FBM.emeasure_real less_top_ennreal by blast moreover have "\a b. a \ b \ emeasure (distr M lborel X) {a<..b} = emeasure (density lborel f) {a<..b}" proof - fix a b :: real assume "a \ b" hence [simp]: "sym_diff {a<..b} {a..b} = {a}" by force - have "emeasure (density lborel f) {a<..b} = \\<^sup>+x\{a<..b}. ennreal (f x) \lborel" + have "emeasure (density lborel f) {a<..b} = (\\<^sup>+x\{a<..b}. ennreal (f x) \lborel)" by (rewrite emeasure_density; simp) - also have "\ = \\<^sup>+x\{a..b}. ennreal (f x) \lborel" by (rewrite nn_integral_null_delta, auto) - also have "\ = \\<^sup>+x. ennreal (indicat_real {a..b} x * f x) \lborel" + also have "\ = (\\<^sup>+x\{a..b}. ennreal (f x) \lborel)" by (rewrite nn_integral_null_delta, auto) + also have "\ = (\\<^sup>+x. ennreal (indicat_real {a..b} x * f x) \lborel)" by (metis indicator_mult_ennreal mult.commute) also have "\ = ennreal (F b - F a)" proof - define g where "g x = (if x \ S then 0 else f x)" for x :: real have [simp]: "\x. g x \ 0" unfolding g_def apply (split if_split, auto) apply (rule mono_on_imp_deriv_nonneg[of UNIV F], auto) unfolding F_def mono_on_def using distrX_FBM.cdf_nondecreasing apply blast using assms unfolding F_def by force have "(\\<^sup>+x. ennreal (indicat_real {a..b} x * f x) \lborel) = \\<^sup>+x. ennreal (indicat_real {a..b} x * g x) \lborel" apply (rule nn_integral_cong_AE) apply (rule AE_mp[where P= "\x. x \ S"]) using assms finite_imp_null_set_lborel AE_not_in apply blast unfolding g_def by simp also have "\ = ennreal (F b - F a)" apply (rewrite nn_integral_has_integral_lebesgue, simp) apply (rule fundamental_theorem_of_calculus_strong[of S], auto simp: \a \ b\ g_def assms) using has_real_derivative_iff_has_vector_derivative assms apply presburger using assms continuous_on_subset subsetI by fastforce finally show ?thesis . qed also have "\ = emeasure (distr M lborel X) {a <.. b}" apply (rewrite distrlX_FBM.emeasure_Ioc, simp add: \a \ b\) unfolding F_def cdf_def apply (rewrite ennreal_minus[THEN sym], simp)+ by (metis distr_borel_lborel) finally show "emeasure (distr M lborel X) {a<..b} = emeasure (density lborel f) {a<..b}" by simp qed ultimately show ?thesis by (intro measure_eqI_Ioc; simp) qed ultimately show ?thesis unfolding distributed_def by auto qed end text \ Define the conditional probability space. This is obtained by rescaling the restricted space of a probability space. \ subsection \Conditional Probability Space\ lemma restrict_prob_space: assumes "measure_space \ A \" "a \ A" and "0 < \ a" "\ a < \" shows "prob_space (scale_measure (1 / \ a) (restrict_space (measure_of \ A \) a))" proof let ?M = "measure_of \ A \" let ?Ma = "restrict_space ?M a" let ?rMa = "scale_measure (1 / \ a) ?Ma" have "emeasure ?rMa (space ?rMa) = 1 / \ a * emeasure ?Ma (space ?rMa)" by simp also have "\ = 1 / \ a * emeasure ?M (space ?rMa)" using assms apply (rewrite emeasure_restrict_space) apply (simp add: measure_space_def sigma_algebra.sets_measure_of_eq) by (simp add: space_restrict_space space_scale_measure) simp also have "\ = 1 / \ a * emeasure ?M (space ?Ma)" by (rewrite space_scale_measure) simp also have "\ = 1 / \ a * emeasure ?M a" using assms apply (rewrite space_restrict_space2) by (simp add: measure_space_closed)+ also have "\ = 1" using assms measure_space_def apply (rewrite emeasure_measure_of_sigma, blast+) by (simp add: ennreal_divide_times) finally show "emeasure ?rMa (space ?rMa) = 1" . qed definition cond_prob_space :: "'a measure \ 'a set \ 'a measure" (infix "\" 200) where "M\A \ scale_measure (1 / emeasure M A) (restrict_space M A)" context prob_space begin lemma cond_prob_space_whole[simp]: "M \ space M = M" unfolding cond_prob_space_def using emeasure_space_1 by simp lemma cond_prob_space_correct: assumes "A \ events" "prob A > 0" shows "prob_space (M\A)" unfolding cond_prob_space_def apply (subst(2) measure_of_of_measure[of M, THEN sym]) using assms by (intro restrict_prob_space; (simp add: measure_space)?; simp_all add: emeasure_eq_measure) lemma space_cond_prob_space: assumes "A \ events" shows "space (M\A) = A" unfolding cond_prob_space_def using assms by (simp add: space_scale_measure) lemma sets_cond_prob_space: "sets (M\A) = (\) A ` events" unfolding cond_prob_space_def by (metis sets_restrict_space sets_scale_measure) lemma measure_cond_prob_space: assumes "A \ events" "B \ events" and "prob A > 0" shows "measure (M\A) (B \ A) = prob (B \ A) / prob A" proof - have "1 / emeasure M A = ennreal (1 / prob A)" using assms by (smt (verit) divide_ennreal emeasure_eq_measure ennreal_1) hence "measure (M\A) (B \ A) = (1 / prob A) * measure (restrict_space M A) (B \ A)" unfolding cond_prob_space_def using measure_scale_measure by force also have "\ = (1 / prob A) * prob (B \ A)" using measure_restrict_space assms by (metis inf.cobounded2 sets.Int_space_eq2) also have "\ = prob (B \ A) / prob A" by simp finally show ?thesis . qed corollary measure_cond_prob_space_subset: assumes "A \ events" "B \ events" "B \ A" and "prob A > 0" shows "measure (M\A) B = prob B / prob A" proof - have "B = B \ A" using assms by auto moreover have "measure (M\A) (B \ A) = prob (B \ A) / prob A" using assms measure_cond_prob_space by simp ultimately show ?thesis by auto qed lemma cond_cond_prob_space: assumes "A \ events" "B \ events" "B \ A" "prob B > 0" shows "(M\A)\B = M\B" proof (rule measure_eqI) have pA_pos[simp]: "prob A > 0" using assms by (smt (verit, ccfv_SIG) finite_measure_mono) interpret MA_PS: prob_space "M\A" using cond_prob_space_correct assms by simp interpret MB_PS: prob_space "M\B" using cond_prob_space_correct assms by simp have "1 / emeasure M A = ennreal (1 / prob A)" using pA_pos by (smt (verit, ccfv_SIG) divide_ennreal emeasure_eq_measure ennreal_1) hence [simp]: "0 < MA_PS.prob B" using assms pA_pos by (metis divide_eq_0_iff measure_cond_prob_space_subset zero_less_measure_iff) have [simp]: "B \ MA_PS.events" using assms by (rewrite sets_cond_prob_space, unfold image_def) blast have [simp]: "finite_measure ((M\A)\B)" by (rule prob_space.finite_measure, rule prob_space.cond_prob_space_correct, simp_all add: MA_PS.prob_space_axioms) show sets_MAB: "sets ((M\A)\B) = sets (M\B)" apply (rewrite prob_space.sets_cond_prob_space) using MA_PS.prob_space_axioms apply presburger apply (rewrite sets_cond_prob_space, unfold image_def)+ using assms by blast show "\C. C \ sets ((M\A)\B) \ emeasure ((M\A)\B) C = emeasure (M\B) C" proof - fix C assume "C \ sets ((M\A)\B)" hence "C \ sets (M\B)" using sets_MAB by simp from this obtain D where "D \ events" "C = B \ D" by (rewrite in asm sets_cond_prob_space, auto) hence [simp]: "C \ events" and [simp]: "C \ B" and [simp]: "C \ A" using assms by auto hence [simp]: "C \ MA_PS.events" using assms by (rewrite sets_cond_prob_space, unfold image_def) blast show "emeasure ((M\A)\B) C = emeasure (M\B) C" apply (rewrite finite_measure.emeasure_eq_measure, simp)+ apply (rewrite ennreal_inj, simp_all) apply (rewrite prob_space.measure_cond_prob_space_subset, simp_all add: assms prob_space_axioms MA_PS.prob_space_axioms)+ using pA_pos by fastforce qed qed lemma cond_prob_space_prob: assumes[measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "measure (M \ {x \ space M. Q x}) {x \ space M. P x \ Q x} = \

(x in M. P x \ Q x)" proof - let ?SetP = "{x \ space M. P x}" let ?SetQ = "{x \ space M. Q x}" have "measure (M\?SetQ) {x \ space M. P x \ Q x} = measure (M\?SetQ) (?SetP \ ?SetQ)" by (smt (verit, ccfv_SIG) Collect_cong Int_def mem_Collect_eq) also have "\ = prob (?SetP \ ?SetQ) / prob ?SetQ" using assms by (rewrite measure_cond_prob_space; simp?) also have "\ = \

(x in M. P x \ Q x)" unfolding cond_prob_def assms by (smt (verit) Collect_cong Int_def mem_Collect_eq) finally show ?thesis . qed lemma cond_prob_space_cond_prob: assumes [measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "\

(x in M. P x \ Q x) = \

(x in (M \ {x \ space M. Q x}). P x)" proof - let ?setQ = "{x \ space M. Q x}" have "\

(x in M. P x \ Q x) = measure (M\?setQ) {x \ space M. P x \ Q x}" using cond_prob_space_prob assms by simp also have "\ = \

(x in (M\?setQ). P x)" proof - have "{x \ space M. P x \ Q x} = {x \ space (M\?setQ). P x}" using space_cond_prob_space assms by force thus ?thesis by simp qed finally show ?thesis . qed lemma cond_prob_neg: assumes[measurable]: "Measurable.pred M P" "Measurable.pred M Q" and "\

(x in M. Q x) > 0" shows "\

(x in M. \ P x \ Q x) = 1 - \

(x in M. P x \ Q x)" proof - let ?setP = "{x \ space M. P x}" let ?setQ = "{x \ space M. Q x}" interpret setQ_PS: prob_space "M\?setQ" using cond_prob_space_correct assms by simp have [simp]: "{x \ space (M\?setQ). P x} \ setQ_PS.events" proof - have "{x \ space (M\?setQ). P x} = ?setQ \ ?setP" using space_cond_prob_space by force thus ?thesis using sets_cond_prob_space by simp qed have "\

(x in M. \ P x \ Q x) = \

(x in M\?setQ. \ P x)" by (rewrite cond_prob_space_cond_prob; simp add: assms) also have "\ = 1 - \

(x in M\?setQ. P x)" by (rewrite setQ_PS.prob_neg, simp_all add: assms) also have "\ = 1 - \

(x in M. P x \ Q x)" by (rewrite cond_prob_space_cond_prob; simp add: assms) finally show ?thesis . qed lemma random_variable_cond_prob_space: assumes "A \ events" "prob A > 0" and [measurable]: "random_variable borel X" shows "X \ borel_measurable (M\A)" proof (rule borel_measurableI) fix S :: "'b set" assume [measurable]: "open S" show "X -` S \ space (M \ A) \ sets (M \ A)" apply (rewrite space_cond_prob_space, simp add: assms) apply (rewrite sets_cond_prob_space, simp add: image_def) apply (rule bexI[of _ "X -` S \ space M"]; measurable) using sets.Int_space_eq2 Int_commute assms by auto qed lemma AE_cond_prob_space_iff: assumes "A \ events" "prob A > 0" shows "(AE x in M\A. P x) \ (AE x in M. x \ A \ P x)" proof - have [simp]: "1 / emeasure M A > 0" using assms divide_ennreal emeasure_eq_measure ennreal_1 by (smt (verit) divide_pos_pos ennreal_eq_0_iff not_gr_zero) show ?thesis unfolding cond_prob_space_def apply (rewrite AE_scale_measure_iff, simp) by (rewrite AE_restrict_space_iff; simp add: assms) qed lemma integral_cond_prob_space_nn: assumes "A \ events" "prob A > 0" and [measurable]: "random_variable borel X" and nonneg: "AE x in M. x \ A \ 0 \ X x" shows "integral\<^sup>L (M\A) X = expectation (\x. indicator A x * X x) / prob A" proof - have [simp]: "X \ borel_measurable (M\A)" by (rule random_variable_cond_prob_space; (simp add: assms)) have [simp]: "AE x in (M\A). 0 \ X x" by (rewrite AE_cond_prob_space_iff; simp add: assms) have [simp]: "random_variable borel (\x. indicat_real A x * X x)" using borel_measurable_indicator assms by force have [simp]: "AE x in M. 0 \ indicat_real A x * X x" using nonneg by fastforce have "integral\<^sup>L (M\A) X = enn2real (\\<^sup>+ x. ennreal (X x) \(M\A))" by (rewrite integral_eq_nn_integral; simp) also have "\ = enn2real (1 / prob A * \\<^sup>+ x. ennreal (X x) \(restrict_space M A))" unfolding cond_prob_space_def apply (rewrite nn_integral_scale_measure, simp add: measurable_restrict_space1) using divide_ennreal emeasure_eq_measure ennreal_1 assms by smt also have "\ = enn2real (1 / prob A * (\\<^sup>+ x. ennreal (indicator A x * X x) \M))" apply (rewrite nn_integral_restrict_space, simp add: assms) by (metis indicator_mult_ennreal mult.commute) also have "\ = integral\<^sup>L M (\x. indicator A x * X x) / prob A" apply (rewrite integral_eq_nn_integral; simp?) by (simp add: divide_nonneg_pos enn2real_mult) finally show ?thesis by simp qed end text \ Define the complementary cumulative distribution function, also known as tail distribution. The theory presented below is a slight modification of the subsection "Properties of cdf's" in the theory "@{text Distribution_Functions}". \ subsection \Complementary Cumulative Distribution Function\ definition ccdf :: "real measure \ real \ real" where "ccdf M \ \x. measure M {x<..}" \ \complementary cumulative distribution function (tail distribution)\ lemma ccdf_def2: "ccdf M x = measure M {x<..}" by (simp add: ccdf_def) context finite_borel_measure begin lemma add_cdf_ccdf: "cdf M x + ccdf M x = measure M (space M)" proof - have "{..x} \ {x<..} = UNIV" by auto moreover have "{..x} \ {x<..} = {}" by auto ultimately have "emeasure M {..x} + emeasure M {x<..} = emeasure M UNIV" using plus_emeasure M_is_borel atMost_borel greaterThan_borel by metis hence "measure M {..x} + measure M {x<..} = measure M UNIV" using finite_emeasure_space emeasure_eq_measure ennreal_inj by (smt (verit, ccfv_SIG) ennreal_plus measure_nonneg) thus ?thesis unfolding cdf_def ccdf_def using borel_UNIV by simp qed lemma ccdf_cdf: "ccdf M = (\x. measure M (space M) - cdf M x)" by (rule ext) (smt add_cdf_ccdf) lemma cdf_ccdf: "cdf M = (\x. measure M (space M) - ccdf M x)" by (rule ext) (smt add_cdf_ccdf) lemma isCont_cdf_ccdf: "isCont (cdf M) x \ isCont (ccdf M) x" proof show "isCont (cdf M) x \ isCont (ccdf M) x" by (rewrite ccdf_cdf) auto next show "isCont (ccdf M) x \ isCont (cdf M) x" by (rewrite cdf_ccdf) auto qed lemma isCont_ccdf: "isCont (ccdf M) x \ measure M {x} = 0" using isCont_cdf_ccdf isCont_cdf by simp lemma continuous_cdf_ccdf: shows "continuous F (cdf M) \ continuous F (ccdf M)" (is "?LHS \ ?RHS") proof assume ?LHS thus ?RHS using continuous_diff continuous_const by (rewrite ccdf_cdf) blast next assume ?RHS thus ?LHS using continuous_diff continuous_const by (rewrite cdf_ccdf) blast qed lemma has_real_derivative_cdf_ccdf: "(cdf M has_real_derivative D) F \ (ccdf M has_real_derivative -D) F" proof assume "(cdf M has_real_derivative D) F" thus "(ccdf M has_real_derivative -D) F" using ccdf_cdf DERIV_const Deriv.field_differentiable_diff by fastforce next assume "(ccdf M has_real_derivative -D) F" thus "(cdf M has_real_derivative D) F" using cdf_ccdf DERIV_const Deriv.field_differentiable_diff by fastforce qed lemma differentiable_cdf_ccdf: "(cdf M differentiable F) \ (ccdf M differentiable F)" unfolding differentiable_def apply (rewrite has_real_derivative_iff[THEN sym])+ apply (rewrite has_real_derivative_cdf_ccdf) by (metis verit_minus_simplify(4)) lemma deriv_cdf_ccdf: assumes "cdf M differentiable at x" shows "deriv (cdf M) x = - deriv (ccdf M) x" using has_real_derivative_cdf_ccdf differentiable_cdf_ccdf assms by (simp add: DERIV_deriv_iff_real_differentiable DERIV_imp_deriv) lemma ccdf_diff_eq2: assumes "x \ y" shows "ccdf M x - ccdf M y = measure M {x<..y}" proof - have "ccdf M x - ccdf M y = cdf M y - cdf M x" using add_cdf_ccdf by (smt (verit)) also have "\ = measure M {x<..y}" using cdf_diff_eq2 assms by simp finally show ?thesis . qed lemma ccdf_nonincreasing: "x \ y \ ccdf M x \ ccdf M y" using add_cdf_ccdf cdf_nondecreasing by smt lemma ccdf_nonneg: "ccdf M x \ 0" using add_cdf_ccdf cdf_bounded by smt lemma ccdf_bounded: "ccdf M x \ measure M (space M)" using add_cdf_ccdf cdf_nonneg by smt lemma ccdf_lim_at_top: "(ccdf M \ 0) at_top" proof - have "((\x. measure M (space M) - cdf M x) \ measure M (space M) - measure M (space M)) at_top" apply (intro tendsto_intros) by (rule cdf_lim_at_top) thus ?thesis by (rewrite ccdf_cdf) simp qed lemma ccdf_lim_at_bot: "(ccdf M \ measure M (space M)) at_bot" proof - have "((\x. measure M (space M) - cdf M x) \ measure M (space M) - 0) at_bot" apply (intro tendsto_intros) by (rule cdf_lim_at_bot) thus ?thesis by (rewrite ccdf_cdf) simp qed lemma ccdf_is_right_cont: "continuous (at_right a) (ccdf M)" proof - have "continuous (at_right a) (\x. measure M (space M) - cdf M x)" apply (intro continuous_intros) by (rule cdf_is_right_cont) thus ?thesis by (rewrite ccdf_cdf) simp qed end context prob_space begin lemma ccdf_distr_measurable [measurable]: assumes [measurable]: "random_variable borel X" shows "ccdf (distr M borel X) \ borel_measurable borel" using real_distribution.finite_borel_measure_M by (rewrite finite_borel_measure.ccdf_cdf; simp) lemma ccdf_distr_P: assumes "random_variable borel X" shows "ccdf (distr M borel X) x = \

(\ in M. X \ > x)" unfolding ccdf_def apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob]) force lemma ccdf_continuous_distr_P_ge: assumes "random_variable borel X" "isCont (ccdf (distr M borel X)) x" shows "ccdf (distr M borel X) x = \

(\ in M. X \ \ x)" proof - have "ccdf (distr M borel X) x = measure (distr M borel X) {x<..}" unfolding ccdf_def by simp also have "\ = measure (distr M borel X) ({x<..} \ {x})" apply (rewrite finite_measure.measure_zero_union, simp_all add: assms finite_measure_distr) using finite_borel_measure.isCont_ccdf real_distribution.finite_borel_measure_M assms by blast also have "\ = measure (distr M borel X) {x..}" by (metis Un_commute ivl_disj_un_singleton(1)) also have "\ = \

(\ in M. X \ \ x)" apply (rewrite measure_distr, simp_all add: assms) unfolding vimage_def by simp (smt (verit) Collect_cong Int_def mem_Collect_eq) finally show ?thesis . qed lemma ccdf_distr_diff_P: assumes "x \ y" and "random_variable borel X" shows "ccdf (distr M borel X) x - ccdf (distr M borel X) y = \

(\ in M. x < X \ \ X \ \ y)" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp have "ccdf (distr M borel X) x - ccdf (distr M borel X) y = measure (distr M borel X) {x<..y}" by (rewrite distrX_FBM.ccdf_diff_eq2; simp add: assms) also have "\ = \

(\ in M. x < X \ \ X \ \ y)" apply (rewrite measure_distr; (simp add: assms)?) unfolding vimage_def by (rule arg_cong[where f=prob], force) finally show ?thesis . qed end context real_distribution begin lemma ccdf_bounded_prob: "\x. ccdf M x \ 1" by (subst prob_space[THEN sym], rule ccdf_bounded) lemma ccdf_lim_at_bot_prob: "(ccdf M \ 1) at_bot" by (subst prob_space[THEN sym], rule ccdf_lim_at_bot) end text \Introduce the hazard rate. This notion will be used to define the force of mortality.\ subsection \Hazard Rate\ context prob_space begin definition hazard_rate :: "('a \ real) \ real \ real" where "hazard_rate X t \ Lim (at_right 0) (\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt)" \ \Here, X is supposed to be a random variable.\ lemma hazard_rate_0_ccdf_0: assumes [measurable]: "random_variable borel X" and "ccdf (distr M borel X) t = 0" shows "hazard_rate X t = 0" \ \Note that division by 0 is calculated as 0 in Isabelle/HOL.\ proof - have "\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) = 0" unfolding cond_prob_def using ccdf_distr_P assms by simp hence "hazard_rate X t = Lim (at_right 0) (\dt::real. 0)" unfolding hazard_rate_def by (rewrite Lim_cong; simp) also have "\ = 0" by (rule tendsto_Lim; simp) finally show ?thesis . qed lemma hazard_rate_deriv_cdf: assumes [measurable]: "random_variable borel X" and "(cdf (distr M borel X)) differentiable at t" shows "hazard_rate X t = deriv (cdf (distr M borel X)) t / ccdf (distr M borel X) t" proof (cases \ccdf (distr M borel X) t = 0\) case True with hazard_rate_0_ccdf_0 show ?thesis by simp next case False let ?F = "cdf (distr M borel X)" have "\\<^sub>F dt in at_right 0. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt = (?F (t + dt) - ?F t) / dt / ccdf (distr M borel X) t" apply (rule eventually_at_rightI[where b=1]; simp) unfolding cond_prob_def apply (rewrite cdf_distr_diff_P; simp) apply (rewrite ccdf_distr_P[THEN sym], simp) by (smt (verit) Collect_cong mult.commute) moreover have "((\dt. (?F (t + dt) - ?F t) / dt / ccdf (distr M borel X) t) \ deriv ?F t / ccdf (distr M borel X) t) (at_right 0)" apply (rule tendsto_intros, simp_all add: False) apply (rule Lim_at_imp_Lim_at_within) using DERIV_deriv_iff_real_differentiable assms DERIV_def by blast ultimately show ?thesis unfolding hazard_rate_def using tendsto_cong by (intro tendsto_Lim; force) qed lemma deriv_cdf_hazard_rate: assumes [measurable]: "random_variable borel X" and "(cdf (distr M borel X)) differentiable at t" shows "deriv (cdf (distr M borel X)) t = ccdf (distr M borel X) t * hazard_rate X t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis proof (cases \ccdf (distr M borel X) t = 0\) case True hence "cdf (distr M borel X) t = 1" using distrX_FBM.cdf_ccdf by simp (metis assms(1) distrX_FBM.borel_UNIV prob_space.prob_space prob_space_distr) moreover obtain D where "(cdf (distr M borel X) has_real_derivative D) (at t)" using assms real_differentiable_def by atomize_elim blast ultimately have "(cdf (distr M borel X) has_real_derivative 0) (at t)" using assms by (smt (verit) DERIV_local_max real_distribution.cdf_bounded_prob real_distribution_distr) thus ?thesis using True by (simp add: DERIV_imp_deriv) next case False thus ?thesis using hazard_rate_deriv_cdf assms by simp qed qed lemma hazard_rate_deriv_ccdf: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" shows "hazard_rate X t = - deriv (ccdf (distr M borel X)) t / ccdf (distr M borel X) t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis using hazard_rate_deriv_cdf distrX_FBM.deriv_cdf_ccdf assms distrX_FBM.differentiable_cdf_ccdf by presburger qed lemma deriv_ccdf_hazard_rate: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" shows "deriv (ccdf (distr M borel X)) t = - ccdf (distr M borel X) t * hazard_rate X t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp show ?thesis using deriv_cdf_hazard_rate distrX_FBM.deriv_cdf_ccdf assms distrX_FBM.differentiable_cdf_ccdf by simp qed lemma hazard_rate_deriv_ln_ccdf: assumes [measurable]: "random_variable borel X" and "(ccdf (distr M borel X)) differentiable at t" and "ccdf (distr M borel X) t \ 0" shows "hazard_rate X t = - deriv (\t. ln (ccdf (distr M borel X) t)) t" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have "?srvl t > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit)) moreover have "(?srvl has_real_derivative (deriv ?srvl t)) (at t)" using DERIV_deriv_iff_real_differentiable assms by blast ultimately have "((\t. ln (?srvl t)) has_real_derivative 1 / ?srvl t * deriv ?srvl t) (at t)" by (rule derivative_intros) hence "deriv (\t. ln (?srvl t)) t = deriv ?srvl t / ?srvl t" by (simp add: DERIV_imp_deriv) also have "\ = - hazard_rate X t" using hazard_rate_deriv_ccdf assms by simp finally show ?thesis by simp qed lemma hazard_rate_has_integral: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "(hazard_rate X has_integral ln (ccdf (distr M borel X) t / ccdf (distr M borel X) u)) {t..u}" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have [simp]: "\s. t \ s \ s \ u \ ?srvl s > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit) atLeastAtMost_iff) have "(deriv (\s. - ln (?srvl s)) has_integral - ln (?srvl u) - - ln (?srvl t)) {t..u}" proof - have "continuous_on {t..u} (\s. - ln (?srvl s))" by (rule continuous_intros, rule continuous_on_ln, auto simp add: assms) moreover hence "(\s. - ln (?srvl s)) piecewise_differentiable_on {t<.. {0<..}" proof - { fix s assume "s \ {t<.. 0" using assms by simp moreover have "?srvl s \ 0" using distrX_FBM.ccdf_nonneg by simp ultimately have "?srvl s > 0" by simp } thus ?thesis by auto qed hence "(\r. - ln r) \ ?srvl piecewise_differentiable_on {t<..s. - ln (?srvl s)) has_integral ln (?srvl t / ?srvl u)) {t..u}" by simp (rewrite ln_div; force simp: assms) thus "((hazard_rate X) has_integral ln (?srvl t / ?srvl u)) {t..u}" proof - from assms obtain S0 where finS0: "finite S0" and diffS0: "\s. s \ {t<.. ?srvl differentiable at s within {t<..s. s \ {t..u} - S \ ?srvl differentiable at s" proof (atomize_elim) let ?S = "S0 \ {t, u}" have "finite ?S" using finS0 by simp moreover have "\s. s \ {t..u} - ?S \ ccdf (distr M borel X) differentiable at s" proof - { fix s assume s_in: "s \ {t..u} - ?S" hence "?srvl differentiable at s within {t<..S. finite S \ (\s. s \ {t..u} - S \ ccdf (distr M borel X) differentiable at s)" by blast qed thus ?thesis apply (rewrite has_integral_spike_finite_eq [of S _ "deriv (\s. - ln (?srvl s))"], simp) apply (rewrite hazard_rate_deriv_ln_ccdf, simp_all add: assms) apply (rewrite deriv_minus, simp_all) apply (rewrite in asm differentiable_eq_field_differentiable_real) apply (rewrite comp_def[THEN sym], rule field_differentiable_compose[of "?srvl"], simp_all) unfolding field_differentiable_def apply (rule exI, rule DERIV_ln, simp) using ln by simp qed qed corollary hazard_rate_integrable: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "hazard_rate X integrable_on {t..u}" using has_integral_integrable hazard_rate_has_integral assms by blast lemma ccdf_exp_cumulative_hazard: assumes [measurable]: "random_variable borel X" and "t \ u" and "(ccdf (distr M borel X)) piecewise_differentiable_on {t<..s. s \ {t..u} \ ccdf (distr M borel X) s \ 0" shows "ccdf (distr M borel X) u / ccdf (distr M borel X) t = exp (- integral {t..u} (hazard_rate X))" proof - interpret distrX_FBM: finite_borel_measure "distr M borel X" using real_distribution.finite_borel_measure_M real_distribution_distr assms by simp let ?srvl = "ccdf (distr M borel X)" have [simp]: "\s. t \ s \ s \ u \ ?srvl s > 0" using distrX_FBM.ccdf_nonneg assms by (smt (verit) atLeastAtMost_iff) have "integral {t..u} (hazard_rate X) = ln (?srvl t / ?srvl u)" using hazard_rate_has_integral has_integral_integrable_integral assms by auto also have "\ = - ln (?srvl u / ?srvl t)" using ln_div assms by simp finally have "- integral {t..u} (hazard_rate X) = ln (?srvl u / ?srvl t)" by simp thus ?thesis using assms by simp qed lemma hazard_rate_density_ccdf: assumes "distributed M lborel X f" and "\s. f s \ 0" "t < u" "continuous_on {t..u} f" shows "hazard_rate X t = f t / ccdf (distr M borel X) t" proof (cases \ccdf (distr M borel X) t = 0\) case True thus ?thesis apply (rewrite hazard_rate_0_ccdf_0, simp_all) using assms(1) distributed_measurable by force next case False have [simp]: "t \ u" using assms by simp have [measurable]: "random_variable borel X" using assms distributed_measurable measurable_lborel1 by blast have [measurable]: "f \ borel_measurable lborel" using assms distributed_real_measurable by metis have [simp]: "integrable lborel f" proof - - have "prob (X -` UNIV \ space M) = LINT x|lborel. indicat_real UNIV x * f x" + have "prob (X -` UNIV \ space M) = (LINT x|lborel. indicat_real UNIV x * f x)" by (rule distributed_measure; simp add: assms) thus ?thesis using prob_space not_integrable_integral_eq by fastforce qed have "((\dt. (LBINT s:{t..t+dt}. f s) / dt) \ f t) (at_right 0)" proof - have "\dt. (\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) < \" proof - fix dt :: real have "(\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) = set_nn_integral lborel {t..t+dt} f" by (metis indicator_mult_ennreal mult.commute) moreover have "emeasure M (X -` {t..t+dt} \ space M) = set_nn_integral lborel {t..t+dt} f" by (rule distributed_emeasure; simp add: assms) moreover have "emeasure M (X -` {t..t+dt} \ space M) < \" using emeasure_eq_measure ennreal_less_top infinity_ennreal_def by presburger ultimately show "(\\<^sup>+ x. ennreal (indicat_real {t..t+dt} x * f x) \lborel) < \" by simp qed - hence "\dt. LBINT s:{t..t+dt}. f s = integral {t..t+dt} f" + hence "\dt. (LBINT s:{t..t+dt}. f s) = integral {t..t+dt} f" apply (intro set_borel_integral_eq_integral) unfolding set_integrable_def apply (rule integrableI_nonneg; simp) by (rule AE_I2, simp add: assms) moreover have "((\dt. (integral {t..t+dt} f) / dt) \ f t) (at_right 0)" proof - have "((\x. integral {t..x} f) has_real_derivative f t) (at t within {t..u})" by (rule integral_has_real_derivative; simp add: assms) moreover have "(at t within {t..u}) = (at (t+0) within (+)t ` {0..u-t})" by simp ultimately have "((\x. integral {t..x} f) \ (+)t has_real_derivative f t) (at 0 within {0..u-t})" by (metis DERIV_at_within_shift_lemma) hence "((\dt. (integral {t..t+dt} f) / dt) \ f t) (at 0 within {0..u-t})" using has_field_derivative_iff by force thus ?thesis using at_within_Icc_at_right assms by smt qed ultimately show ?thesis by simp qed - moreover have "\dt. dt > 0 \ \

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t..t+dt}. f s" + moreover have "\dt. dt > 0 \ \

(x in M. X x \ {t <.. t+dt}) = (LBINT s:{t..t+dt}. f s)" proof - fix dt :: real assume "dt > 0" hence [simp]: "sym_diff {t<..t + dt} {t..t + dt} = {t}" by force have "prob (X -` {t<..t+dt} \ space M) = \s. indicator {t<..t+dt} s * f s \lborel" by (rule distributed_measure; simp add: assms) - hence "\

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t<..t+dt}. f s" + hence "\

(x in M. X x \ {t <.. t+dt}) = (LBINT s:{t<..t+dt}. f s)" unfolding set_lebesgue_integral_def vimage_def Int_def by simp (smt (verit) Collect_cong) - moreover have "LBINT s:{t<..t+dt}. f s = LBINT s:{t..t+dt}. f s" + moreover have "(LBINT s:{t<..t+dt}. f s) = (LBINT s:{t..t+dt}. f s)" by (rule set_integral_null_delta; force) - ultimately show "\

(x in M. X x \ {t <.. t+dt}) = LBINT s:{t..t+dt}. f s" by simp + ultimately show "\

(x in M. X x \ {t <.. t+dt}) = (LBINT s:{t..t+dt}. f s)" by simp qed ultimately have "((\dt. \

(x in M. t < X x \ X x \ t + dt) / dt) \ f t) (at_right 0)" by simp (smt (verit) Lim_cong_within greaterThan_iff) hence "((\dt. \

(x in M. t < X x \ X x \ t + dt \ X x > t) / dt) \ f t / ccdf (distr M borel X) t) (at_right 0)" unfolding cond_prob_def apply (rewrite ccdf_distr_P[THEN sym]; simp) apply (rewrite mult.commute, rewrite divide_divide_eq_left[THEN sym]) by (rule tendsto_intros; (simp add: False)?) (smt (verit) Collect_cong Lim_cong_within) thus ?thesis unfolding hazard_rate_def by (intro tendsto_Lim; simp) qed end end diff --git a/thys/Actuarial_Mathematics/Survival_Model.thy b/thys/Actuarial_Mathematics/Survival_Model.thy --- a/thys/Actuarial_Mathematics/Survival_Model.thy +++ b/thys/Actuarial_Mathematics/Survival_Model.thy @@ -1,2732 +1,2730 @@ theory Survival_Model imports "HOL-Library.Rewrite" "HOL-Library.Extended_Nonnegative_Real" "HOL-Library.Extended_Real" "HOL-Probability.Probability" Preliminaries begin section \Survival Model\ text \ The survival model is built on the probability space @{text "\"}. Additionally, the random variable @{text "X : space \ \ \"} is introduced, which represents the age at death. \ locale prob_space_actuary = MM_PS: prob_space \ for \ \ \Since the letter M may be used as a commutation function, adopt the letter @{text "\"} instead as a notation for the measure space.\ locale survival_model = prob_space_actuary + fixes X :: "'a \ real" assumes X_RV[simp]: "MM_PS.random_variable (borel :: real measure) X" and X_pos_AE[simp]: "AE \ in \. X \ > 0" begin subsection \General Theory of Survival Model\ interpretation distrX_RD: real_distribution "distr \ borel X" using MM_PS.real_distribution_distr by simp lemma X_le_event[simp]: "{\ \ space \. X \ \ x} \ MM_PS.events" by measurable simp lemma X_gt_event[simp]: "{\ \ space \. X \ > x} \ MM_PS.events" by measurable simp lemma X_compl_le_gt: "space \ - {\ \ space \. X \ \ x} = {\ \ space \. X \ > x}" for x::real proof - have "space \ - {\ \ space \. X \ \ x} = space \ - (X -` {..x})" by blast also have "\ = (X -` {x<..}) \ space \" using vimage_compl_atMost by fastforce also have "\ = {\ \ space \. X \ > x}" by blast finally show ?thesis . qed lemma X_compl_gt_le: "space \ - {\ \ space \. X \ > x} = {\ \ space \. X \ \ x}" for x::real using X_compl_le_gt by blast subsubsection \Introduction of Survival Function for X\ text \Note that @{text "ccdf (distr \ borel X)"} is the survival (distributive) function for X.\ lemma ccdfX_0_1: "ccdf (distr \ borel X) 0 = 1" apply (rewrite MM_PS.ccdf_distr_P, simp) using X_pos_AE MM_PS.prob_space using MM_PS.prob_Collect_eq_1 X_gt_event by presburger lemma ccdfX_unborn_1: "ccdf (distr \ borel X) x = 1" if "x \ 0" proof (rule antisym) show "ccdf (distr \ borel X) x \ 1" using MM_PS.ccdf_distr_P by simp show "ccdf (distr \ borel X) x \ 1" proof - have "ccdf (distr \ borel X) x \ ccdf (distr \ borel X) 0" using finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M that by simp also have "ccdf (distr \ borel X) 0 = 1" using ccdfX_0_1 that by simp finally show ?thesis . qed qed definition death_pt :: ereal ("$\") where "$\ \ Inf (ereal ` {x \ \. ccdf (distr \ borel X) x = 0})" \ \This is my original notation, which is used to develop life insurance mathematics rigorously.\ lemma psi_nonneg: "$\ \ 0" unfolding death_pt_def proof (rule Inf_greatest) fix x'::ereal assume "x' \ ereal ` {x \ \. ccdf (distr \ borel X) x = 0}" then obtain x::real where "x' = ereal x" and "ccdf (distr \ borel X) x = 0" by blast hence "ccdf (distr \ borel X) 0 > ccdf (distr \ borel X) x" using ccdfX_0_1 X_pos_AE by simp hence "x \ 0" using mono_invE finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M X_RV by (smt(verit)) thus "x' \ 0" using \x' = ereal x\ by simp qed lemma ccdfX_beyond_0: "ccdf (distr \ borel X) x = 0" if "x > $\" for x::real proof - have "ereal ` {y \ \. ccdf (distr \ borel X) y = 0} \ {}" using death_pt_def that by force hence "\y'\(ereal ` {y \ \. ccdf (distr \ borel X) y = 0}). y' < ereal x" using that unfolding death_pt_def by (rule cInf_lessD) then obtain "y'" where "y' \ (ereal ` {y \ \. ccdf (distr \ borel X) y = 0})" and "y' < ereal x" by blast then obtain y::real where "y' = ereal y" and "ccdf (distr \ borel X) y = 0" and "ereal y < ereal x" by blast hence "ccdf (distr \ borel X) y = 0" and "y < x" by simp_all hence "ccdf (distr \ borel X) x \ 0" using finite_borel_measure.ccdf_nonincreasing distrX_RD.finite_borel_measure_M X_RV by (metis order_less_le) thus ?thesis using finite_borel_measure.ccdf_nonneg distrX_RD.finite_borel_measure_M X_RV by smt qed lemma ccdfX_psi_0: "ccdf (distr \ borel X) (real_of_ereal $\) = 0" if "$\ < \" proof - have "\$\\ \ \" using that psi_nonneg by simp then obtain u::real where "$\ = ereal u" using ereal_real' by blast hence "real_of_ereal $\ = u" by simp moreover have "ccdf (distr \ borel X) u = 0" proof - have "\x::real. x \ u \ x \ {u<..} \ ccdf (distr \ borel X) x = 0" by (rule ccdfX_beyond_0, simp add: \$\ = ereal u\) hence "(ccdf (distr \ borel X) \ 0) (at_right u)" apply - by (rule iffD2[OF Lim_cong_within[where ?g="(\x.0)"]], simp_all+) moreover have "(ccdf (distr \ borel X) \ ccdf (distr \ borel X) u) (at_right u)" using finite_borel_measure.ccdf_is_right_cont distrX_RD.finite_borel_measure_M continuous_within X_RV by blast ultimately show ?thesis using tendsto_unique trivial_limit_at_right_real by blast qed ultimately show ?thesis by simp qed lemma ccdfX_0_equiv: "ccdf (distr \ borel X) x = 0 \ x \ $\" for x::real proof assume "ccdf (distr \ borel X) x = 0" thus "ereal x \ $\" unfolding death_pt_def by (simp add: INF_lower) next assume "$\ \ ereal x" hence "$\ = ereal x \ $\ < ereal x" unfolding less_eq_ereal_def by auto thus "ccdf (distr \ borel X) x = 0" proof assume \: "$\ = ereal x" hence "$\ < \" by simp moreover have "real_of_ereal $\ = x" using \ by simp ultimately show "ccdf (distr \ borel X) x = 0" using ccdfX_psi_0 by simp next assume "$\ < ereal x" thus "ccdf (distr \ borel X) x = 0" by (rule ccdfX_beyond_0) qed qed lemma psi_pos[simp]: "$\ > 0" proof (rule not_le_imp_less, rule notI) show "$\ \ (0::ereal) \ False" proof - assume "$\ \ (0::ereal)" hence "ccdf (distr \ borel X) 0 = 0" using ccdfX_0_equiv by (simp add: zero_ereal_def) moreover have "ccdf (distr \ borel X) 0 = 1" using ccdfX_0_1 by simp ultimately show "False" by simp qed qed corollary psi_pos'[simp]: "$\ > ereal 0" using psi_pos zero_ereal_def by presburger subsubsection \Introdution of Future-Lifetime Random Variable T(x)\ definition alive :: "real \ 'a set" where "alive x \ {\ \ space \. X \ > x}" lemma alive_event[simp]: "alive x \ MM_PS.events" for x::real unfolding alive_def by simp lemma X_alivex_measurable[measurable, simp]: "X \ borel_measurable (\ \ alive x)" for x::real unfolding cond_prob_space_def by (measurable; simp add: measurable_restrict_space1) definition futr_life :: "real \ ('a \ real)" ("T") where "T x \ (\\. X \ - x)" \ \Note that @{text "T(x) : space \ \ \"} represents the time until death of a person aged x.\ lemma T0_eq_X[simp]: "T 0 = X" unfolding futr_life_def by simp lemma Tx_measurable[measurable, simp]: "T x \ borel_measurable \" for x::real unfolding futr_life_def by (simp add: borel_measurable_diff) lemma Tx_alivex_measurable[measurable, simp]: "T x \ borel_measurable (\ \ alive x)" for x::real unfolding futr_life_def by (simp add: borel_measurable_diff) lemma alive_T: "alive x = {\ \ space \. T x \ > 0}" for x::real unfolding alive_def futr_life_def by force lemma alivex_Tx_pos[simp]: "0 < T x \" if "\ \ space (\ \ alive x)" for x::real using MM_PS.space_cond_prob_space alive_T that by simp lemma PT0_eq_PX_lborel: "\

(\ in \. T 0 \ \ A \ T 0 \ > 0) = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" apply (rewrite MM_PS.cond_prob_AE_prob, simp_all) using that X_RV measurable_lborel1 predE pred_sets2 by blast subsubsection \Actuarial Notations on the Survival Model\ definition survive :: "real \ real \ real" ("$p'_{_&_}" [0,0] 200) where "$p_{t&x} \ ccdf (distr (\ \ alive x) borel (T x)) t" \ \the probability that a person aged x will survive for t years\ \ \Note that the function @{text "$p_{\&x}"} is the survival function on @{text "(\ \ alive x)"} for the random variable T(x).\ \ \The parameter t is usually nonnegative, but theoretically it can be negative.\ abbreviation survive_1 :: "real \ real" ("$p'__" [101] 200) where "$p_x \ $p_{1&x}" definition die :: "real \ real \ real" ("$q'_{_&_}" [0,0] 200) where "$q_{t&x} \ cdf (distr (\ \ alive x) borel (T x)) t" \ \the probability that a person aged x will die within t years\ \ \Note that the function @{text "$q_{\&x}"} is the cumulative distributive function on @{text "(\ \ alive x)"} for the random variable T(x).\ \ \The parameter t is usually nonnegative, but theoretically it can be negative.\ abbreviation die_1 :: "real \ real" ("$q'__" [101] 200) where "$q_x \ $q_{1&x}" definition die_defer :: "real \ real \ real \ real" ("$q'_{_\_&_}" [0,0,0] 200) where "$q_{f\t&x} = \$q_{f+t&x} - $q_{f&x}\" \ \the probability that a person aged x will die within t years, deferred f years\ \ \The parameters f and t are usually nonnegative, but theoretically they can be negative.\ abbreviation die_defer_1 :: "real \ real \ real" ("$q'_{_\&_}" [0,0] 200) where "$q_{f\&x} \ $q_{f\1&x}" definition life_expect :: "real \ real" ("$e`\'__" [101] 200) where "$e`\_x \ integral\<^sup>L (\ \ alive x) (T x)" \ \complete life expectation\ \ \Note that @{text "$e`\_x"} is calculated as 0 when @{text "nn_integral (\ \ alve x) (T x) = \"}.\ definition temp_life_expect :: "real \ real \real" ("$e`\'_{_:_}" [0,0] 200) where "$e`\_{x:n} \ integral\<^sup>L (\ \ alive x) (\\. min (T x \) n)" \ \temporary complete life expectation\ definition curt_life_expect :: "real \ real" ("$e'__" [101] 200) where "$e_x \ integral\<^sup>L (\ \ alive x) (\\. \T x \\)" \ \curtate life expectation\ \ \Note that @{text "$e_x"} is calculated as 0 when @{text "nn_integral (\ \ alive x) \T x\ = \"}.\ definition temp_curt_life_expect :: "real \ real \ real" ("$e'_{_:_}" [0,0] 200) where "$e_{x:n} \ integral\<^sup>L (\ \ alive x) (\\. \min (T x \) n\)" \ \temporary curtate life expectation\ \ \In the definition n can be a real number, but in practice n is usually a natural number.\ subsubsection \Properties of Survival Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma PXx_pos[simp]: "\

(\ in \. X \ > x) > 0" proof - have "\

(\ in \. X \ > x) = ccdf (distr \ borel X) x" unfolding alive_def using MM_PS.ccdf_distr_P by simp also have "\ > 0" using ccdfX_0_equiv distrX_RD.ccdf_nonneg x_lt_psi by (smt (verit) linorder_not_le) finally show ?thesis . qed lemma PTx_pos[simp]: "\

(\ in \. T x \ > 0) > 0" apply (rewrite alive_T[THEN sym]) unfolding alive_def by simp interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma ccdfTx_cond_prob: "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. T x \ > t \ T x \ > 0)" for t::real apply (rewrite alivex_PS.ccdf_distr_P, simp) unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], simp_all add: pred_def) unfolding futr_life_def by simp lemma ccdfTx_0_1: "ccdf (distr (\ \ alive x) borel (T x)) 0 = 1" apply (rewrite ccdfTx_cond_prob) unfolding futr_life_def cond_prob_def by (smt (verit, best) Collect_cong PXx_pos divide_eq_1_iff) lemma ccdfTx_nonpos_1: "ccdf (distr (\ \ alive x) borel (T x)) t = 1" if "t \ 0" for t :: real using antisym ccdfTx_0_1 that by (metis distrTx_RD.ccdf_bounded_prob distrTx_RD.ccdf_nonincreasing) lemma ccdfTx_0_equiv: "ccdf (distr (\ \ alive x) borel (T x)) t = 0 \ x+t \ $\" for t::real proof - have "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. X \ > x+t \ X \ > x) / \

(\ in \. X \ > x)" apply (rewrite ccdfTx_cond_prob) unfolding cond_prob_def futr_life_def by (smt (verit) Collect_cong) hence "ccdf (distr (\ \ alive x) borel (T x)) t = 0 \ \

(\ in \. X \ > x+t \ X \ > x) / \

(\ in \. X \ > x) = 0" by simp also have "\ \ \

(\ in \. X \ > x+t \ X \ > x) = 0" using x_lt_psi PXx_pos by (smt (verit) divide_eq_0_iff) also have "\ \ x+t \ $\" using ccdfX_0_equiv MM_PS.ccdf_distr_P by (smt (verit) Collect_cong X_RV le_ereal_le linorder_not_le x_lt_psi) finally show ?thesis . qed lemma ccdfTx_continuous_on_nonpos[simp]: "continuous_on {..0} (ccdf (distr (\ \ alive x) borel (T x)))" by (metis atMost_iff ccdfTx_nonpos_1 continuous_on_cong continuous_on_const) lemma ccdfTx_differentiable_on_nonpos[simp]: "(ccdf (distr (\ \ alive x) borel (T x))) differentiable_on {..0}" by (rewrite differentiable_on_cong[where f="\_. 1"]; simp add: ccdfTx_nonpos_1) lemma ccdfTx_has_real_derivative_0_at_neg: "(ccdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" if "t < 0" for t::real apply (rewrite has_real_derivative_iff_has_vector_derivative) apply (rule has_vector_derivative_transform_within_open[of "\_. 1" _ _ "{..<0}"]) using ccdfTx_nonpos_1 that by simp_all lemma ccdfTx_integrable_Icc: "set_integrable lborel {a..b} (ccdf (distr (\ \ alive x) borel (T x)))" for a b :: real proof - have "(\\<^sup>+t. ennreal (indicat_real {a..b} t * ccdf (distr (\ \ alive x) borel (T x)) t) \lborel) < \" proof - have "(\\<^sup>+t. ennreal (indicat_real {a..b} t * ccdf (distr (\ \ alive x) borel (T x)) t) \lborel) \ (\\<^sup>+t. ennreal (indicat_real {a..b} t) \lborel)" apply (rule nn_integral_mono) using distrTx_RD.ccdf_bounded by (simp add: distrTx_RD.ccdf_bounded_prob indicator_times_eq_if(1)) also have "\ = nn_integral lborel (indicator {a..b})" by (meson ennreal_indicator) also have "\ = emeasure lborel {a..b}" by (rewrite nn_integral_indicator; simp) also have "\ < \" using emeasure_lborel_Icc_eq ennreal_less_top infinity_ennreal_def by presburger finally show ?thesis . qed thus ?thesis unfolding set_integrable_def apply (intro integrableI_nonneg, simp_all) using distrTx_RD.ccdf_nonneg by (intro always_eventually) auto qed corollary ccdfTx_integrable_on_Icc: "ccdf (distr (\ \ alive x) borel (T x)) integrable_on {a..b}" for a b :: real using set_borel_integral_eq_integral ccdfTx_integrable_Icc by force lemma ccdfTx_PX: "ccdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. X \ > x+t) / \

(\ in \. X \ > x)" if "t \ 0" for t::real apply (rewrite ccdfTx_cond_prob) unfolding cond_prob_def futr_life_def PXx_pos by (smt (verit) Collect_cong that) lemma ccdfTx_ccdfX: "ccdf (distr (\ \ alive x) borel (T x)) t = ccdf (distr \ borel X) (x + t) / ccdf (distr \ borel X) x" if "t \ 0" for t::real using ccdfTx_PX that MM_PS.ccdf_distr_P X_RV by presburger lemma ccdfT0_eq_ccdfX: "ccdf (distr (\ \ alive 0) borel (T 0)) = ccdf (distr \ borel X)" proof fix x show "ccdf (distr (\ \ alive 0) borel (T 0)) x = ccdf (distr \ borel X) x" proof (cases \x \ 0\) case True thus ?thesis using survival_model.ccdfTx_ccdfX[where x=0] ccdfX_0_1 survival_model_axioms by fastforce next case False hence "x \ 0" by simp thus ?thesis apply (rewrite ccdfX_unborn_1, simp) by (rewrite survival_model.ccdfTx_nonpos_1; simp add: survival_model_axioms) qed qed lemma continuous_ccdfX_ccdfTx: "continuous (at (x+t) within {x..}) (ccdf (distr \ borel X)) \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" if "t \ 0" for t::real proof - let ?srvl = "ccdf (distr \ borel X)" have [simp]: "\

(\ in \. X \ > x) \ 0" using PXx_pos by force have \: "\u. u \ 0 \ ccdf (distr (\ \ alive x) borel (T x)) u = ?srvl (x + u) / \

(\ in \. X \ > x)" using survive_def MM_PS.ccdf_distr_P ccdfTx_PX that by simp have "continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x))) \ continuous (at t within {0..}) (\u. ?srvl (x+u) / \

(\ in \. x < X \))" proof - have "\\<^sub>F u in at t within {0..}. ccdf (distr (\ \ alive x) borel (T x)) u = ?srvl (x+u) / \

(\ in \. X \ > x)" using \ by (rewrite eventually_at_topological, simp_all) blast thus ?thesis by (intro continuous_at_within_cong, simp_all add: \ that) qed also have "\ \ continuous (at t within {0..}) (\u. ?srvl (x+u))" by (rewrite at "_ = \" continuous_cdivide_iff[of "\

(\ in \. X \ > x)"], simp_all) also have "\ \ continuous (at (x+t) within {x..}) ?srvl" proof let ?subx = "\v. v-x" assume LHS: "continuous (at t within {0..}) (\u. ?srvl (x+u))" hence "continuous (at (?subx (x+t)) within ?subx ` {x..}) (\u. ?srvl (x+u))" proof - have "?subx ` {x..} = {0..}" by (metis (no_types, lifting) add.commute add_uminus_conv_diff diff_self image_add_atLeast image_cong) thus ?thesis using LHS by simp qed moreover have "continuous (at (x+t) within {x..}) ?subx" by (simp add: continuous_diff) ultimately have "continuous (at (x+t) within {x..}) (\u. ?srvl (x + (?subx u)))" using continuous_within_compose2 by force thus "continuous (at (x+t) within {x..}) ?srvl" by simp next assume RHS: "continuous (at (x+t) within {x..}) ?srvl" hence "continuous (at ((plus x) t) within (plus x) ` {0..}) ?srvl" by simp moreover have "continuous (at t within {0..}) (plus x)" by (simp add: continuous_add) ultimately show "continuous (at t within {0..}) (\u. ?srvl (x+u))" using continuous_within_compose2 by force qed finally show ?thesis by simp qed lemma isCont_ccdfX_ccdfTx: "isCont (ccdf (distr \ borel X)) (x+t) \ isCont (ccdf (distr (\ \ alive x) borel (T x))) t" if "t > 0" for t::real proof - have "isCont (ccdf (distr \ borel X)) (x+t) \ continuous (at (x+t) within {x<..}) (ccdf (distr \ borel X))" by (smt (verit) at_within_open greaterThan_iff open_greaterThan that) also have "\ \ continuous (at (x+t) within {x..}) (ccdf (distr \ borel X))" by (meson Ioi_le_Ico calculation continuous_within_subset top_greatest) also have "\ \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" using that continuous_ccdfX_ccdfTx by force also have "\ \ continuous (at t within {0<..}) (ccdf (distr (\ \ alive x) borel (T x)))" by (metis Ioi_le_Ico at_within_open continuous_at_imp_continuous_at_within continuous_within_subset greaterThan_iff open_greaterThan that) also have "\ \ isCont (ccdf (distr (\ \ alive x) borel (T x))) t" by (metis at_within_open greaterThan_iff open_greaterThan that) finally show ?thesis . qed lemma has_real_derivative_ccdfX_ccdfTx: "((ccdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ ((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" if "t > 0" for t D :: real proof - have "((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t) \ ((\t. (ccdf (distr \ borel X)) (x+t) / \

(\ in \. X \ > x)) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" proof - let ?d = "t/2" { fix u::real assume "dist u t < ?d" hence "u > 0" by (smt (verit) dist_real_def dist_triangle_half_r) hence "ccdf (distr (\ \ alive x) borel (T x)) u = ccdf (distr \ borel X) (x + u) / MM_PS.prob {\::'a \ space \. x < X \}" using survive_def MM_PS.ccdf_distr_P ccdfTx_PX that by simp } moreover have "?d > 0" using that by simp ultimately show ?thesis apply - apply (rule DERIV_cong_ev, simp) apply (rewrite eventually_nhds_metric, blast) by simp qed also have "\ \ ((\t. (ccdf (distr \ borel X)) (x+t)) has_real_derivative D) (at t)" using PXx_pos by (rewrite DERIV_cdivide_iff[of "\

(\ in \. X \ > x)", THEN sym]; force) also have "\ \ (ccdf (distr \ borel X) has_real_derivative D) (at (x+t))" by (simp add: DERIV_shift add.commute) finally show ?thesis by simp qed lemma differentiable_ccdfX_ccdfTx: "(ccdf (distr \ borel X)) differentiable at (x+t) \ (ccdf (distr (\ \ alive x) borel (T x))) differentiable at t" if "t > 0" for t::real apply (rewrite differentiable_eq_field_differentiable_real)+ unfolding field_differentiable_def using has_real_derivative_ccdfX_ccdfTx that by (smt (verit, del_insts) PXx_pos nonzero_mult_div_cancel_left) subsubsection \Properties of @{text "$p_{t&x}"}\ lemma p_0_1: "$p_{0&x} = 1" unfolding survive_def using ccdfTx_0_1 by simp lemma p_nonneg[simp]: "$p_{t&x} \ 0" for t::real unfolding survive_def using distrTx_RD.ccdf_nonneg by simp lemma p_0_equiv: "$p_{t&x} = 0 \ x+t \ $\" for t::real unfolding survive_def by (rule ccdfTx_0_equiv) lemma p_PTx: "$p_{t&x} = \

(\ in \. T x \ > t \ T x \ > 0)" for t::real unfolding survive_def using ccdfTx_cond_prob by simp lemma p_PX: "$p_{t&x} = \

(\ in \. X \ > x + t) / \

(\ in \. X \ > x)" if "t \ 0" for t::real unfolding survive_def using ccdfTx_PX that by simp lemma p_mult: "$p_{t+t' & x} = $p_{t&x} * $p_{t' & x+t}" if "t \ 0" "t' \ 0" "x+t < $\" for t t' :: real proof - have "$p_{t+t' & x} = \

(\ in \. X \ > x+t+t') / \

(\ in \. X \ > x)" apply (rewrite p_PX; (simp add: that)?) by (rule disjI2, smt (verit, best) Collect_cong) also have "\ = (\

(\ in \. X \ > x+t+t') / \

(\ in \. X \ > x+t)) * (\

(\ in \. X \ > x+t) / \

(\ in \. X \ > x))" using that survival_model.PXx_pos survival_model_axioms by fastforce also have "\ = $p_{t&x} * $p_{t' & x+t}" apply (rewrite p_PX, simp add: that) by (rewrite survival_model.p_PX, simp_all add: that survival_model_axioms) finally show ?thesis . qed lemma p_PTx_ge_ccdf_isCont: "$p_{t&x} = \

(\ in \. T x \ \ t \ T x \ > 0)" if "isCont (ccdf (distr \ borel X)) (x+t)" "t > 0" for t::real unfolding survive_def using that isCont_ccdfX_ccdfTx apply (rewrite alivex_PS.ccdf_continuous_distr_P_ge, simp_all) by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: alive_T) end subsubsection \Properties of Survival Function for X\ lemma ccdfX_continuous_unborn[simp]: "continuous_on {..0} (ccdf (distr \ borel X))" using ccdfTx_continuous_on_nonpos by (metis ccdfT0_eq_ccdfX psi_pos') lemma ccdfX_differentiable_unborn[simp]: "(ccdf (distr \ borel X)) differentiable_on {..0}" using ccdfTx_differentiable_on_nonpos by (metis ccdfT0_eq_ccdfX psi_pos') lemma ccdfX_has_real_derivative_0_unborn: "(ccdf (distr \ borel X) has_real_derivative 0) (at x)" if "x < 0" for x::real using ccdfTx_has_real_derivative_0_at_neg by (metis ccdfT0_eq_ccdfX psi_pos' that) lemma ccdfX_integrable_Icc: "set_integrable lborel {a..b} (ccdf (distr \ borel X))" for a b :: real using ccdfTx_integrable_Icc by (metis ccdfT0_eq_ccdfX psi_pos') corollary ccdfX_integrable_on_Icc: "ccdf (distr \ borel X) integrable_on {a..b}" for a b :: real using set_borel_integral_eq_integral ccdfX_integrable_Icc by force lemma ccdfX_p: "ccdf (distr \ borel X) x = $p_{x&0}" for x::real by (metis ccdfT0_eq_ccdfX survive_def psi_pos') subsubsection \Introduction of Cumulative Distributive Function for X\ lemma cdfX_0_0: "cdf (distr \ borel X) 0 = 0" using ccdfX_0_1 distrX_RD.ccdf_cdf distrX_RD.prob_space by fastforce lemma cdfX_unborn_0: "cdf (distr \ borel X) x = 0" if "x \ 0" using ccdfX_unborn_1 cdfX_0_0 distrX_RD.cdf_ccdf that by fastforce lemma cdfX_beyond_1: "cdf (distr \ borel X) x = 1" if "x > $\" for x::real using ccdfX_beyond_0 distrX_RD.cdf_ccdf that distrX_RD.prob_space by force lemma cdfX_psi_1: "cdf (distr \ borel X) (real_of_ereal $\) = 1" if "$\ < \" using ccdfX_psi_0 distrX_RD.cdf_ccdf distrX_RD.prob_space that by fastforce lemma cdfX_1_equiv: "cdf (distr \ borel X) x = 1 \ x \ $\" for x::real using ccdfX_0_equiv distrX_RD.cdf_ccdf distrX_RD.prob_space by force subsubsection \Properties of Cumulative Distributive Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma cdfTx_cond_prob: "cdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. T x \ \ t \ T x \ > 0)" for t::real apply (rewrite distrTx_RD.cdf_ccdf, rewrite distrTx_RD.prob_space) apply (rewrite ccdfTx_cond_prob, simp) by (rewrite not_less[THEN sym], rewrite MM_PS.cond_prob_neg; simp) lemma cdfTx_0_0: "cdf (distr (\ \ alive x) borel (T x)) 0 = 0" using ccdfTx_0_1 distrTx_RD.cdf_ccdf distrTx_RD.prob_space by force lemma cdfTx_nonpos_0: "cdf (distr (\ \ alive x) borel (T x)) t = 0" if "t \ 0" for t :: real using ccdfTx_nonpos_1 distrTx_RD.cdf_ccdf distrTx_RD.prob_space that by force lemma cdfTx_1_equiv: "cdf (distr (\ \ alive x) borel (T x)) t = 1 \ x+t \ $\" for t::real using ccdfTx_0_equiv distrTx_RD.cdf_ccdf distrTx_RD.prob_space by force lemma cdfTx_continuous_on_nonpos[simp]: "continuous_on {..0} (cdf (distr (\ \ alive x) borel (T x)))" by (rewrite continuous_on_cong[where g="\t. 0"]) (simp_all add: cdfTx_nonpos_0)+ lemma cdfTx_differentiable_on_nonpos[simp]: "(cdf (distr (\ \ alive x) borel (T x))) differentiable_on {..0}" by (rewrite differentiable_on_cong[where f="\t. 0"]; simp add: cdfTx_nonpos_0) lemma cdfTx_has_real_derivative_0_at_neg: "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" if "t < 0" for t::real apply (rewrite has_real_derivative_iff_has_vector_derivative) apply (rule has_vector_derivative_transform_within_open[of "\_. 0" _ _ "{..<0}"]) using cdfTx_nonpos_0 that by simp_all lemma cdfTx_integrable_Icc: "set_integrable lborel {a..b} (cdf (distr (\ \ alive x) borel (T x)))" for a b :: real proof - have "set_integrable lborel {a..b} (\_. 1::real)" unfolding set_integrable_def using emeasure_compact_finite by (simp, intro integrable_real_indicator; force) thus ?thesis apply (rewrite distrTx_RD.cdf_ccdf, rewrite distrTx_RD.prob_space) using ccdfTx_integrable_Icc by (rewrite set_integral_diff; simp) qed corollary cdfTx_integrable_on_Icc: "cdf (distr (\ \ alive x) borel (T x)) integrable_on {a..b}" for a b :: real using cdfTx_integrable_Icc set_borel_integral_eq_integral by force lemma cdfTx_PX: "cdf (distr (\ \ alive x) borel (T x)) t = \

(\ in \. x < X \ \ X \ \ x+t) / \

(\ in \. X \ > x)" for t::real apply (rewrite cdfTx_cond_prob) unfolding cond_prob_def futr_life_def PXx_pos by (smt (verit) Collect_cong) lemma cdfT0_eq_cdfX: "cdf (distr (\ \ alive 0) borel (T 0)) = cdf (distr \ borel X)" proof interpret alive0_PS: prob_space "\ \ alive 0" apply (rule MM_PS.cond_prob_space_correct, simp) using PXx_pos alive_def psi_pos' by presburger interpret distrT0_RD: real_distribution "distr (\ \ alive 0) borel (T 0)" by simp show "\x. cdf (distr (\ \ alive 0) borel (T 0)) x = cdf (distr \ borel X) x" using ccdfT0_eq_ccdfX distrX_RD.ccdf_cdf distrT0_RD.ccdf_cdf by (smt (verit, best) distrT0_RD.prob_space distrX_RD.prob_space psi_pos') qed lemma continuous_cdfX_cdfTx: "continuous (at (x+t) within {x..}) (cdf (distr \ borel X)) \ continuous (at t within {0..}) (cdf (distr (\ \ alive x) borel (T x)))" if "t \ 0" for t::real proof - have "continuous (at (x+t) within {x..}) (cdf (distr \ borel X)) \ continuous (at (x+t) within {x..}) (ccdf (distr \ borel X))" by (rule distrX_RD.continuous_cdf_ccdf) also have "\ \ continuous (at t within {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" using continuous_ccdfX_ccdfTx that by simp also have "\ \ continuous (at t within {0..}) (cdf (distr (\ \ alive x) borel (T x)))" using distrTx_RD.continuous_cdf_ccdf by simp finally show ?thesis . qed lemma isCont_cdfX_cdfTx: "isCont (cdf (distr \ borel X)) (x+t) \ isCont (cdf (distr (\ \ alive x) borel (T x))) t" if "t > 0" for t::real apply (rewrite distrX_RD.isCont_cdf_ccdf) apply (rewrite isCont_ccdfX_ccdfTx, simp_all add: that) by (rule distrTx_RD.isCont_cdf_ccdf[THEN sym]) lemma has_real_derivative_cdfX_cdfTx: "((cdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ ((cdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" if "t > 0" for t D :: real proof - have "((cdf (distr \ borel X)) has_real_derivative D) (at (x+t)) \ (ccdf (distr \ borel X) has_real_derivative -D) (at (x+t))" using distrX_RD.has_real_derivative_cdf_ccdf by force also have "\ \ ((ccdf (distr (\ \ alive x) borel (T x))) has_real_derivative (-D / \

(\ in \. X \ > x))) (at t)" using has_real_derivative_ccdfX_ccdfTx that by simp also have "\ \ ((cdf (distr (\ \ alive x) borel (T x))) has_real_derivative (D / \

(\ in \. X \ > x))) (at t)" by (simp add: distrTx_RD.has_real_derivative_cdf_ccdf) finally show ?thesis . qed lemma differentiable_cdfX_cdfTx: "(cdf (distr \ borel X)) differentiable at (x+t) \ (cdf (distr (\ \ alive x) borel (T x))) differentiable at t" if "t > 0" for t::real apply (rewrite differentiable_eq_field_differentiable_real)+ unfolding field_differentiable_def using has_real_derivative_cdfX_cdfTx that by (meson differentiable_ccdfX_ccdfTx distrTx_RD.finite_borel_measure_axioms distrX_RD.finite_borel_measure_axioms finite_borel_measure.differentiable_cdf_ccdf real_differentiable_def x_lt_psi) subsubsection \Properties of @{text "$q_{t&x}"}\ lemma q_0_0: "$q_{0&x} = 0" unfolding die_def using cdfTx_0_0 by simp lemma q_nonneg[simp]: "$q_{t&x} \ 0" for t::real unfolding die_def using distrTx_RD.cdf_nonneg by simp lemma q_1_equiv: "$q_{t&x} = 1 \ x+t \ $\" for t::real unfolding die_def using cdfTx_1_equiv by simp lemma q_PTx: "$q_{t&x} = \

(\ in \. T x \ \ t \ T x \ > 0)" for t::real unfolding die_def using cdfTx_cond_prob by simp lemma q_PX: "$q_{t&x} = \

(\ in \. x < X \ \ X \ \ x + t) / \

(\ in \. X \ > x)" unfolding die_def using cdfTx_PX by simp lemma q_defer_0_q[simp]: "$q_{0\t&x} = $q_{t&x}" for t::real unfolding die_defer_def using q_0_0 by simp lemma q_defer_0_0: "$q_{f\0&x} = 0" for f::real unfolding die_defer_def by simp lemma q_defer_nonneg[simp]: "$q_{f\t&x} \ 0" for f t :: real unfolding die_defer_def by simp lemma q_defer_q: "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" if "t \ 0" for f t :: real unfolding die_defer_def die_def using distrTx_RD.cdf_nondecreasing that by simp lemma q_defer_PTx: "$q_{f\t&x} = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0)" if "t \ 0" for f t :: real proof - have "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" using q_defer_q that by simp also have "\ = \

(\ in \. T x \ \ f + t \ T x \ > 0) - \

(\ in \. T x \ \ f \ T x \ > 0)" using q_PTx by simp also have "\ = \

(\ in (\ \ alive x). T x \ \ f + t) - \

(\ in (\ \ alive x). T x \ \ f)" using MM_PS.cond_prob_space_cond_prob alive_T by simp also have "\ = \

(\ in (\ \ alive x). f < T x \ \ T x \ \ f + t)" proof - have "{\ \ space (\ \ alive x). T x \ \ f + t} - {\ \ space (\ \ alive x). T x \ \ f} = {\ \ space (\ \ alive x). f < T x \ \ T x \ \ f + t}" using that by force hence "alivex_PS.prob ({\ \ space (\ \ alive x). T x \ \ f + t} - {\ \ space (\ \ alive x). T x \ \ f}) = \

(\ in (\ \ alive x). f < T x \ \ T x \ \ f + t)" by simp moreover have "{\ \ space (\ \ alive x). T x \ \ f} \ {\ \ space (\ \ alive x). T x \ \ f + t}" using that by force ultimately show ?thesis by (rewrite alivex_PS.finite_measure_Diff[THEN sym]; simp) qed also have "\ = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0)" using MM_PS.cond_prob_space_cond_prob alive_T by simp finally show ?thesis . qed lemma q_defer_PX: "$q_{f\t&x} = \

(\ in \. x + f < X \ \ X \ \ x + f + t) / \

(\ in \. X \ > x)" if "f \ 0" "t \ 0" for f t :: real proof - have "$q_{f\t&x} = \

(\ in \. f < T x \ \ T x \ \ f + t \ T x \ > 0) / \

(\ in \. T x \ > 0)" apply (rewrite q_defer_PTx; (simp add: that)?) unfolding cond_prob_def by simp also have "\ = \

(\ in \. f < T x \ \ T x \ \ f + t) / \

(\ in \. T x \ > 0)" proof - have "\\. \ \ space \ \ f < T x \ \ T x \ \ f + t \ T x \ > 0 \ f < T x \ \ T x \ \ f + t" using that by auto hence "{\ \ space \. f < T x \ \ T x \ \ f + t \ T x \ > 0} = {\ \ space \. f < T x \ \ T x \ \ f + t}" by blast thus ?thesis by simp qed also have "\ = \

(\ in \. x + f < X \ \ X \ \ x + f + t) / \

(\ in \. X \ > x)" unfolding futr_life_def by (smt (verit) Collect_cong) finally show ?thesis . qed lemma q_defer_old_0: "$q_{f\t&x} = 0" if "x+f \ $\" "t \ 0" for f t :: real proof - have "$q_{f\t&x} = $q_{f+t & x} - $q_{f&x}" using q_defer_q that by simp moreover have "$q_{f+t & x} = 1" using q_1_equiv that le_ereal_le by auto moreover have "$q_{f&x} = 1" using q_1_equiv that by simp ultimately show ?thesis by simp qed end subsubsection \Properties of Cumulative Distributive Function for X\ lemma cdfX_continuous_unborn[simp]: "continuous_on {..0} (cdf (distr \ borel X))" using cdfTx_continuous_on_nonpos by (metis cdfT0_eq_cdfX psi_pos') lemma cdfX_differentiable_unborn[simp]: "(cdf (distr \ borel X)) differentiable_on {..0}" using cdfTx_differentiable_on_nonpos by (metis cdfT0_eq_cdfX psi_pos') lemma cdfX_has_real_derivative_0_unborn: "(cdf (distr \ borel X) has_real_derivative 0) (at x)" if "x < 0" for x::real using cdfTx_has_real_derivative_0_at_neg by (metis cdfT0_eq_cdfX psi_pos' that) lemma cdfX_integrable_Icc: "set_integrable lborel {a..b} (cdf (distr \ borel X))" for a b :: real using cdfTx_integrable_Icc by (metis cdfT0_eq_cdfX psi_pos') corollary cdfX_integrable_on_Icc: "cdf (distr \ borel X) integrable_on {a..b}" for a b :: real using cdfX_integrable_Icc set_borel_integral_eq_integral by force lemma cdfX_q: "cdf (distr \ borel X) x = $q_{x&0}" if "x \ 0" for x::real by (metis cdfT0_eq_cdfX die_def psi_pos') subsubsection \Relations between @{text "$p_{t&x}"} and @{text "$q_{t&x}"}\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct, simp_all add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma p_q_1: "$p_{t&x} + $q_{t&x} = 1" for t::real unfolding survive_def die_def using distrTx_RD.add_cdf_ccdf by (smt (verit) distrTx_RD.prob_space x_lt_psi) lemma q_defer_p: "$q_{f\t&x} = $p_{f&x} - $p_{f+t & x}" if "t \ 0" for f t :: real using q_defer_q p_q_1 that x_lt_psi by smt lemma q_defer_p_q_defer: "$p_{f&x} * $q_{f'\t & x+f} = $q_{f+f'\t & x}" if "x+f < $\" "f \ 0" "f' \ 0" "t \ 0" for f f' t :: real proof - have "$p_{f&x} * $q_{f'\t & x+f} = \

(\ in \. X \ > x+f) / \

(\ in \. X \ > x) * \

(\ in \. x+f+f' < X \ \ X \ \ x+f+f'+t) / \

(\ in \. X \ > x+f)" apply (rewrite p_PX, (simp_all add: that)[2]) by (rewrite survival_model.q_defer_PX, simp_all add: that survival_model_axioms) also have "\ = \

(\ in \. x+f+f' < X \ \ X \ \ x+f+f'+t) / \

(\ in \. X \ > x)" using survival_model.PXx_pos[of \ X "x+f"] nonzero_mult_div_cancel_left that by (smt (verit, ccfv_SIG) survival_model_axioms times_divide_eq_left times_divide_eq_right) also have "\ = $q_{f+f'\t & x}" by (rewrite q_defer_PX; simp add: that group_cancel.add1) finally show ?thesis . qed lemma q_defer_pq: "$q_{f\t&x} = $p_{f&x} * $q_{t & x+f}" if "x+f < $\" "t \ 0" "f \ 0" for f t :: real using q_defer_p_q_defer[where f'=0] that by (simp add: survival_model.q_defer_0_q survival_model_axioms) subsubsection \Properties of Life Expectation\ lemma e_nonneg: "$e`\_x \ 0" unfolding life_expect_def by (rule Bochner_Integration.integral_nonneg, simp add: less_eq_real_def) lemma e_P: "$e`\_x = MM_PS.expectation (\\. indicator (alive x) \ * T x \) / \

(\ in \. T x \ > 0)" unfolding life_expect_def by (rewrite MM_PS.integral_cond_prob_space_nn, auto simp add: alive_T) proposition nn_integral_T_p: - "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = \\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel" + "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = (\\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel)" apply (rewrite alivex_PS.expectation_nonneg_tail, simp_all add: less_imp_le) apply (rule nn_integral_cong) unfolding survive_def using distrTx_RD.prob_space distrTx_RD.ccdf_cdf by presburger lemma nn_integral_T_pos: "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) > 0" proof - let ?f = "\t. - ccdf (distr (\ \ alive x) borel (T x)) t" have "\t u. t \ u \ ?f t \ ?f u" using distrTx_RD.ccdf_nonincreasing by simp moreover have "continuous (at_right 0) ?f" using distrTx_RD.ccdf_is_right_cont by (intro continuous_intros) ultimately have "\e>0. \d>0. ?f (0 + d) - ?f 0 < e" using continuous_at_right_real_increasing by simp hence "\d>0. ?f (0 + d) - ?f 0 < 1/2" by (smt (verit, del_insts) field_sum_of_halves) from this obtain d where d_pos: "d > 0" and "$p_{d&x} \ 1/2" using p_0_1 unfolding survive_def by auto hence "\t. t\{0..d} \ $p_{t&x} \ 1/2" unfolding survive_def using distrTx_RD.ccdf_nonincreasing by force - hence "\\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel \ \\<^sup>+t\{0..d}. ennreal (1/2) \lborel" + hence "(\\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel) \ (\\<^sup>+t\{0..d}. ennreal (1/2) \lborel)" apply (intro nn_set_integral_mono, simp_all) unfolding survive_def using Tx_alivex_measurable apply force by (rule AE_I2) (smt (verit) ennreal_half ennreal_leI half_bounded_equal) - moreover have "\\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel \ \\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel" + moreover have "(\\<^sup>+t\{0..}. ennreal ($p_{t&x}) \lborel) \ (\\<^sup>+t\{0..d}. ennreal ($p_{t&x}) \lborel)" by (rule nn_set_integral_set_mono) simp - moreover have "\\<^sup>+t\{0..d}. ennreal (1/2) \lborel > 0" + moreover have "(\\<^sup>+t\{0..d}. ennreal (1/2) \lborel) > 0" apply (rewrite nn_integral_cmult_indicator, simp_all) using d_pos emeasure_lborel_Icc ennreal_zero_less_mult_iff by fastforce ultimately show ?thesis using nn_integral_T_p by simp qed -proposition e_LBINT_p: "$e`\_x = LBINT t:{0..}. $p_{t&x}" +proposition e_LBINT_p: "$e`\_x = (LBINT t:{0..}. $p_{t&x})" \ \Note that 0 = 0 holds when the integral diverges.\ unfolding life_expect_def apply (rewrite integral_eq_nn_integral, simp_all add: less_imp_le) unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) apply (measurable, simp add: survive_def) by (rewrite nn_integral_T_p) (simp add: indicator_mult_ennreal mult.commute) corollary e_integral_p: "$e`\_x = integral {0..} (\t. $p_{t&x})" \ \Note that 0 = 0 holds when the integral diverges.\ proof - - have "$e`\_x = LBINT t:{0..}. $p_{t&x}" using e_LBINT_p by simp + have "$e`\_x = (LBINT t:{0..}. $p_{t&x})" using e_LBINT_p by simp also have "\ = integral {0..} (\t. $p_{t&x})" apply (rule set_borel_integral_eq_integral_nonneg, simp_all) unfolding survive_def by simp finally show ?thesis . qed -lemma e_LBINT_p_Icc: "$e`\_x = LBINT t:{0..n}. $p_{t&x}" if "x+n \ $\" for n::real +lemma e_LBINT_p_Icc: "$e`\_x = (LBINT t:{0..n}. $p_{t&x})" if "x+n \ $\" for n::real proof - have [simp]: "{0..n} \ {n<..} = {}" using ivl_disj_int_one(7) by blast have [simp]: "{0..n} \ {n<..} = {0..}" by (smt (verit) ereal_less_le ivl_disj_un_one(7) leD that x_lt_psi) have [simp]: "\t. n < t \ 0 \ t" using that x_lt_psi by (smt (verit) ereal_less_le leD) have [simp]: "\t. n < t \ $\ \ ereal (x+t)" using that by (simp add: le_ereal_le) have gt_n_0: "has_bochner_integral lborel (\t. indicat_real {n<..} t * $p_{t&x}) 0" apply (rewrite has_bochner_integral_cong[where N=lborel and g="\t.0" and y=0], simp_all) using p_0_equiv that x_lt_psi apply (smt (verit, ccfv_SIG) greaterThan_iff indicator_simps le_ereal_le linorder_not_le) by (rule has_bochner_integral_zero) hence gt_n: "set_integrable lborel {n<..} (\t. $p_{t&x})" unfolding set_integrable_def using integrable.simps by auto moreover have le_n: "set_integrable lborel {0..n} (\t. $p_{t&x})" unfolding survive_def by (intro ccdfTx_integrable_Icc) simp ultimately have "set_integrable lborel ({0..n} \ {n<..}) (\t. $p_{t&x})" using set_integrable_Un by force hence "set_integrable lborel {0..} (\t. $p_{t&x})" by force thus ?thesis apply (rewrite e_LBINT_p, simp) apply (rewrite set_integral_Un[of "{0..n}" "{n<..}", simplified], simp_all add: gt_n le_n) unfolding set_lebesgue_integral_def using gt_n_0 has_bochner_integral_integral_eq by fastforce qed lemma e_integral_p_Icc: "$e`\_x = integral {0..n} (\t. $p_{t&x})" if "x+n \ $\" for n::real using that apply (rewrite e_LBINT_p_Icc, simp_all) using ccdfTx_integrable_Icc unfolding survive_def by (rewrite set_borel_integral_eq_integral; simp) lemma temp_e_P: "$e`\_{x:n} = MM_PS.expectation (\\. indicator (alive x) \ * min (T x \) n) / \

(\ in \. T x \ > 0)" if "n \ 0" for n::real unfolding temp_life_expect_def by (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T that) -lemma temp_e_LBINT_p: "$e`\_{x:n} = LBINT t:{0..n}. $p_{t&x}" if "n \ 0" for n::real +lemma temp_e_LBINT_p: "$e`\_{x:n} = (LBINT t:{0..n}. $p_{t&x})" if "n \ 0" for n::real proof - let ?minTxn = "\\. min (T x \) n" let ?F = "cdf (distr (\ \ alive x) borel (T x))" let ?Fn = "cdf (distr (\ \ alive x) borel ?minTxn)" interpret distrTxn_RD: real_distribution "distr (\ \ alive x) borel ?minTxn" by (simp add: that) have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) - have "(\\<^sup>+\. ennreal (min (T x \) n) \(\ \ alive x)) = \\<^sup>+t\{0..}. ennreal (1 - ?Fn t) \lborel" + have "(\\<^sup>+\. ennreal (min (T x \) n) \(\ \ alive x)) = (\\<^sup>+t\{0..}. ennreal (1 - ?Fn t) \lborel)" by (rewrite alivex_PS.expectation_nonneg_tail; simp add: that) - also have "\ = \\<^sup>+t\{0..}. (ennreal (1 - ?F t) * indicator {..lborel" + also have "\ = (\\<^sup>+t\{0..}. (ennreal (1 - ?F t) * indicator {..lborel)" apply (rule nn_integral_cong) by (rewrite alivex_PS.cdf_distr_min; simp add: indicator_mult_ennreal mult.commute) - also have "\ = \\<^sup>+t\{0..lborel" + also have "\ = (\\<^sup>+t\{0..lborel)" apply (rule nn_integral_cong) using nn_integral_set_ennreal by (smt (verit, best) Int_def atLeastLessThan_def ennreal_mult_right_cong indicator_simps mem_Collect_eq mult.commute mult_1) - also have "\ = \\<^sup>+t\{0..n}. ennreal (1 - ?F t) \lborel" + also have "\ = (\\<^sup>+t\{0..n}. ennreal (1 - ?F t) \lborel)" proof - have "sym_diff {0.. = ennreal (LBINT t:{0..n}. $p_{t&x})" proof - have "set_integrable lborel {0..n} (\t. $p_{t&x})" unfolding survive_def by (intro ccdfTx_integrable_Icc) simp thus ?thesis unfolding set_lebesgue_integral_def unfolding set_integrable_def apply (rewrite nn_integral_eq_integral[THEN sym]; simp) apply (rule nn_integral_cong, simp) unfolding survive_def using distrTx_RD.ccdf_cdf distrTx_RD.prob_space nn_integral_set_ennreal by (simp add: indicator_mult_ennreal mult.commute) qed finally have "(\\<^sup>+\. ennreal (min (T x \) n) \(\ \ alive x)) = ennreal (LBINT t:{0..n}. $p_{t&x})" . thus ?thesis unfolding temp_life_expect_def apply (rewrite integral_eq_nn_integral; simp add: that) apply (rewrite enn2real_ennreal; simp?) unfolding set_lebesgue_integral_def by (simp add: that) qed lemma temp_e_integral_p: "$e`\_{x:n} = integral {0..n} (\t. $p_{t&x})" if "n \ 0" for n::real using that apply (rewrite temp_e_LBINT_p, simp_all) using ccdfTx_integrable_Icc unfolding survive_def by (rewrite set_borel_integral_eq_integral; simp) lemma e_eq_temp: "$e`\_x = $e`\_{x:n}" if "n \ 0" "x+n \ $\" for n::real using that e_LBINT_p_Icc temp_e_LBINT_p by simp lemma curt_e_P: "$e_x = MM_PS.expectation (\\. indicator (alive x) \ * \T x \\) / \

(\ in \. T x \ > 0)" unfolding curt_life_expect_def apply (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T) by (metis (no_types, lifting) Bochner_Integration.integral_cong indicator_simps of_int_0 of_int_1) lemma curt_e_sum_P: "$e_x = (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "summable (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" proof - let ?F_flrTx = "cdf (distr (\ \ alive x) borel (\\. \T x \\))" have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) - have "integral\<^sup>N (\ \ alive x) (\\. ennreal \T x \\) = \\<^sup>+t\{0..}. ennreal (1 - ?F_flrTx t) \lborel" + have "integral\<^sup>N (\ \ alive x) (\\. ennreal \T x \\) = (\\<^sup>+t\{0..}. ennreal (1 - ?F_flrTx t) \lborel)" by (rewrite alivex_PS.expectation_nonneg_tail; simp) - also have "\ = \\<^sup>+t\{0::real..}. ennreal \

(\ in \. T x \ \ \t\ + 1 \ T x \ > 0) \lborel" + also have "\ = (\\<^sup>+t\{0::real..}. ennreal \

(\ in \. T x \ \ \t\ + 1 \ T x \ > 0) \lborel)" proof - { fix t::real assume "t \ 0" hence "1 - ?F_flrTx t = \

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0)" proof - have "1 - ?F_flrTx t = 1 - \

(\ in (\ \ alive x). T x \ < real_of_int \t\ + 1)" by (rewrite alivex_PS.cdf_distr_floor_P; simp) also have "\ = 1 - \

(\ in \. T x \ < real_of_int \t\ + 1 \ T x \ > 0)" using alive_T by (rewrite MM_PS.cond_prob_space_cond_prob; simp) also have "\ = \

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0)" by (rewrite not_le[THEN sym], rewrite MM_PS.cond_prob_neg; simp) finally show ?thesis . qed } thus ?thesis - apply - - by (rule nn_set_integral_cong2, rule AE_I2) simp + by (intro nn_set_integral_cong2 AE_I2) simp qed also have "\ = (\k. \\<^sup>+t\{k..(\ in \. T x \ \ \t\ + 1 \ T x \ > 0) \lborel)" apply (rewrite nn_integral_disjoint_family[THEN sym]; simp) apply (rewrite add.commute, rule Ico_nat_disjoint) by (rewrite Ico_nat_union[THEN sym], simp add: add.commute) also have "\ = (\k. \\<^sup>+t\{k..(\ in \. T x \ \ k + 1 \ T x \ > 0) \lborel)" proof - { fix k::nat and t::real assume "real k \ t" and "t < 1 + real k" hence "real_of_int \t\ = real k" by (metis add.commute floor_eq2 of_int_of_nat_eq) hence "\

(\ in \. T x \ \ real_of_int \t\ + 1 \ T x \ > 0) = \

(\ in \. T x \ \ 1 + real k \ T x \ > 0)" by (simp add: add.commute) } thus ?thesis apply - apply (rule suminf_cong, rule nn_set_integral_cong2, rule AE_I2) by (rule impI) simp qed also have "\ = (\k. ennreal \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" by (rewrite nn_integral_cmult_indicator; simp add: add.commute) also have "\ = ennreal (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" by (rewrite suminf_ennreal2; simp add: that) finally have "integral\<^sup>N (\ \ alive x) (\\. ennreal \T x \\) = ennreal (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" . hence "integral\<^sup>L (\ \ alive x) (\\. \T x \\) = (\k. \

(\ in \. T x \ \ k + 1 \ T x \ > 0))" apply (rewrite integral_eq_nn_integral; simp) apply (rewrite enn2real_ennreal; simp add: add.commute) apply (rule suminf_nonneg; simp?) by (rewrite add.commute, simp add: that) thus ?thesis unfolding curt_life_expect_def by (simp add: add.commute) qed lemma curt_e_sum_P_finite: "$e_x = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "x+n+1 > $\" for n::nat proof - from that have psi_fin: "$\ < \" by force let ?P = "\k::nat. \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" let ?P_fin = "\k::nat. if k\{..k. ?P k = ?P_fin k" proof - fix k show "?P k = ?P_fin k" proof (cases \k \ {..) case True thus ?thesis by simp next case False hence "\ k < n" by simp hence "x + k + 1 > real_of_ereal $\" using that psi_nonneg real_of_ereal_ord_simps(4) by fastforce hence "{\ \ space \. T x \ \ k + 1 \ T x \ > 0} \ {\ \ space \. X \ > real_of_ereal $\}" unfolding futr_life_def using that less_ereal_le of_nat_1 of_nat_add by force hence "\

(\ in \. T x \ \ k + 1 \ T x \ > 0) \ \

(\ in \. X \ > real_of_ereal $\)" by (intro MM_PS.finite_measure_mono, simp_all) also have "\ = 0" using MM_PS.ccdf_distr_P X_RV ccdfX_psi_0 psi_fin by presburger finally have "\

(\ in \. T x \ \ k + 1 \ T x \ > 0) = 0" using measure_le_0_iff by blast hence "?P k = 0" unfolding cond_prob_def by (simp add: add.commute) thus ?thesis by simp qed qed moreover have "?P_fin sums (\k: "?P sums (\kk) thus ?thesis by (simp add: add.commute) qed lemma curt_e_sum_p: "$e_x = (\k. $p_{k+1&x})" if "summable (\k. $p_{k+1&x})" "\k::nat. isCont (\t. $p_{t&x}) (k+1)" proof - have "\k::nat. $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that(2) isCont_ccdfX_ccdfTx unfolding survive_def by simp thus ?thesis using that p_PTx_ge_ccdf_isCont curt_e_sum_P by presburger qed lemma curt_e_rec: "$e_x = $p_x * (1 + $e_(x+1))" if "summable (\k. $p_{k+1&x})" "\k::nat. isCont (\t. $p_{t&x}) (real k + 1)" "x+1 < $\" proof - have px_neq_0[simp]: "$p_x \ 0" using p_0_equiv that by auto have "(\k. $p_{k+1&x}) sums $e_x" using that apply (rewrite curt_e_sum_p, simp_all add: add.commute) by (rule summable_sums, simp add: that) hence "(\k. $p_x * $p_{k&x+1}) sums $e_x" apply (rewrite sums_cong[where g="\k. $p_{k+1&x}"]; simp?) using p_mult by (smt (verit) of_nat_0_le_iff that(3) x_lt_psi) hence "(\k. $p_{k&x+1}) sums ($e_x / $p_x)" using sums_mult_D that by (smt (verit, best) linorder_not_le p_0_equiv sums_cong x_lt_psi) hence p_e_p: "(\k. $p_{Suc k & x+1}) sums ($e_x / $p_x - $p_{0&x+1})" using sums_split_initial_segment[where n=1] by force moreover have "(\k. $p_{Suc k & x+1}) sums $e_(x+1)" proof - have [simp]: "summable (\k::nat. $p_{real k + 1 & x + 1})" apply (intro sums_summable[where l="$e_x / $p_x - $p_{0&x+1}"]) using p_e_p by (simp add: add.commute) have [simp]: "\k::nat. isCont (\t. $p_{t&x+1}) (real k + 1)" proof - fix k::nat have "isCont (\t. $p_x * $p_{t-1&x+1}) (real k + 2)" proof - let ?S="{real k + 1 <..< real k + 3}" have "open ?S" by simp moreover have "real k + 2 \ ?S" by simp moreover have "\t. t \ ?S \ $p_x * $p_{t-1&x+1} = $p_{t&x}" using p_mult by (smt (verit, del_insts) greaterThanLessThan_iff of_nat_0_le_iff that(3) x_lt_psi) ultimately show ?thesis apply (rewrite isCont_cong[where g="\t. $p_{t&x}"]) apply (rewrite eventually_nhds, blast) using that by (smt (verit) of_nat_1 of_nat_add) qed hence "isCont (\t. $p_x * $p_{t-1&x+1} / $p_x) (real k + 2)" by (intro isCont_divide[where g="\t. $p_x"], auto) hence "isCont ((\t. $p_{t-1&x+1}) \ (\t. t+1)) (real k + 1)" by simp (rule continuous_at_compose, simp_all add: add.commute) thus "isCont (\t. $p_{t&x+1}) (real k + 1)" unfolding comp_def by simp qed show ?thesis apply (rewrite survival_model.curt_e_sum_p; simp add: survival_model_axioms that) using summable_sums by (rewrite add.commute) force qed ultimately have "$e_x / $p_x - $p_{0&x+1} = $e_(x+1)" by (rule sums_unique2) thus ?thesis using p_0_1 that by (smt (verit) px_neq_0 divide_mult_cancel mult.commute mult_cancel_left2 p_mult that(3)) qed lemma curt_e_sum_p_finite: "$e_x = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+n+1 > $\" for n::nat proof - have "\k::nat. k < n \ $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that isCont_ccdfX_ccdfTx unfolding survive_def by (smt (verit) of_nat_0_le_iff x_lt_psi) thus ?thesis using that p_PTx_ge_ccdf_isCont curt_e_sum_P_finite by simp qed lemma temp_curt_e_P: "$e_{x:n} = MM_PS.expectation (\\. indicator (alive x) \ * \min (T x \) n\) / \

(\ in \. T x \ > 0)" if "n \ 0" for n::real unfolding temp_curt_life_expect_def apply (rewrite MM_PS.integral_cond_prob_space_nn; simp add: alive_T that) apply (rule disjI2, rule Bochner_Integration.integral_cong; simp) using indicator_simps of_int_0 of_int_1 by smt lemma temp_curt_e_sum_P: "$e_{x:n} = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" for n::nat proof - let ?F_flrminTx = "cdf (distr (\ \ alive x) borel (\\. \min (T x \) n\))" have [simp]: "\\. \ \ space (\ \ alive x) \ 0 \ T x \" by (smt (verit) alivex_Tx_pos) have "integral\<^sup>N (\ \ alive x) (\\. ennreal \min (T x \) n\) = (\\<^sup>+t\{0..}. ennreal (1 - ?F_flrminTx t) \lborel)" by (rewrite alivex_PS.expectation_nonneg_tail; simp) - also have "\ = \\<^sup>+t\{0::real..}. ennreal - ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n)) \lborel" + also have "\ = (\\<^sup>+t\{0::real..}. ennreal + ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n)) \lborel)" proof - have "\t. t \ 0 \ ennreal (1 - ?F_flrminTx t) = ennreal ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n))" proof - fix t::real assume "t \ 0" thus "ennreal (1 - ?F_flrminTx t) = ennreal ((1 - \

(\ in (\ \ alive x). T x \ < \t\ + 1)) * of_bool (\t\ + 1 \ n))" proof (cases \\t\ + 1 \ n\) case True thus ?thesis apply (rewrite alivex_PS.cdf_distr_floor_P; simp) using min_less_iff_disj by (smt (verit, ccfv_SIG) Collect_cong floor_eq floor_less_cancel floor_of_nat of_int_floor_le) next case False thus ?thesis apply (rewrite alivex_PS.cdf_distr_floor_P; simp) using min_less_iff_disj by (smt (verit, ccfv_SIG) Collect_cong MM_PS.space_cond_prob_space alive_T alive_event alivex_PS.prob_space mem_Collect_eq of_int_of_nat_eq of_nat_less_of_int_iff) qed qed thus ?thesis by (intro nn_set_integral_cong2, intro AE_I2) auto qed - also have "\ = \\<^sup>+t\{0..(\ in (\ \ alive x). T x \ < \t\ + 1)) \lborel" + also have "\ = (\\<^sup>+t\{0..(\ in (\ \ alive x). T x \ < \t\ + 1)) \lborel)" proof - { fix t::real have "(\t\ + 1 \ n) = (t < n)" by linarith hence "\r::real. ennreal (r * of_bool (\t\ + 1 \ n)) * indicator {0..} t = ennreal r * indicator {0.. = \\<^sup>+t\{0..(\ in (\ \ alive x). T x \ \ \t\ + 1) \lborel" + also have "\ = (\\<^sup>+t\{0..(\ in (\ \ alive x). T x \ \ \t\ + 1) \lborel)" by (rewrite alivex_PS.prob_neg[THEN sym]; simp add: not_less) also have "\ = (\k\<^sup>+t\{k..(\ in (\ \ alive x). T x \ \ \t\ + 1) \lborel)" apply (rewrite Ico_nat_union_finite[of n, THEN sym]) apply (rewrite nn_integral_disjoint_family_on_finite; simp add: add.commute) apply (rule disjoint_family_on_mono[of _ "{0..}"]; simp) by (rewrite add.commute, rule Ico_nat_disjoint) also have "\ = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" proof - { fix k::nat assume "k < n" - hence "\\<^sup>+t\{k..<(1 + real k)}. - ennreal \

(\ in (\ \ alive x). T x \ \ real_of_int \t\ + 1) \lborel = + hence "(\\<^sup>+t\{k..<(1 + real k)}. + ennreal \

(\ in (\ \ alive x). T x \ \ real_of_int \t\ + 1) \lborel) = \

(\ in \. T x \ \ 1 + real k \ T x \ > 0)" (is "?LHS = ?RHS") proof - - have "?LHS = \\<^sup>+t\{k..<(1 + real k)}. ennreal \

(\ in (\ \ alive x). T x \ \ k + 1) \lborel" + have "?LHS = (\\<^sup>+t\{k..<(1 + real k)}. ennreal \

(\ in (\ \ alive x). T x \ \ k + 1) \lborel)" proof - { fix t::real assume "k \ t" "t < 1 + real k" hence "real_of_int \t\ = real k" by (smt (verit) floor_eq2 of_int_of_nat_eq) hence "\

(\ in (\ \ alive x). T x \ \ real_of_int \t\ + 1) = \

(\ in (\ \ alive x). T x \ \ 1 + real k)" by (simp add: add.commute) } thus ?thesis by (intro nn_set_integral_cong2, intro AE_I2) auto qed also have "\ = ennreal \

(\ in (\ \ alive x). T x \ \ k + 1)" by (rewrite nn_integral_cmult_indicator; simp) also have "\ = ?RHS" by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: alive_T) finally show ?thesis . qed } thus ?thesis by (intro sum.cong) auto qed also have "\ = ennreal (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" by simp finally have "integral\<^sup>N (\ \ alive x) (\\. ennreal \min (T x \) n\) = ennreal (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" . hence "integral\<^sup>L (\ \ alive x) (\\. \min (T x \) n\) = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" apply (intro nn_integral_eq_integrable[THEN iffD1, THEN conjunct2]; simp) using MM_PS.cond_prob_nonneg by (meson sum_nonneg) thus ?thesis unfolding temp_curt_life_expect_def by simp qed corollary curt_e_eq_temp: "$e_x = $e_{x:n}" if "x+n+1 > $\" for n::nat using curt_e_sum_P_finite temp_curt_e_sum_P that by simp lemma temp_curt_e_sum_p: "$e_{x:n} = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" for n::nat proof - have "\k::nat. k < n \ $p_{k+1&x} = \

(\ in \. T x \ \ k + 1 \ T x \ > 0)" apply (rewrite p_PTx_ge_ccdf_isCont, simp_all) using that isCont_ccdfX_ccdfTx unfolding survive_def by (smt (verit) of_nat_0_le_iff x_lt_psi) thus ?thesis apply (rewrite temp_curt_e_sum_P) by (rule sum.cong) simp_all qed lemma temp_curt_e_rec: "$e_{x:n} = $p_x * (1 + $e_{x+1:n-1})" if "\k::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+1 < $\" "n \ 0" for n::nat proof - have [simp]: "$p_x \ 0" using p_0_equiv that by auto have [simp]: "\k. k < n-1 \ isCont (\t. $p_{t&x+1}) (real k + 1)" proof - fix k::nat assume k_n: "k < n-1" have "isCont (\t. $p_x * $p_{t-1&x+1}) (real k + 2)" proof - let ?S="{real k + 1 <..< real k + 3}" have "open ?S" by simp moreover have "real k + 2 \ ?S" by simp moreover have "\t. t \ ?S \ $p_x * $p_{t-1&x+1} = $p_{t&x}" using p_mult by (smt (verit, del_insts) greaterThanLessThan_iff of_nat_0_le_iff that(2) x_lt_psi) ultimately show ?thesis apply (rewrite isCont_cong[where g="\t. $p_{t&x}"]) apply (rewrite eventually_nhds, blast) using that k_n by (smt (verit) less_diff_conv of_nat_1 of_nat_add) qed hence "isCont (\t. $p_x * $p_{t-1&x+1} / $p_x) (real k + 2)" by (intro isCont_divide[where g="\t. $p_x"], auto) hence "isCont ((\t. $p_{t-1&x+1}) \ (\t. t+1)) (real k + 1)" by simp (rule continuous_at_compose, simp_all add: add.commute) thus "isCont (\t. $p_{t&x+1}) (real k + 1)" unfolding comp_def by simp qed have "$p_x * (1 + $e_{x+1:n-1}) = $p_x * (1 + (\k<(n-1). $p_{k+1&x+1}))" by (rewrite survival_model.temp_curt_e_sum_p; simp add: survival_model_axioms that) also have "\ = $p_x + (\k<(n-1). $p_x * $p_{k+1&x+1})" apply (rewrite distrib_left, simp) by (rewrite vector_space_over_itself.scale_sum_right, simp) also have "\ = $p_x + (\k<(n-1). $p_{k+2&x})" apply (rewrite sum.cong, simp_all add: add.commute) using p_mult by (smt (verit) of_nat_0_le_iff that(2) x_lt_psi) also have "\ = (\k = $e_{x:n}" using that by (rewrite temp_curt_e_sum_p; simp) finally show ?thesis by simp qed end end subsection \Piecewise Differentiable Survival Function\ locale smooth_survival_function = survival_model + assumes ccdfX_piecewise_differentiable[simp]: "(ccdf (distr \ borel X)) piecewise_differentiable_on UNIV" begin interpretation distrX_RD: real_distribution "distr \ borel X" using MM_PS.real_distribution_distr by simp subsubsection \Properties of Survival Function for X\ lemma ccdfX_continuous[simp]: "continuous_on UNIV (ccdf (distr \ borel X))" using ccdfX_piecewise_differentiable piecewise_differentiable_on_imp_continuous_on by fastforce corollary ccdfX_borel_measurable[measurable]: "ccdf (distr \ borel X) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma ccdfX_nondifferentiable_finite_set[simp]: "finite {x. \ ccdf (distr \ borel X) differentiable at x}" proof - obtain S where fin: "finite S" and "\x. x \ S \ (ccdf (distr \ borel X)) differentiable at x" using ccdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast hence "{x. \ ccdf (distr \ borel X) differentiable at x} \ S" by blast thus ?thesis using finite_subset fin by blast qed lemma ccdfX_differentiable_open_set: "open {x. ccdf (distr \ borel X) differentiable at x}" using ccdfX_nondifferentiable_finite_set finite_imp_closed by (metis (mono_tags, lifting) Collect_cong open_Collect_neg) lemma ccdfX_differentiable_borel_set[measurable, simp]: "{x. ccdf (distr \ borel X) differentiable at x} \ sets borel" using ccdfX_differentiable_open_set by simp lemma ccdfX_differentiable_AE: "AE x in lborel. (ccdf (distr \ borel X)) differentiable at x" apply (rule AE_I'[of "{x. \ ccdf (distr \ borel X) differentiable at x}"], simp_all) using ccdfX_nondifferentiable_finite_set by (simp add: finite_imp_null_set_lborel) lemma deriv_ccdfX_measurable[measurable]: "deriv (ccdf (distr \ borel X)) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (ccdf (distr \ borel X)))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Cumulative Distributive Function for X\ lemma cdfX_piecewise_differentiable[simp]: "(cdf (distr \ borel X)) piecewise_differentiable_on UNIV" by (rewrite distrX_RD.cdf_ccdf) (rule piecewise_differentiable_diff; simp) lemma cdfX_continuous[simp]: "continuous_on UNIV (cdf (distr \ borel X))" using cdfX_piecewise_differentiable piecewise_differentiable_on_imp_continuous_on by fastforce corollary cdfX_borel_measurable[measurable]: "cdf (distr \ borel X) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma cdfX_nondifferentiable_finite_set[simp]: "finite {x. \ cdf (distr \ borel X) differentiable at x}" using distrX_RD.differentiable_cdf_ccdf ccdfX_nondifferentiable_finite_set by simp lemma cdfX_differentiable_open_set: "open {x. cdf (distr \ borel X) differentiable at x}" using distrX_RD.differentiable_cdf_ccdf ccdfX_differentiable_open_set by simp lemma cdfX_differentiable_borel_set[measurable, simp]: "{x. cdf (distr \ borel X) differentiable at x} \ sets borel" using cdfX_differentiable_open_set by simp lemma cdfX_differentiable_AE: "AE x in lborel. (cdf (distr \ borel X)) differentiable at x" using ccdfX_differentiable_AE distrX_RD.differentiable_cdf_ccdf AE_iffI by simp lemma deriv_cdfX_measurable[measurable]: "deriv (cdf (distr \ borel X)) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (cdf (distr \ borel X)))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Introduction of Probability Density Functions of X and T(x)\ definition pdfX :: "real \ real" where "pdfX x \ if cdf (distr \ borel X) differentiable at x then deriv (cdf (distr \ borel X)) x else 0" \ \This function is defined to be always nonnegative for future application.\ definition pdfT :: "real \ real \ real" where "pdfT x t \ if cdf (distr (\ \ alive x) borel (T x)) differentiable at t then deriv (cdf (distr (\ \ alive x) borel (T x))) t else 0" \ \This function is defined to be always nonnegative for future application.\ lemma pdfX_measurable[measurable]: "pdfX \ borel_measurable borel" proof - let ?cdfX = "cdf (distr \ borel X)" have "countable {x. \ ?cdfX differentiable at x}" using cdfX_nondifferentiable_finite_set uncountable_infinite by force thus ?thesis unfolding pdfX_def apply - by (rule measurable_discrete_difference [where X="{x. \ ?cdfX differentiable at x}" and f="deriv ?cdfX"]; simp) qed lemma distributed_pdfX: "distributed \ lborel X pdfX" proof - let ?cdfX = "cdf (distr \ borel X)" obtain S where fin: "finite S" and diff: "\t. t \ S \ ?cdfX differentiable at t" using cdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast { fix t::real assume t_notin: "t \ S" have "?cdfX differentiable at t" using diff t_notin by simp hence "(?cdfX has_real_derivative pdfX t) (at t)" unfolding pdfX_def using DERIV_deriv_iff_real_differentiable by auto } thus ?thesis by (intro MM_PS.distributed_deriv_cdf[where S=S]; simp add: fin) qed lemma pdfT0_X: "pdfT 0 = pdfX" unfolding pdfT_def pdfX_def using cdfT0_eq_cdfX psi_pos' by fastforce subsubsection \Properties of Survival Function for T(x)\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma ccdfTx_continuous_on_nonneg[simp]: "continuous_on {0..} (ccdf (distr (\ \ alive x) borel (T x)))" apply (rewrite continuous_on_eq_continuous_within, auto) apply (rewrite continuous_ccdfX_ccdfTx[THEN sym], auto) by (metis UNIV_I ccdfX_continuous continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma ccdfTx_continuous[simp]: "continuous_on UNIV (ccdf (distr (\ \ alive x) borel (T x)))" proof - have [simp]: "{..0::real} \ {0..} = UNIV" by auto have "continuous_on {..0} (ccdf (distr (\ \ alive x) borel (T x)))" by (rule ccdfTx_continuous_on_nonpos) simp moreover have "continuous_on {0..} (ccdf (distr (\ \ alive x) borel (T x)))" by simp ultimately have "continuous_on ({..0} \ {0..}) (ccdf (distr (\ \ alive x) borel (T x)))" by (intro continuous_on_closed_Un) simp_all thus ?thesis by simp qed corollary ccdfTx_borel_measurable[measurable]: "ccdf (distr (\ \ alive x) borel (T x)) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma ccdfTx_nondifferentiable_finite_set[simp]: "finite {t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}" proof - let ?P = "\t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t" have "{t. t < 0 \ \ ?P t} = {}" proof (rule equals0I) fix t assume asm: "t \ {t. t < 0 \ \ ?P t}" hence "?P t" using ccdfTx_has_real_derivative_0_at_neg real_differentiable_def x_lt_psi by blast with asm show False by simp qed hence "{t. \ ?P t} \ insert 0 {t. t > 0 \ \ ?P t}" by force moreover have "finite {t. t > 0 \ \ ?P t}" proof - have "{t. \ ccdf (distr \ borel X) differentiable at (x+t)} = plus (-x) ` {s. \ ccdf (distr \ borel X) differentiable at s}" by force hence "finite {t. \ ccdf (distr \ borel X) differentiable at (x+t)}" using ccdfX_nondifferentiable_finite_set by simp thus ?thesis using finite_subset differentiable_ccdfX_ccdfTx by (metis (no_types, lifting) mem_Collect_eq subsetI x_lt_psi) qed ultimately show ?thesis using finite_subset by auto qed lemma ccdfTx_differentiable_open_set: "open {t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using ccdfTx_nondifferentiable_finite_set finite_imp_closed by (metis (mono_tags, lifting) Collect_cong open_Collect_neg) lemma ccdfTx_differentiable_borel_set[measurable, simp]: "{t. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t} \ sets borel" using ccdfTx_differentiable_open_set by simp lemma ccdfTx_differentiable_AE: "AE t in lborel. (ccdf (distr (\ \ alive x) borel (T x))) differentiable at t" apply (rule AE_I'[of "{t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}"]; simp?) using ccdfTx_nondifferentiable_finite_set by (simp add: finite_imp_null_set_lborel) lemma ccdfTx_piecewise_differentiable[simp]: "(ccdf (distr (\ \ alive x) borel (T x))) piecewise_differentiable_on UNIV" proof - have "\t\UNIV-{t. \ ccdf (distr (\ \ alive x) borel (T x)) differentiable at t}. ccdf (distr (\ \ alive x) borel (T x)) differentiable at t" by auto thus ?thesis unfolding piecewise_differentiable_on_def using ccdfTx_continuous ccdfTx_nondifferentiable_finite_set by blast qed lemma deriv_ccdfTx_measurable[measurable]: "deriv (ccdf (distr (\ \ alive x) borel (T x))) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (ccdf (distr (\ \ alive x) borel (T x))))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Cumulative Distributive Function for T(x)\ lemma cdfTx_continuous[simp]: "continuous_on UNIV (cdf (distr (\ \ alive x) borel (T x)))" using distrTx_RD.cdf_ccdf ccdfTx_continuous by (simp add: continuous_on_eq_continuous_within) corollary cdfTx_borel_measurable[measurable]: "cdf (distr (\ \ alive x) borel (T x)) \ borel_measurable borel" by (rule borel_measurable_continuous_onI) simp lemma cdfTx_nondifferentiable_finite_set[simp]: "finite {t. \ cdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using distrTx_RD.differentiable_cdf_ccdf by simp lemma cdfTx_differentiable_open_set: "open {t. cdf (distr (\ \ alive x) borel (T x)) differentiable at t}" using distrTx_RD.differentiable_cdf_ccdf ccdfTx_differentiable_open_set by simp lemma cdfTx_differentiable_borel_set[measurable, simp]: "{t. cdf (distr (\ \ alive x) borel (T x)) differentiable at t} \ sets borel" using cdfTx_differentiable_open_set by simp lemma cdfTx_differentiable_AE: "AE t in lborel. (cdf (distr (\ \ alive x) borel (T x))) differentiable at t" by (rewrite distrTx_RD.differentiable_cdf_ccdf, simp add: ccdfTx_differentiable_AE) lemma cdfTx_piecewise_differentiable[simp]: "(cdf (distr (\ \ alive x) borel (T x))) piecewise_differentiable_on UNIV" using piecewise_differentiable_diff piecewise_differentiable_const ccdfTx_piecewise_differentiable by (rewrite distrTx_RD.cdf_ccdf) blast lemma deriv_cdfTx_measurable[measurable]: "deriv (cdf (distr (\ \ alive x) borel (T x))) \ borel_measurable borel" proof - have "set_borel_measurable borel UNIV (deriv (cdf (distr (\ \ alive x) borel (T x))))" by (rule piecewise_differentiable_on_deriv_measurable_real; simp) thus ?thesis unfolding set_borel_measurable_def by simp qed subsubsection \Properties of Probability Density Function of T(x)\ lemma pdfTx_nonneg: "pdfT x t \ 0" for t::real proof - fix t show "pdfT x t \ 0" proof (cases \cdf (distr (\ \ alive x) borel (T x)) differentiable at t\) case True have "mono_on UNIV (cdf (distr (\ \ alive x) borel (T x)))" unfolding mono_on_def using distrTx_RD.cdf_nondecreasing by blast moreover have "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative pdfT x t) (at t)" unfolding pdfT_def using True DERIV_deriv_iff_real_differentiable by presburger ultimately show ?thesis by (rule mono_on_imp_deriv_nonneg) simp next case False thus ?thesis unfolding pdfT_def by simp qed qed lemma pdfTx_neg_0: "pdfT x t = 0" if "t < 0" for t::real proof - let ?e = "-t/2" have "(cdf (distr (\ \ alive x) borel (T x)) has_real_derivative 0) (at t)" apply (rewrite DERIV_cong_ev[of t t _ "\_. 0" 0 0], simp_all add: that) apply (rewrite eventually_nhds) apply (rule exI[of _ "ball t ?e"], auto simp add: that) by (rule cdfTx_nonpos_0; simp add: dist_real_def) thus ?thesis unfolding pdfT_def by (meson DERIV_imp_deriv) qed lemma pdfTx_0_0: "pdfT x 0 = 0" proof (cases \cdf (distr (\ \ alive x) borel (T x)) differentiable at 0\) case True let ?cdfx = "cdf (distr (\ \ alive x) borel (T x))" have "(?cdfx has_real_derivative 0) (at 0)" proof - from True obtain c where cdfx_deriv: "(?cdfx has_real_derivative c) (at 0)" unfolding differentiable_def using has_real_derivative by blast hence "(?cdfx has_real_derivative c) (at 0 within {..0})" by (rule has_field_derivative_at_within) moreover have "(?cdfx has_real_derivative 0) (at 0 within {..0})" proof - have "\\<^sub>F t in at 0 within {..0}. ?cdfx t = 0" proof - { fix t assume "t \ {..0::real}" "t \ 0" "dist t 0 < 1" hence "?cdfx t = 0" using cdfTx_nonpos_0 x_lt_psi by blast } hence "\d>0::real. \t\{..0}. t\0 \ dist t 0 < d \ ?cdfx t = 0" by (smt (verit)) thus ?thesis by (rewrite eventually_at) simp qed moreover have "?cdfx 0 = 0" proof - have "continuous (at 0 within {..0}) ?cdfx" using True differentiable_imp_continuous_within differentiable_subset by blast thus ?thesis by (simp add: cdfTx_nonpos_0) qed ultimately show ?thesis by (rewrite has_field_derivative_cong_eventually[of _ "\_. 0::real" 0 "{..0}" 0]; simp) qed ultimately have "c = 0" using has_real_derivative_iff_has_vector_derivative apply (intro vector_derivative_unique_within[of 0 "{..0}" ?cdfx c 0]; blast?) by (rewrite at_within_eq_bot_iff) (metis closure_lessThan islimpt_in_closure limpt_of_closure trivial_limit_at_left_real trivial_limit_within) thus "(?cdfx has_real_derivative 0) (at 0)" using cdfx_deriv by simp qed thus ?thesis unfolding pdfT_def by (meson DERIV_imp_deriv) next case False thus ?thesis unfolding pdfT_def by simp qed lemma pdfTx_nonpos_0: "pdfT x t = 0" if "t \ 0" for t::real using pdfTx_neg_0 pdfTx_0_0 that by force lemma pdfTx_beyond_0: "pdfT x t = 0" if "x+t \ $\" for t::real proof (cases \t \ 0\) case True thus ?thesis using pdfTx_nonpos_0 by simp next let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" case False hence t_pos: "t > 0" by simp thus ?thesis proof - have "(?cdfTx has_real_derivative 0) (at_right t)" apply (rewrite has_field_derivative_cong_eventually[where g="\_. 1"], simp_all) apply (rewrite eventually_at_right_field) using that cdfTx_1_equiv by (intro exI[of _"t+1"], auto simp add: le_ereal_less less_eq_ereal_def) thus ?thesis unfolding pdfT_def by (meson has_real_derivative_iff_has_vector_derivative has_vector_derivative_at_within DERIV_deriv_iff_real_differentiable trivial_limit_at_right_real vector_derivative_unique_within) qed qed lemma pdfTx_pdfX: "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" if "t > 0" for t::real proof (cases \cdf (distr \ borel X) differentiable at (x+t)\) case True let ?cdfX = "cdf (distr \ borel X)" let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have [simp]: "?cdfTx differentiable at t" using differentiable_cdfX_cdfTx that True by simp have "pdfT x t = deriv ?cdfTx t" unfolding pdfT_def using that differentiable_cdfX_cdfTx by simp hence "(?cdfTx has_field_derivative (pdfT x t)) (at t)" using True DERIV_deriv_iff_real_differentiable by simp moreover have "\u. dist u t < t \ ?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = ?cdfTx u" proof - fix u::real assume "dist u t < t" hence [simp]: "u > 0" using dist_real_def by fastforce have "?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = (1 - \

(\ in \. X \ > x+u)) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1)" using MM_PS.ccdf_distr_P X_RV distrX_RD.cdf_ccdf distrX_RD.prob_space by presburger also have "\ = 1 - \

(\ in \. X \ > x+u) / \

(\ in \. X \ > x)" by (simp add: diff_divide_distrib) also have "\ = ?cdfTx u" apply (rewrite ccdfTx_PX[THEN sym], simp_all add: less_eq_real_def) using distrTx_RD.cdf_ccdf distrTx_RD.prob_space by presburger finally show "?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1) = ?cdfTx u" . qed ultimately have "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x) - (1 / \

(\ in \. X \ > x) - 1)) has_field_derivative (pdfT x t)) (at t)" apply - by (rule has_field_derivative_transform_within[where d=t]; simp add: that) hence "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x)) has_field_derivative (pdfT x t)) (at t)" unfolding has_field_derivative_def using has_derivative_add_const[where c="1 / \

(\ in \. X \ > x) - 1"] by force hence "((\u. ?cdfX (x+u) / \

(\ in \. X \ > x) * \

(\ in \. X \ > x)) has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at t)" using DERIV_cmult_right[where c="\

(\ in \. X \ > x)"] by force hence "((\u. ?cdfX (x+u)) has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at t)" unfolding has_field_derivative_def using has_derivative_transform PXx_pos x_lt_psi by (smt (verit) Collect_cong UNIV_I nonzero_mult_div_cancel_right times_divide_eq_left) hence "(?cdfX has_field_derivative pdfT x t * \

(\ in \. X \ > x)) (at (x+t))" using DERIV_at_within_shift by force moreover have "(?cdfX has_field_derivative deriv ?cdfX (x+t)) (at (x+t))" using True DERIV_deriv_iff_real_differentiable by blast ultimately have "pdfT x t * \

(\ in \. X \ > x) = deriv ?cdfX (x+t)" by (simp add: DERIV_imp_deriv) thus "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" unfolding pdfX_def using True by simp (metis PXx_pos nonzero_mult_div_cancel_right x_lt_psi zero_less_measure_iff) next case False hence [simp]: "\ cdf (distr (\ \ alive x) borel (T x)) differentiable at t" using differentiable_cdfX_cdfTx that by simp thus "pdfT x t = pdfX (x+t) / \

(\ in \. X \ > x)" unfolding pdfT_def pdfX_def using False by simp qed lemma pdfTx_measurable[measurable]: "pdfT x \ borel_measurable borel" proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have "countable {x. \ ?cdfTx differentiable at x}" using cdfX_nondifferentiable_finite_set uncountable_infinite by force thus ?thesis unfolding pdfT_def apply - by (rule measurable_discrete_difference [where X="{x. \ ?cdfTx differentiable at x}" and f="deriv ?cdfTx"]; simp) qed lemma distributed_pdfTx: "distributed (\ \ alive x) lborel (T x) (pdfT x)" proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" obtain S where fin: "finite S" and diff: "\t. t \ S \ ?cdfTx differentiable at t" using cdfTx_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast { fix t::real assume t_notin: "t \ S" have "?cdfTx differentiable at t" using diff t_notin by simp hence "(?cdfTx has_real_derivative pdfT x t) (at t)" unfolding pdfT_def using DERIV_deriv_iff_real_differentiable by force } thus ?thesis by (intro alivex_PS.distributed_deriv_cdf[where S=S]; simp add: fin) qed lemma nn_integral_pdfTx_1: "(\\<^sup>+s. pdfT x s \lborel) = 1" proof - have "(\\<^sup>+s. pdfT x s \lborel) = emeasure (density lborel (pdfT x)) UNIV" by (rewrite emeasure_density) simp_all also have "\ = emeasure (distr (\ \ alive x) lborel (T x)) UNIV" using distributed_pdfTx unfolding distributed_def by simp also have "\ = 1" by (metis distrTx_RD.emeasure_space_1 distrTx_RD.space_eq_univ distr_cong sets_lborel) finally show ?thesis . qed corollary has_bochner_integral_pdfTx_1: "has_bochner_integral lborel (pdfT x) 1" by (rule has_bochner_integral_nn_integral; simp add: pdfTx_nonneg nn_integral_pdfTx_1) -corollary LBINT_pdfTx_1: "LBINT s. pdfT x s = 1" +corollary LBINT_pdfTx_1: "(LBINT s. pdfT x s) = 1" using has_bochner_integral_pdfTx_1 by (simp add: has_bochner_integral_integral_eq) corollary pdfTx_has_integral_1: "(pdfT x has_integral 1) UNIV" by (rule nn_integral_has_integral; simp add: pdfTx_nonneg nn_integral_pdfTx_1) -lemma set_nn_integral_pdfTx_1: "\\<^sup>+s\{0..}. pdfT x s \lborel = 1" +lemma set_nn_integral_pdfTx_1: "(\\<^sup>+s\{0..}. pdfT x s \lborel) = 1" apply (rewrite nn_integral_pdfTx_1[THEN sym]) apply (rule nn_integral_cong) using pdfTx_nonpos_0 by (metis atLeast_iff ennreal_0 indicator_simps(1) linorder_le_cases mult.commute mult_1 mult_zero_left) corollary has_bochner_integral_pdfTx_1_nonpos: "has_bochner_integral lborel (\s. pdfT x s * indicator {0..} s) 1" apply (rule has_bochner_integral_nn_integral, simp_all) using pdfTx_nonneg apply simp using set_nn_integral_pdfTx_1 by (simp add: nn_integral_set_ennreal) corollary set_LBINT_pdfTx_1: "(LBINT s:{0..}. pdfT x s) = 1" unfolding set_lebesgue_integral_def using has_bochner_integral_pdfTx_1_nonpos has_bochner_integral_integral_eq apply (simp, rewrite mult.commute) by (smt (verit) Bochner_Integration.integral_cong has_bochner_integral_integral_eq) corollary pdfTx_has_integral_1_nonpos: "(pdfT x has_integral 1) {0..}" proof - have "set_integrable lebesgue {0..} (pdfT x)" unfolding set_integrable_def apply (rewrite integrable_completion, simp_all) using has_bochner_integral_pdfTx_1_nonpos by (rewrite mult.commute, rule integrable.intros) - moreover have "LINT s:{0..}|lebesgue. pdfT x s = 1" + moreover have "(LINT s:{0..}|lebesgue. pdfT x s) = 1" using set_LBINT_pdfTx_1 unfolding set_lebesgue_integral_def by (rewrite integral_completion; simp) ultimately show ?thesis using has_integral_set_lebesgue by force qed -lemma set_nn_integral_pdfTx_PTx: "\\<^sup>+s\A. pdfT x s \lborel = \

(\ in \. T x \ \ A \ T x \ > 0)" +lemma set_nn_integral_pdfTx_PTx: "(\\<^sup>+s\A. pdfT x s \lborel) = \

(\ in \. T x \ \ A \ T x \ > 0)" if "A \ sets lborel" for A :: "real set" proof - have [simp]: "A \ sets borel" using that by simp - have "\\<^sup>+s\A. pdfT x s \lborel = emeasure (density lborel (pdfT x)) A" + have "(\\<^sup>+s\A. pdfT x s \lborel) = emeasure (density lborel (pdfT x)) A" using that by (rewrite emeasure_density; force) also have "\ = emeasure (distr (\ \ alive x) lborel (T x)) A" using distributed_pdfTx unfolding distributed_def by simp also have "\ = ennreal \

(\ in (\ \ alive x). T x \ \ A)" apply (rewrite emeasure_distr, simp_all) apply (rewrite finite_measure.emeasure_eq_measure, simp) by (smt (verit) Collect_cong vimage_eq Int_def) also have "\ = ennreal \

(\ in \. T x \ \ A \ T x \ > 0)" unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], simp_all add: pred_def futr_life_def) using borel_measurable_diff X_RV that by measurable finally show ?thesis . qed lemma pdfTx_set_integrable: "set_integrable lborel A (pdfT x)" if "A \ sets lborel" unfolding set_integrable_def using that pdfTx_nonneg apply (intro integrableI_nonneg, simp_all) apply (rewrite mult.commute) using set_nn_integral_pdfTx_PTx that by (metis (no_types, lifting) ennreal_indicator ennreal_less_top ennreal_mult' nn_integral_cong) -lemma set_integral_pdfTx_PTx: "LBINT s:A. pdfT x s = \

(\ in \. T x \ \ A \ T x \ > 0)" +lemma set_integral_pdfTx_PTx: "(LBINT s:A. pdfT x s) = \

(\ in \. T x \ \ A \ T x \ > 0)" if "A \ sets lborel" for A :: "real set" unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral) using that apply (simp_all add: pdfTx_nonneg) apply (rewrite indicator_mult_ennreal[THEN sym], rewrite mult.commute) by (rewrite set_nn_integral_pdfTx_PTx; simp) lemma pdfTx_has_integral_PTx: "(pdfT x has_integral \

(\ in \. T x \ \ A \ T x \ > 0)) A" if "A \ sets lborel" for A :: "real set" proof - have "((\s. pdfT x s * indicat_real A s) has_integral \

(\ in \. T x \ \ A \ T x \ > 0)) UNIV" using that pdfTx_nonneg apply (intro nn_integral_has_integral, simp_all) using set_nn_integral_pdfTx_PTx that by (simp add: nn_integral_set_ennreal) thus ?thesis by (smt (verit) has_integral_cong has_integral_restrict_UNIV indicator_eq_0_iff indicator_simps(1) mult_cancel_left2 mult_eq_0_iff) qed corollary pdfTx_has_integral_PTx_Icc: "(pdfT x has_integral \

(\ in \. a \ T x \ \ T x \ \ b \ T x \ > 0)) {a..b}" for a b :: real using pdfTx_has_integral_PTx[of "{a..b}"] by simp corollary pdfTx_integrable_on_Icc: "pdfT x integrable_on {a..b}" for a b :: real using pdfTx_has_integral_PTx_Icc by auto end subsubsection \Properties of Probability Density Function of X\ lemma pdfX_nonneg: "pdfX x \ 0" for x::real using pdfTx_nonneg pdfT0_X psi_pos' by smt lemma pdfX_nonpos_0: "pdfX x = 0" if "x \ 0" for x::real using pdfTx_nonpos_0 pdfT0_X psi_pos' that by smt lemma pdfX_beyond_0: "pdfX x = 0" if "x \ $\" for x::real using pdfTx_beyond_0 pdfT0_X psi_pos' that by smt lemma nn_integral_pdfX_1: "integral\<^sup>N lborel pdfX = 1" using nn_integral_pdfTx_1 pdfT0_X psi_pos' by metis corollary has_bochner_integral_pdfX_1: "has_bochner_integral lborel pdfX 1" by (rule has_bochner_integral_nn_integral; simp add: pdfX_nonneg nn_integral_pdfX_1) -corollary LBINT_pdfX_1: "LBINT s. pdfX s = 1" +corollary LBINT_pdfX_1: "(LBINT s. pdfX s) = 1" using has_bochner_integral_pdfX_1 by (simp add: has_bochner_integral_integral_eq) corollary pdfX_has_integral_1: "(pdfX has_integral 1) UNIV" by (rule nn_integral_has_integral; simp add: pdfX_nonneg nn_integral_pdfX_1) lemma set_nn_integral_pdfX_PX: "set_nn_integral lborel A pdfX = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" using PT0_eq_PX_lborel that by (rewrite pdfT0_X[THEN sym], rewrite set_nn_integral_pdfTx_PTx; simp) lemma pdfX_set_integrable: "set_integrable lborel A pdfX" if "A \ sets lborel" for A :: "real set" using pdfTx_set_integrable pdfT0_X psi_pos' that by smt -lemma set_integral_pdfX_PX: "LBINT s:A. pdfX s = \

(\ in \. X \ \ A)" +lemma set_integral_pdfX_PX: "(LBINT s:A. pdfX s) = \

(\ in \. X \ \ A)" if "A \ sets lborel" for A :: "real set" using PT0_eq_PX_lborel that by (rewrite pdfT0_X[THEN sym], rewrite set_integral_pdfTx_PTx; simp) lemma pdfX_has_integral_PX: "(pdfX has_integral \

(\ in \. X \ \ A)) A" if "A \ sets lborel" for A :: "real set" using that apply (rewrite pdfT0_X[THEN sym], rewrite PT0_eq_PX_lborel[THEN sym], simp) by (rule pdfTx_has_integral_PTx; simp) corollary pdfX_has_integral_PX_Icc: "(pdfX has_integral \

(\ in \. a \ X \ \ X \ \ b)) {a..b}" for a b :: real using pdfX_has_integral_PX[of "{a..b}"] by simp corollary pdfX_integrable_on_Icc: "pdfX integrable_on {a..b}" for a b :: real using pdfX_has_integral_PX_Icc by auto subsubsection \Relations between Life Expectation and Probability Density Function\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp proposition nn_integral_T_pdfT: - "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = \\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel" + "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel)" if "g \ borel_measurable lborel" for g :: "real \ real" proof - - have "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = \\<^sup>+s. ennreal (pdfT x s) * ennreal (g s) \lborel" + have "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) = (\\<^sup>+s. ennreal (pdfT x s) * ennreal (g s) \lborel)" proof - have "distributed (\ \ alive x) lborel (T x) (\s. ennreal (pdfT x s))" by (intro distributed_pdfTx) simp moreover have "(\s. ennreal (g s)) \ borel_measurable borel" using that by measurable ultimately show ?thesis by (rewrite distributed_nn_integral; simp) qed - also have "\ = \\<^sup>+s. ennreal (pdfT x s * g s) \lborel" using ennreal_mult' pdfTx_nonneg by force - also have "\ = \\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel" + also have "\ = (\\<^sup>+s. ennreal (pdfT x s * g s) \lborel)" using ennreal_mult' pdfTx_nonneg by force + also have "\ = (\\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel)" apply (rule nn_integral_cong, simp) by (metis atLeast_iff ennreal_0 indicator_simps linorder_not_le mult_1 mult_commute_abs mult_zero_left pdfTx_neg_0 x_lt_psi) finally show ?thesis . qed lemma expectation_LBINT_pdfT_nonneg: - "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" + "alivex_PS.expectation (\\. g (T x \)) = (LBINT s:{0..}. pdfT x s * g s)" if "\s. s \ 0 \ g s \ 0" "g \ borel_measurable lborel" for g :: "real \ real" \ \Note that 0 = 0 holds when the integral diverges.\ using that apply (rewrite integral_eq_nn_integral, simp) apply (rule AE_I2, metis alivex_Tx_pos less_imp_le) unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) apply (rule AE_I2, metis indicator_pos_le pdfTx_nonpos_0 x_lt_psi zero_le_mult_iff zero_le_square pdfTx_nonneg) by (rewrite nn_integral_T_pdfT) (simp_all add: indicator_mult_ennreal mult.commute) corollary expectation_integral_pdfT_nonneg: "alivex_PS.expectation (\\. g (T x \)) = integral {0..} (\s. pdfT x s * g s)" if "\s. s \ 0 \ g s \ 0" "g \ borel_measurable lborel" for g :: "real \ real" \ \Note that 0 = 0 holds when the integral diverges.\ proof - - have "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" + have "alivex_PS.expectation (\\. g (T x \)) = (LBINT s:{0..}. pdfT x s * g s)" using expectation_LBINT_pdfT_nonneg that by simp also have "\ = integral {0..} (\s. pdfT x s * g s)" using that pdfTx_nonneg by (intro set_borel_integral_eq_integral_nonneg; simp) finally show ?thesis . qed proposition expectation_LBINT_pdfT: - "alivex_PS.expectation (\\. g (T x \)) = LBINT s:{0..}. pdfT x s * g s" + "alivex_PS.expectation (\\. g (T x \)) = (LBINT s:{0..}. pdfT x s * g s)" if "set_integrable lborel {0..} (\s. pdfT x s * g s)" "g \ borel_measurable lborel" for g :: "real \ real" proof - have pg_fin: "(\\<^sup>+\. ennreal (g (T x \)) \(\ \ alive x)) \ \" using that apply (rewrite nn_integral_T_pdfT, simp) using that unfolding set_integrable_def apply (rewrite in asm real_integrable_def, simp) by (simp add: indicator_mult_ennreal mult.commute) moreover have mg_fin: "(\\<^sup>+\. ennreal (- g (T x \)) \(\ \ alive x)) \ \" using that apply (rewrite nn_integral_T_pdfT[of "\s. - g s"], simp) using that unfolding set_integrable_def apply (rewrite in asm real_integrable_def, simp) by (simp add: indicator_mult_ennreal mult.commute) ultimately have [simp]: "integrable (\ \ alive x) (\\. g (T x \))" using that by (rewrite real_integrable_def; simp) have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (g s)) \lborel) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * g s) \lborel)" using SPMF.ennreal_max_0 ennreal_mult' pdfTx_nonneg x_lt_psi by presburger also have "\ < \" using pg_fin nn_integral_T_pdfT that top.not_eq_extremum by auto finally have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (g s)) \lborel) < \" . hence [simp]: "set_integrable lborel {0..} (\s. pdfT x s * max 0 (g s))" unfolding set_integrable_def using that apply (intro integrableI_nonneg, simp_all) using pdfTx_nonneg apply (intro AE_I2, force) by (metis (no_types, lifting) indicator_mult_ennreal mult.commute nn_integral_cong) have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (- g s)) \lborel) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * - g s) \lborel)" using SPMF.ennreal_max_0 ennreal_mult' pdfTx_nonneg x_lt_psi by presburger also have "\ < \" using mg_fin nn_integral_T_pdfT[of "\s. - g s"] that top.not_eq_extremum by force finally have "(\\<^sup>+s\{0..}. ennreal (pdfT x s * max 0 (- g s)) \lborel) < \" . hence [simp]: "set_integrable lborel {0..} (\s. pdfT x s * max 0 (- g s))" unfolding set_integrable_def using that apply (intro integrableI_nonneg, simp_all) using pdfTx_nonneg apply (intro AE_I2, force) by (metis (no_types, lifting) indicator_mult_ennreal mult.commute nn_integral_cong) have "alivex_PS.expectation (\\. g (T x \)) = alivex_PS.expectation (\\. max 0 (g (T x \))) - alivex_PS.expectation (\\. max 0 (- g (T x \)))" by (rewrite Bochner_Integration.integral_cong [where N="\ \ alive x" and g="\\. max 0 (g (T x \)) - max 0 (- g (T x \))"]; simp) moreover have "alivex_PS.expectation (\\. max 0 (g (T x \))) = (LBINT s:{0..}. pdfT x s * max 0 (g s))" using that by (rewrite expectation_LBINT_pdfT_nonneg[where g="\s. max 0 (g s)"]) simp_all moreover have "alivex_PS.expectation (\\. max 0 (- g (T x \))) = (LBINT s:{0..}. pdfT x s * max 0 (- g s))" using that by (rewrite expectation_LBINT_pdfT_nonneg[where g="\s. max 0 (- g s)"]) simp_all ultimately have "alivex_PS.expectation (\\. g (T x \)) = (LBINT s:{0..}. pdfT x s * max 0 (g s)) - (LBINT s:{0..}. pdfT x s * max 0 (- g s))" by simp - also have "\ = LBINT s:{0..}. (pdfT x s * max 0 (g s) - pdfT x s * max 0 (- g s))" + also have "\ = (LBINT s:{0..}. (pdfT x s * max 0 (g s) - pdfT x s * max 0 (- g s)))" by (rewrite set_integral_diff; simp) - also have "\ = LBINT s:{0..}. pdfT x s * (max 0 (g s) - max 0 (- g s))" + also have "\ = (LBINT s:{0..}. pdfT x s * (max 0 (g s) - max 0 (- g s)))" by (simp add: right_diff_distrib) - also have "\ = LBINT s:{0..}. pdfT x s * g s" + also have "\ = (LBINT s:{0..}. pdfT x s * g s)" using minus_max_eq_min by (metis (no_types, opaque_lifting) diff_zero max_def min_def minus_diff_eq) finally show ?thesis . qed corollary expectation_integral_pdfT: "alivex_PS.expectation (\\. g (T x \)) = integral {0..} (\s. pdfT x s * g s)" if "(\s. pdfT x s * g s) absolutely_integrable_on {0..}" "g \ borel_measurable lborel" for g :: "real \ real" proof - have [simp]: "set_integrable lborel {0..} (\s. pdfT x s * g s)" using that by (rewrite absolutely_integrable_on_iff_set_integrable; simp) show ?thesis apply (rewrite set_borel_integral_eq_integral(2)[THEN sym], simp) using that by (rewrite expectation_LBINT_pdfT; simp) qed -corollary e_LBINT_pdfT: "$e`\_x = LBINT s:{0..}. pdfT x s * s" +corollary e_LBINT_pdfT: "$e`\_x = (LBINT s:{0..}. pdfT x s * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ unfolding life_expect_def using expectation_LBINT_pdfT_nonneg by force corollary e_integral_pdfT: "$e`\_x = integral {0..} (\s. pdfT x s * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ unfolding life_expect_def using expectation_integral_pdfT_nonneg by force end corollary e_LBINT_pdfX: "$e`\_0 = (LBINT x:{0..}. pdfX x * x)" \ \Note that 0 = 0 holds when the life expectation diverges.\ using e_LBINT_pdfT pdfT0_X psi_pos' by presburger corollary e_integral_pdfX: "$e`\_0 = integral {0..} (\x. pdfX x * x)" \ \Note that 0 = 0 holds when the life expectation diverges.\ using e_integral_pdfT pdfT0_X psi_pos' by presburger subsubsection \Introduction of Force of Mortality\ definition force_mortal :: "real \ real" ("$\'__" [101] 200) where "$\_x \ MM_PS.hazard_rate X x" lemma mu_pdfX: "$\_x = pdfX x / ccdf (distr \ borel X) x" if "(cdf (distr \ borel X)) differentiable at x" for x::real unfolding force_mortal_def pdfX_def by (rewrite MM_PS.hazard_rate_deriv_cdf, simp_all add: that) lemma mu_unborn_0: "$\_x = 0" if "x < 0" for x::real apply (rewrite mu_pdfX) using cdfX_has_real_derivative_0_unborn real_differentiable_def that apply blast using pdfX_nonpos_0 that by auto lemma mu_beyond_0: "$\_x = 0" if "x \ $\" for x::real \ \Note that division by 0 is defined as 0 in Isabelle/HOL.\ unfolding force_mortal_def using MM_PS.hazard_rate_0_ccdf_0 ccdfX_0_equiv that by simp lemma mu_nonneg_differentiable: "$\_x \ 0" if "(cdf (distr \ borel X)) differentiable at x" for x::real apply (rewrite mu_pdfX, simp add: that) using pdfX_nonneg distrX_RD.ccdf_nonneg by simp lemma mu_nonneg_AE: "AE x in lborel. $\_x \ 0" using cdfX_differentiable_AE mu_nonneg_differentiable by auto lemma mu_measurable[measurable]: "(\x. $\_x) \ borel_measurable borel" proof - obtain S where "finite S" and "\x. x \ S \ (cdf (distr \ borel X)) differentiable at x" using cdfX_piecewise_differentiable unfolding piecewise_differentiable_on_def by blast thus ?thesis apply (rewrite measurable_discrete_difference [where f="\x. pdfX x / ccdf (distr \ borel X) x" and X=S], simp_all) by (simp_all add: MM_PS.ccdf_distr_measurable borel_measurable_divide countable_finite mu_pdfX) qed lemma mu_deriv_ccdf: "$\_x = - deriv (ccdf (distr \ borel X)) x / ccdf (distr \ borel X) x" if "(ccdf (distr \ borel X)) differentiable at x" "x < $\" for x::real unfolding force_mortal_def by (rewrite MM_PS.hazard_rate_deriv_ccdf; simp add: that) lemma mu_deriv_ln: "$\_x = - deriv (\x. ln (ccdf (distr \ borel X) x)) x" if "(ccdf (distr \ borel X)) differentiable at x" "x < $\" for x::real unfolding force_mortal_def apply (rewrite MM_PS.hazard_rate_deriv_ln_ccdf, simp_all add: that) using ccdfX_0_equiv that by force lemma p_exp_integral_mu: "$p_{t&x} = exp (- integral {x..x+t} (\y. $\_y))" if "x \ 0" "t \ 0" "x+t < $\" proof - have [simp]: "x < $\" using that by (meson add_increasing2 ereal_less linorder_not_le) have "$p_{t&x} = (ccdf (distr \ borel X) (x+t)) / (ccdf (distr \ borel X) x)" apply (rewrite p_PX, simp_all add: that) by (rewrite MM_PS.ccdf_distr_P, simp)+ simp also have "\ = exp (- integral {x..x+t} (MM_PS.hazard_rate X))" apply (rule MM_PS.ccdf_exp_cumulative_hazard, simp_all add: that) using ccdfX_piecewise_differentiable piecewise_differentiable_on_subset apply blast using ccdfX_continuous continuous_on_subset apply blast using ccdfX_0_equiv that ereal_less_le linorder_not_le by force also have "\ = exp (- integral {x..x+t} (\y. $\_y))" unfolding force_mortal_def by simp finally show ?thesis . qed corollary ccdfX_exp_integral_mu: "ccdf (distr \ borel X) x = exp (- integral {0..x} (\y. $\_y))" if "0 \ x \ x < $\" for x::real by (rewrite p_exp_integral_mu[where t=x and x=0, simplified, THEN sym]; simp add: that ccdfX_p) subsubsection \Properties of Force of Mortality\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin interpretation alivex_PS: prob_space "\ \ alive x" by (rule MM_PS.cond_prob_space_correct; simp add: alive_def) interpretation distrTx_RD: real_distribution "distr (\ \ alive x) borel (T x)" by simp lemma hazard_rate_Tx_mu: "alivex_PS.hazard_rate (T x) t = $\_(x+t)" if "t \ 0" "x+t < $\" for t::real proof - have [simp]: "\

(\ in \. X \ > x) > 0" by force moreover have [simp]: "\

(\ in \. X \ > x + t) > 0" using that by force moreover have [simp]: "{\ \ space (\ \ alive x). X \ > x + t} = {\ \ space \. X \ > x + t}" unfolding alive_def using that by (rewrite MM_PS.space_cond_prob_space, simp_all, force) hence [simp]: "{\ \ space (\ \ alive x). X \ > x + t} \ MM_PS.events" by simp ultimately have \[simp]: "\

(\ in (\ \ alive x). X \ > x + t) > 0" unfolding alive_def apply (rewrite MM_PS.cond_prob_space_cond_prob[THEN sym], (simp_all add: pred_def)[3]) unfolding cond_prob_def by (smt (verit) Collect_cong divide_pos_pos) have "alivex_PS.hazard_rate (T x) t = Lim (at_right 0) (\dt. \

(\ in (\ \ alive x). t < T x \ \ T x \ \ t + dt \ T x \ > t) / dt)" unfolding alivex_PS.hazard_rate_def by simp also have "\ = Lim (at_right 0) (\dt. \

(\ in (\ \ alive x). x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) / dt)" apply (rule Lim_cong, rule eventually_at_rightI[of 0 1], simp_all) unfolding futr_life_def cond_prob_def using Collect_cong by smt also have "\ = Lim (at_right 0) (\dt. \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) / dt)" proof - have "\dt. \

(\ in (\ \ alive x). x + t < X \ \ X \ \ x + t + dt \ X \ > x + t) = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" proof - fix dt let ?rngX = "\\. x + t < X \ \ X \ \ x + t + dt" have "\

(\ in (\ \ alive x). ?rngX \ \ X \ > x + t) = \

(\ in ((\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t}). ?rngX \)" using \ by (rewrite alivex_PS.cond_prob_space_cond_prob) simp_all also have "\ = \

(\ in (\ \ {\ \ space \. X \ > x + t}). ?rngX \)" proof - have "(\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t} = \ \ {\ \ space (\ \ alive x). X \ > x + t}" apply (rewrite MM_PS.cond_cond_prob_space, simp_all) unfolding alive_def using that by force also have "\ = \ \ {\ \ space \. X \ > x + t}" by simp finally have "(\ \ alive x) \ {\ \ space (\ \ alive x). X \ > x + t} = \ \ {\ \ space \. X \ > x + t}" . thus ?thesis by simp qed also have "\ = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" by (rewrite MM_PS.cond_prob_space_cond_prob, simp_all add: pred_def sets.sets_Collect_conj) finally show "\

(\ in (\ \ alive x). ?rngX \ \ X \ > x + t) = \

(\ in \. x + t < X \ \ X \ \ x + t + dt \ X \ > x + t)" . qed thus ?thesis apply - by (rule Lim_cong, rule eventually_at_rightI[of 0 1]; simp) qed also have "\ = $\_(x+t)" unfolding force_mortal_def MM_PS.hazard_rate_def by simp finally show ?thesis . qed lemma pdfTx_p_mu: "pdfT x t = $p_{t&x} * $\_(x+t)" if "(cdf (distr (\ \ alive x) borel (T x))) differentiable at t" "t > 0" for t::real proof (cases \x+t < $\\) case True hence [simp]: "t \ 0" and "(ccdf (distr (\ \ alive x) borel (T x))) t \ 0" using that p_0_equiv unfolding survive_def by auto have "deriv (cdf (distr (\ \ alive x) borel (T x))) t = ccdf (distr (\ \ alive x) borel (T x)) t * alivex_PS.hazard_rate (T x) t" by (rule alivex_PS.deriv_cdf_hazard_rate; simp add: that) thus ?thesis unfolding survive_def pdfT_def using hazard_rate_Tx_mu True that by simp next case False hence "x+t \ $\" by simp thus ?thesis using pdfTx_beyond_0 mu_beyond_0 by simp qed lemma deriv_t_p_mu: "deriv (\s. $p_{s&x}) t = - $p_{t&x} * $\_(x+t)" if "(\s. $p_{s&x}) differentiable at t" "t > 0" for t::real proof - let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" have diff: "?cdfTx differentiable at t" using that distrTx_RD.differentiable_cdf_ccdf unfolding survive_def by blast hence "deriv ?cdfTx t = $p_{t&x} * $\_(x+t)" using that pdfTx_p_mu pdfT_def by simp moreover have "deriv ?cdfTx t = - deriv (\s. $p_{s&x}) t" using that diff distrTx_RD.deriv_cdf_ccdf unfolding survive_def by presburger ultimately show ?thesis by simp qed lemma pdfTx_p_mu_AE: "AE s in lborel. s > 0 \ pdfT x s = $p_{s&x} * $\_(x+s)" proof - let ?cdfX = "cdf (distr \ borel X)" let ?cdfTx = "cdf (distr (\ \ alive x) borel (T x))" from cdfX_differentiable_AE obtain N where diff: "\r. r \ space lborel - N \ ?cdfX differentiable at r" and null: "N \ null_sets lborel" using AE_E3 by blast let ?N' = "{s. x+s \ N}" have "?N' \ null_sets lborel" using null_sets_translation[of N "-x", simplified] null by (simp add: add.commute) moreover have "{s \ space lborel. \ (s > 0 \ pdfT x s = $p_{s&x} * $\_(x+s))} \ ?N'" proof (rewrite Compl_subset_Compl_iff[THEN sym], safe) fix s::real assume "s \ space lborel" and "x+s \ N" and "s > 0" thus "pdfT x s = $p_{s&x} * $\_(x+s)" apply (intro pdfTx_p_mu, simp_all) by (rewrite differentiable_cdfX_cdfTx[THEN sym]; simp add: diff) qed ultimately show ?thesis using AE_I'[of ?N'] by simp qed -lemma LBINT_p_mu_q: "LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s) = $q_{f\t&x}" +lemma LBINT_p_mu_q: "(LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s)) = $q_{f\t&x}" if "t \ 0" "f \ 0" for t f :: real proof - - have "LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s) = LBINT s:{f<..f+t}. pdfT x s" + have "(LBINT s:{f<..f+t}. $p_{s&x} * $\_(x+s)) = (LBINT s:{f<..f+t}. pdfT x s)" apply (rule set_lebesgue_integral_cong_AE; simp) apply (simp add: survive_def) using pdfTx_p_mu_AE apply (rule AE_mp) using that by (intro always_eventually; simp add: ereal_less_le) also have "\ = enn2real (\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel)" proof - - have "\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel < \" + have "(\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel) < \" proof - - have "\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel \ \\<^sup>+s. ennreal (pdfT x s) \lborel" + have "(\\<^sup>+s\{f<..f+t}. ennreal (pdfT x s) \lborel) \ (\\<^sup>+s. ennreal (pdfT x s) \lborel)" by (smt (verit) indicator_simps le_zero_eq linorder_le_cases mult.right_neutral mult_zero_right nn_integral_mono) also have "\ < \" using nn_integral_pdfTx_1 by simp finally show ?thesis . qed thus ?thesis unfolding set_lebesgue_integral_def by (rewrite enn2real_nn_integral_eq_integral[where g="\s. pdfT x s * indicator {f<..f+t} s"]) (simp_all add: pdfTx_nonneg mult.commute ennreal_indicator ennreal_mult') qed also have "\ = measure (density lborel (pdfT x)) {f<..f+t}" unfolding measure_def by (rewrite emeasure_density; simp) also have "\ = measure (distr (\ \ alive x) lborel (T x)) {f<..f+t}" using distributed_pdfTx unfolding distributed_def by simp also have "\ = cdf (distr (\ \ alive x) lborel (T x)) (f+t) - cdf (distr (\ \ alive x) lborel (T x)) f" using that finite_borel_measure.cdf_diff_eq2 by (smt (verit) distrTx_RD.finite_borel_measure_axioms distr_cong sets_lborel) also have "\ = $q_{f\t&x}" using q_defer_q die_def that by (metis distr_cong sets_lborel x_lt_psi) finally show ?thesis . qed lemma set_integrable_p_mu: "set_integrable lborel {f<..f+t} (\s. $p_{s&x} * $\_(x+s))" if "t \ 0" "f \ 0" for t f :: real proof - have "(\s. $p_{s&x}) \ borel_measurable borel" unfolding survive_def by simp moreover have "AE s in lborel. f < s \ s \ f + t \ $p_{s&x} * $\_(x+s) = pdfT x s" using pdfTx_p_mu_AE apply (rule AE_mp) using that by (intro always_eventually; simp add: ereal_less_le) moreover have "set_integrable lborel {f<..f+t} (pdfT x)" using pdfTx_set_integrable by simp ultimately show ?thesis by (rewrite set_integrable_cong_AE[where g="pdfT x"]; simp) qed lemma p_mu_has_integral_q_defer_Ioc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{f\t&x}) {f<..f+t}" if "t \ 0" "f \ 0" for t f :: real apply (rewrite LBINT_p_mu_q[THEN sym], simp_all add: that) apply (rewrite set_borel_integral_eq_integral, simp add: set_integrable_p_mu that) by (rewrite has_integral_integral[THEN sym]; simp add: set_borel_integral_eq_integral set_integrable_p_mu that) lemma p_mu_has_integral_q_defer_Icc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{f\t&x}) {f..f+t}" if "t \ 0" "f \ 0" for t f :: real proof - have "negligible {f}" by simp hence [simp]: "negligible ({f..f+t} - {f<..f+t})" by (smt (verit) Diff_iff atLeastAtMost_iff greaterThanAtMost_iff insertI1 negligible_sing negligible_subset subsetI) have [simp]: "negligible ({f<..f+t} - {f..f+t})" by (simp add: subset_eq) show ?thesis apply (rewrite has_integral_spike_set_eq[where T="{f<..f+t}"]) apply (rule negligible_subset[of "{f..f+t} - {f<..f+t}"], simp, blast) apply (rule negligible_subset[of "{f<..f+t} - {f..f+t}"], simp, blast) using p_mu_has_integral_q_defer_Ioc that by simp qed corollary p_mu_has_integral_q_Icc: "((\s. $p_{s&x} * $\_(x+s)) has_integral $q_{t&x}) {0..t}" if "t \ 0" for t::real using p_mu_has_integral_q_defer_Icc[where f=0] that by simp corollary p_mu_integrable_on_Icc: "(\s. $p_{s&x} * $\_(x+s)) integrable_on {0..t}" if "t \ 0" for t::real using p_mu_has_integral_q_Icc that by auto lemma e_ennreal_p_mu: "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = - \\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" + (\\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel)" proof - have [simp]: "sym_diff {0..} {0<..} = {0::real}" by force - have "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = \\<^sup>+s\{0..}. ennreal (pdfT x s * s) \lborel" + have "(\\<^sup>+\. ennreal (T x \) \(\ \ alive x)) = (\\<^sup>+s\{0..}. ennreal (pdfT x s * s) \lborel)" by (rewrite nn_integral_T_pdfT[where g="\s. s"]; simp) - also have "\ = \\<^sup>+s\{0<..}. ennreal (pdfT x s * s) \lborel" + also have "\ = (\\<^sup>+s\{0<..}. ennreal (pdfT x s * s) \lborel)" by (rewrite nn_integral_null_delta; force) - also have "\ = \\<^sup>+s\{0<..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" + also have "\ = (\\<^sup>+s\{0<..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel)" apply (rule nn_integral_cong_AE) using pdfTx_p_mu_AE apply (rule AE_mp, intro AE_I2) by force - also have "\ = \\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel" + also have "\ = (\\<^sup>+s\{0..}. ennreal ($p_{s&x} * $\_(x+s) * s) \lborel)" by (rewrite nn_integral_null_delta[THEN sym]; force) finally show ?thesis . qed -lemma e_LBINT_p_mu: "$e`\_x = LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s" +lemma e_LBINT_p_mu: "$e`\_x = (LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - let ?f = "\s. $p_{s&x} * $\_(x+s) * s" have [simp]: "(\s. ?f s * indicat_real {0..} s) \ borel_measurable borel" by measurable (simp_all add: survive_def) hence [simp]: "(\s. indicat_real {0..} s * ?f s) \ borel_measurable borel" by (meson measurable_cong mult.commute) have [simp]: "AE s in lborel. ?f s * indicat_real {0..} s \ 0" proof - have "AE s in lborel. pdfT x s * s * indicat_real {0..} s \ 0" using pdfTx_nonneg pdfTx_nonpos_0 x_lt_psi by (intro AE_I2) (smt (verit, del_insts) indicator_pos_le mult_eq_0_iff mult_nonneg_nonneg) thus ?thesis apply (rule AE_mp) using pdfTx_p_mu_AE apply (rule AE_mp) by (rule AE_I2) (metis atLeast_iff indicator_simps(2) mult_eq_0_iff order_le_less) qed hence [simp]: "AE s in lborel. indicat_real {0..} s * ?f s \ 0" by (metis (no_types, lifting) AE_cong mult.commute) show ?thesis proof (cases \integrable (\ \ alive x) (T x)\) case True hence "ennreal ($e`\_x) = \\<^sup>+\. ennreal (T x \) \(\ \ alive x)" unfolding life_expect_def apply (rewrite nn_integral_eq_integral, simp_all) by (smt (verit) AE_I2 alivex_Tx_pos) also have "\ = \\<^sup>+s. ennreal (?f s * indicat_real {0..} s) \lborel" by (simp add: nn_integral_set_ennreal e_ennreal_p_mu) also have "\ = ennreal (LBINT s:{0..}. ?f s)" proof - have "integrable lborel (\s. ?f s * indicat_real {0..} s)" proof - have "(\\<^sup>+s. ennreal (?f s * indicat_real {0..} s) \lborel) < \" by (metis calculation ennreal_less_top infinity_ennreal_def) thus ?thesis by (intro integrableI_nonneg; simp) qed thus ?thesis unfolding set_lebesgue_integral_def by (rewrite nn_integral_eq_integral, simp_all) (meson mult.commute) qed finally have "ennreal ($e`\_x) = ennreal (LBINT s:{0..}. ?f s)" . - moreover have "LBINT s:{0..}. ?f s \ 0" + moreover have "(LBINT s:{0..}. ?f s) \ 0" unfolding set_lebesgue_integral_def by (rule integral_nonneg_AE) simp ultimately show ?thesis using e_nonneg by simp next case False hence "$e`\_x = 0" unfolding life_expect_def using not_integrable_integral_eq by force - also have "\ = LBINT s:{0..}. ?f s" + also have "\ = (LBINT s:{0..}. ?f s)" proof - have "\ = \\<^sup>+\. ennreal (T x \) \(\ \ alive x)" using nn_integral_nonneg_infinite False by (smt (verit) AE_cong Tx_alivex_measurable alivex_PS.AE_prob_1 alivex_PS.prob_space alivex_Tx_pos nn_integral_cong) hence "0 = enn2real (\\<^sup>+s\{0..}. ennreal (?f s) \lborel)" using e_ennreal_p_mu by simp - also have "\ = LBINT s:{0..}. ?f s" + also have "\ = (LBINT s:{0..}. ?f s)" unfolding set_lebesgue_integral_def apply (rewrite integral_eq_nn_integral, simp_all) by (simp add: indicator_mult_ennreal mult.commute) finally show ?thesis by simp qed finally show ?thesis . qed qed lemma e_integral_p_mu: "$e`\_x = integral {0..} (\s. $p_{s&x} * $\_(x+s) * s)" \ \Note that 0 = 0 holds when the life expectation diverges.\ proof - - have "LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s = integral {0..} (\s. $p_{s&x} * $\_(x+s) * s)" + have "(LBINT s:{0..}. $p_{s&x} * $\_(x+s) * s) = integral {0..} (\s. $p_{s&x} * $\_(x+s) * s)" proof - have "AE s in lborel. s \ 0 \ $p_{s&x} * $\_(x+s) * s \ 0" proof - have "AE s in lborel. $\_(x+s) \ 0" by (rule AE_translation, rule mu_nonneg_AE) with p_nonneg show ?thesis by force qed moreover have "(\s. $p_{s&x} * $\_(x+s) * s) \ borel_measurable borel" unfolding survive_def by simp ultimately show ?thesis by (intro set_borel_integral_eq_integral_nonneg_AE; simp) qed thus ?thesis using e_LBINT_p_mu by simp qed end lemma p_has_real_derivative_x_ccdfX: "((\y. $p_{t&y}) has_real_derivative ((deriv svl (x+t) * svl x - svl (x+t) * deriv svl x) / (svl x)\<^sup>2)) (at x)" if "svl \ ccdf (distr \ borel X)" "svl differentiable at x" "svl differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real proof - have "open {y. svl differentiable at y}" using ccdfX_differentiable_open_set that by simp with that obtain e0 where e0_pos: "e0 > 0" and ball_e0: "ball x e0 \ {y. svl differentiable at y}" using open_contains_ball by blast define e where "e \ if $\ < \ then min e0 (real_of_ereal $\ - x) else e0" have "e > 0 \ ball x e \ {y. svl y \ 0 \ svl differentiable at y}" proof (cases \$\ < \\) case True hence "e > 0" proof - have "real_of_ereal $\ - x > 0" using not_inftyI True that by force with e0_pos show ?thesis unfolding e_def using True by simp qed moreover have "ball x e \ {y. svl y \ 0}" proof - have "e \ real_of_ereal $\ - x" unfolding e_def using True by simp hence "ball x e \ {..< real_of_ereal $\}" unfolding ball_def dist_real_def by force thus ?thesis using that ccdfX_0_equiv using True ereal_less_real_iff psi_nonneg by auto qed moreover have "ball x e \ {y. svl differentiable at y}" proof - have "e \ e0" unfolding e_def using True by simp hence "ball x e \ ball x e0" by force with ball_e0 show ?thesis by simp qed ultimately show ?thesis by blast next case False hence "ball x e \ {y. svl y \ 0}" using ccdfX_0_equiv that by simp with False show ?thesis unfolding e_def using e0_pos ball_e0 by force qed hence e_pos: "e > 0" and ball_e: "ball x e \ {y. svl y \ 0 \ svl differentiable at y}" by auto hence ball_svl: "\y. y \ ball x e \ svl y \ 0 \ svl field_differentiable at y" using differentiable_eq_field_differentiable_real by blast hence "\y. y \ ball x e \ $p_{t&y} = svl (y+t) / svl y" unfolding survive_def using that ccdfX_0_equiv by (rewrite ccdfTx_ccdfX, simp_all) force moreover from ball_svl have "((\y. svl (y+t) / svl y) has_real_derivative ((deriv svl (x+t) * svl x - svl (x+t) * deriv svl x) / (svl x)\<^sup>2)) (at x)" apply (rewrite power2_eq_square) apply (rule DERIV_divide) using DERIV_deriv_iff_real_differentiable DERIV_shift that apply blast using that DERIV_deriv_iff_real_differentiable apply simp by (simp add: e_pos) ultimately show ?thesis using e_pos apply (rewrite has_field_derivative_cong_eventually[where g="\y. svl (y+t) / svl y"], simp_all) by (smt (verit) dist_commute eventually_at) qed lemma p_has_real_derivative_x_p_mu: "((\y. $p_{t&y}) has_real_derivative $p_{t&x} * ($\_x - $\_(x+t))) (at x)" if "ccdf (distr \ borel X) differentiable at x" "ccdf (distr \ borel X) differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real proof (cases \x+t < $\\) case True let ?svl = "ccdf (distr \ borel X)" have [simp]: "?svl x \ 0" using that ccdfX_0_equiv by (smt (verit) le_ereal_le linorder_not_le) have [simp]: "?svl field_differentiable at (x+t)" using that differentiable_eq_field_differentiable_real by simp have "((\y. $p_{t&y}) has_real_derivative ((deriv ?svl (x+t) * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2)) (at x)" using p_has_real_derivative_x_ccdfX that by simp moreover have "(deriv ?svl (x+t) * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2 = $p_{t&x} * ($\_x - $\_(x+t))" (is "?LHS = ?RHS") proof - have "deriv ?svl (x+t) = deriv (\y. ?svl (y+t)) x" using that by (metis DERIV_deriv_iff_real_differentiable DERIV_imp_deriv DERIV_shift) hence "?LHS = (deriv (\y. ?svl (y+t)) x * ?svl x - ?svl (x+t) * deriv ?svl x) / (?svl x)\<^sup>2" by simp also have "\ = (deriv (\y. ?svl (y+t)) x - ?svl (x+t) * deriv ?svl x / ?svl x) / ?svl x" by (simp add: add_divide_eq_if_simps(4) power2_eq_square) also have "\ = (- ?svl (x+t) * $\_(x+t) + ?svl (x+t) * $\_x) / ?svl x" proof - have "deriv (\y. ?svl (y+t)) x = - ?svl (x+t) * $\_(x+t)" apply (rewrite add.commute, rewrite deriv_shift[THEN sym], rewrite add.commute, simp) using add.commute that by (metis MM_PS.deriv_ccdf_hazard_rate X_RV force_mortal_def) moreover have "- ?svl (x+t) * deriv ?svl x / ?svl x = ?svl (x+t) * $\_x" using that by (simp add: MM_PS.deriv_ccdf_hazard_rate force_mortal_def) ultimately show ?thesis by simp qed also have "\ = ?svl (x+t) * ($\_x - $\_(x+t)) / ?svl x" by (simp add: mult_diff_mult) also have "\ = ?RHS" unfolding survive_def using ccdfTx_ccdfX that by simp ultimately show ?thesis by simp qed ultimately show ?thesis by simp next case False hence ptx_0: "$p_{t&x} = 0" using p_0_equiv that by simp moreover have "((\y. $p_{t&y}) has_real_derivative 0) (at x)" proof - have "\y. x < y \ y < $\ \ $p_{t&y} = 0" using False p_0_equiv that by (smt (verit, ccfv_SIG) ereal_less_le linorder_not_le) hence "\\<^sub>F x in at_right x. $p_{t&x} = 0" apply (rewrite eventually_at_right_field) using that by (meson ereal_dense2 ereal_le_less less_eq_ereal_def less_ereal.simps) hence "((\y. $p_{t&y}) has_real_derivative 0) (at_right x)" using ptx_0 by (rewrite has_field_derivative_cong_eventually[where g="\_. 0"]; simp) thus ?thesis using vector_derivative_unique_within p_has_real_derivative_x_ccdfX that by (metis has_field_derivative_at_within has_real_derivative_iff_has_vector_derivative trivial_limit_at_right_real) qed ultimately show ?thesis by simp qed corollary deriv_x_p_mu: "deriv (\y. $p_{t&y}) x = $p_{t&x} * ($\_x - $\_(x+t))" if "ccdf (distr \ borel X) differentiable at x" "ccdf (distr \ borel X) differentiable at (x+t)" "t \ 0" "x < $\" for x t :: real using that p_has_real_derivative_x_p_mu DERIV_imp_deriv by blast lemma e_has_derivative_mu_e: "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. set_integrable lborel {x..} (ccdf (distr \ borel X))" "ccdf (distr \ borel X) differentiable at x" "x\{a<.. $\" for a b x :: real proof - let ?svl = "ccdf (distr \ borel X)" have x_lt_psi[simp]: "x < $\" using that ereal_le_less by simp hence svlx_neq0[simp]: "?svl x \ 0" by (simp add: ccdfX_0_equiv linorder_not_le) define d ::real where "d \ min (b-x) (x-a)" have d_pos: "d > 0" unfolding d_def using that ereal_less_real_iff by force have e_svl: "\y. y < $\ \ $e`\_y = (LBINT t:{0..}. ?svl (y+t)) / ?svl y" apply (rewrite e_LBINT_p, simp) apply (rewrite set_integral_divide_zero[THEN sym]) apply (rule set_lebesgue_integral_cong, simp_all) unfolding survive_def using ccdfTx_ccdfX by force have "((\y. LBINT t:{0..}. ?svl (y+t)) has_real_derivative (- ?svl x)) (at x)" proof - { fix y assume "dist y x < d" hence y_ab: "y \ {a<..) ?svl" by (rewrite set_integrable_discrete_difference[where X="{y}"]; simp) force moreover have "\u. ((\u. u-y) has_real_derivative (1-0)) (at u)" by (rule derivative_intros) auto moreover have "\u. y < u \ isCont (\t. ?svl (y+t)) (u-y)" apply (rewrite add.commute, rewrite isCont_shift, simp) using ccdfX_continuous continuous_on_eq_continuous_within by blast moreover have "((ereal \ (\u. u-y) \ real_of_ereal) \ 0) (at_right (ereal y))" by (smt (verit, ccfv_SIG) LIM_zero Lim_cong_within ereal_tendsto_simps1(2) ereal_tendsto_simps2(1) tendsto_ident_at zero_ereal_def) moreover have "((ereal \ (\u. u-y) \ real_of_ereal) \ \) (at_left \)" proof - have "\r u. r+y+1 \ u \ r < u-y" by auto hence "\r. \\<^sub>F u in at_top. r < u - y" by (rule eventually_at_top_linorderI) thus ?thesis by (rewrite ereal_tendsto_simps, rewrite tendsto_PInfty) simp qed ultimately have "(LBINT t=0..\. ?svl (y+t)) = (LBINT u=y..\. ?svl (y+(u-y)) * 1)" using distrX_RD.ccdf_nonneg by (intro interval_integral_substitution_nonneg(2); simp) moreover have "(LBINT t:{0..}. ?svl (y+t)) = (LBINT t=0..\. ?svl (y+t))" unfolding interval_lebesgue_integral_def einterval_def apply simp by (rule set_integral_discrete_difference[where X="{0}"]; force) moreover have "(LBINT u=y..\. ?svl (y+(u-y))*1) = (LBINT u:{y..}. ?svl u)" unfolding interval_lebesgue_integral_def einterval_def apply simp by (rule set_integral_discrete_difference[where X="{y}"]; force) ultimately have "(LBINT t:{0..}. ?svl (y+t)) = (LBINT u:{y..}. ?svl u)" by simp } - hence "\\<^sub>F y in nhds x. LBINT t:{0..}. ?svl (y+t) = LBINT u:{y..}. ?svl u" + hence "\\<^sub>F y in nhds x. (LBINT t:{0..}. ?svl (y+t)) = (LBINT u:{y..}. ?svl u)" using d_pos by (rewrite eventually_nhds_metric) auto moreover have "((\y. LBINT u:{y..}. ?svl u) has_real_derivative (- ?svl x)) (at x)" proof - have "((\y. integral {y..b} ?svl) has_real_derivative (- ?svl x)) (at x within {a..b})" using that apply (intro integral_has_real_derivative'; simp) using ccdfX_continuous continuous_on_subset by blast hence "((\y. integral {y..b} ?svl) has_real_derivative (- ?svl x)) (at x)" using that apply (rewrite at_within_open[where S="{a<..\<^sub>F y in nhds x. LBINT u:{y..b}. ?svl u = integral {y..b} ?svl" + moreover have "\\<^sub>F y in nhds x. (LBINT u:{y..b}. ?svl u) = integral {y..b} ?svl" apply (rewrite eventually_nhds_metric) using d_pos by (metis ccdfX_integrable_Icc set_borel_integral_eq_integral(2)) ultimately have "((\y. LBINT u:{y..b}. ?svl u) has_real_derivative (- ?svl x)) (at x)" by (rewrite DERIV_cong_ev; simp) hence "((\y. (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)) has_real_derivative (- ?svl x)) (at x)" by (rewrite to "- ?svl x + 0" add_0_right[THEN sym], rule DERIV_add; simp) - moreover have "\\<^sub>F y in nhds x. - LBINT u:{y..}. ?svl u = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" + moreover have "\\<^sub>F y in nhds x. (LBINT u:{y..}. ?svl u) = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" proof - { fix y assume "dist y x < d" hence y_ab: "y \ {a<.. {b<..} = {}" "{y..} = {y..b} \ {b<..}" using y_ab by force+ - ultimately have "LBINT u:{y..}. ?svl u = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" + ultimately have "(LBINT u:{y..}. ?svl u) = (LBINT u:{y..b}. ?svl u) + (LBINT u:{b<..}. ?svl u)" using set_integral_Un by simp } thus ?thesis using d_pos by (rewrite eventually_nhds_metric) blast qed ultimately show ?thesis by (rewrite has_field_derivative_cong_ev; simp) qed ultimately show ?thesis by (rewrite DERIV_cong_ev; simp) qed moreover have "(?svl has_real_derivative (deriv ?svl x)) (at x)" using that DERIV_deriv_iff_real_differentiable by blast ultimately have "((\y. (LBINT t:{0..}. ?svl (y+t)) / ?svl y) has_real_derivative ((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x)) (at x)" by (rule DERIV_divide) simp moreover have "eventually (\y. (LBINT t:{0..}. ?svl (y+t)) / ?svl y = $e`\_y) (at x)" proof - { fix y assume "dist y x < d" "y \ x" hence "y < $\" unfolding dist_real_def d_def using that ereal_le_less by fastforce hence "$e`\_y = (LBINT t:{0..}. ?svl (y+t)) / ?svl y" by (rule e_svl) } thus ?thesis apply (rewrite eventually_at_filter, rewrite eventually_nhds_metric) using d_pos that by metis qed ultimately have "((\y. $e`\_y) has_real_derivative ((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x)) (at x)" using e_svl by (rewrite has_field_derivative_cong_eventually[THEN sym]; simp) moreover have "((- ?svl x) * ?svl x - (LBINT t:{0..}. ?svl (x+t)) * deriv ?svl x) / (?svl x * ?svl x) = $\_x * $e`\_x - 1" (is "?LHS = ?RHS") proof - have "?LHS = -1 + (LBINT t:{0..}. ?svl (x+t)) / ?svl x * (- deriv ?svl x / ?svl x)" by simp (smt (verit) svlx_neq0 add_divide_distrib divide_eq_minus_1_iff mult_minus_left real_divide_square_eq) also have "\ = -1 + $e`\_x * $\_x" using e_svl mu_deriv_ccdf that by force also have "\ = ?RHS" by simp finally show ?thesis . qed ultimately show ?thesis by simp qed corollary e_has_derivative_mu_e': "((\x. $e`\_x) has_real_derivative ($\_x * $e`\_x - 1)) (at x)" if "\x. x\{a<.. ccdf (distr \ borel X) integrable_on {x..}" "ccdf (distr \ borel X) differentiable at x" "x\{a<.. $\" for a b x :: real using that apply (intro e_has_derivative_mu_e[where a=a], simp_all) using distrX_RD.ccdf_nonneg by (rewrite integrable_on_iff_set_integrable_nonneg; simp) subsubsection \Properties of Curtate Life Expectation\ context fixes x::real assumes x_lt_psi[simp]: "x < $\" begin lemma isCont_p_nat: "isCont (\t. $p_{t&x}) (k + (1::real))" for k::nat proof - fix k::nat have "continuous_on {0<..} (\t. $p_{t&x})" unfolding survive_def using ccdfTx_continuous_on_nonneg continuous_on_subset Ioi_le_Ico x_lt_psi by blast hence "\t\{0<..}. isCont (\t. $p_{t&x}) t" using continuous_on_eq_continuous_at open_greaterThan by blast thus "isCont (\t. $p_{t&x}) (real k+1)" by force qed lemma curt_e_sum_p_smooth: "$e_x = (\k. $p_{k+1&x})" if "summable (\k. $p_{k+1&x})" using curt_e_sum_p isCont_p_nat that by simp lemma curt_e_rec_smooth: "$e_x = $p_x * (1 + $e_(x+1))" if "summable (\k. $p_{k+1&x})" "x+1 < $\" using curt_e_rec isCont_p_nat that by simp lemma curt_e_sum_p_finite_smooth: "$e_x = (\k $\" for n::nat using curt_e_sum_p_finite isCont_p_nat that by simp lemma temp_curt_e_sum_p_smooth: "$e_{x:n} = (\k" "n \ 0" for n::nat using temp_curt_e_rec isCont_p_nat that by simp end end subsection \Finite Survival Function\ locale finite_survival_function = survival_model + assumes psi_finite[simp]: "$\ < \" begin definition ult_age :: nat ("$\") where "$\ \ LEAST x::nat. ccdf (distr \ borel X) x = 0" \ \the conventional notation for ultimate age\ lemma ccdfX_ceil_psi_0: "ccdf (distr \ borel X) \real_of_ereal $\\ = 0" proof - have "real_of_ereal $\ \ \real_of_ereal $\\" by simp thus ?thesis using ccdfX_0_equiv psi_finite ccdfX_psi_0 le_ereal_le by presburger qed lemma ccdfX_omega_0: "ccdf (distr \ borel X) $\ = 0" unfolding ult_age_def proof (rule LeastI_ex) have "\real_of_ereal $\\ \ 0" using psi_nonneg real_of_ereal_pos by fastforce thus "\x::nat. ccdf (distr \ borel X) (real x) = 0" using ccdfX_ceil_psi_0 by (metis of_int_of_nat_eq zero_le_imp_eq_int) qed corollary psi_le_omega: "$\ \ $\" using ccdfX_0_equiv ccdfX_omega_0 by blast corollary omega_pos: "$\ > 0" using psi_le_omega order.strict_iff_not by fastforce lemma omega_ceil_psi: "$\ = \real_of_ereal $\\" proof (rule antisym) let ?cpsi = "\real_of_ereal $\\" have \: "?cpsi \ 0" using psi_nonneg real_of_ereal_pos by fastforce moreover have "ccdf (distr \ borel X) ?cpsi = 0" by (rule ccdfX_ceil_psi_0) ultimately have "$\ \ nat ?cpsi" unfolding ult_age_def using wellorder_Least_lemma of_nat_nat by smt thus "int $\ \ ?cpsi" using le_nat_iff \ by blast next show "\real_of_ereal $\\ \ int $\" using psi_le_omega by (simp add: ceiling_le_iff real_of_ereal_ord_simps(2)) qed lemma ccdfX_0_equiv_nat: "ccdf (distr \ borel X) x = 0 \ x \ $\" for x::nat proof assume "ccdf (distr \ borel X) (real x) = 0" thus "x \ $\" unfolding ult_age_def using wellorder_Least_lemma by fastforce next assume "x \ $\" hence "ereal (real x) \ $\" using psi_le_omega le_ereal_le of_nat_mono by blast thus "ccdf (distr \ borel X) (real x) = 0" using ccdfX_0_equiv by simp qed lemma psi_le_iff_omega_le: "$\ \ x \ $\ \ x" for x::nat using ccdfX_0_equiv ccdfX_0_equiv_nat by presburger context fixes x::nat assumes x_lt_omega[simp]: "x < $\" begin lemma x_lt_psi[simp]: "x < $\" using x_lt_omega psi_le_iff_omega_le by (meson linorder_not_less) lemma p_0_1_nat: "$p_{0&x} = 1" using p_0_1 by simp lemma p_0_equiv_nat: "$p_{t&x} = 0 \ x+t \ $\" for t::nat using psi_le_iff_omega_le p_0_equiv by (metis of_nat_add x_lt_psi) lemma q_0_0_nat: "$q_{0&x} = 0" using p_q_1 p_0_1_nat by (smt (verit) x_lt_psi) lemma q_1_equiv_nat: "$q_{t&x} = 1 \ x+t \ $\" for t::nat using p_q_1 p_0_equiv_nat by (smt (verit) x_lt_psi) lemma q_defer_old_0_nat: "$q_{f\t&x} = 0" if "$\ \ x+f" for f t :: nat using q_defer_old_0 psi_le_iff_omega_le that by (metis of_nat_0_le_iff of_nat_add x_lt_psi) lemma curt_e_sum_P_finite_nat: "$e_x = (\k(\ in \. T x \ \ k + 1 \ T x \ > 0))" if "x+n \ $\" for n::nat apply (rule curt_e_sum_P_finite, simp) using that psi_le_iff_omega_le by (smt (verit) le_ereal_less of_nat_add) lemma curt_e_sum_p_finite_nat: "$e_x = (\kk::nat. k < n \ isCont (\t. $p_{t&x}) (real k + 1)" "x+n \ $\" for n::nat apply (rule curt_e_sum_p_finite, simp_all add: that) using that psi_le_iff_omega_le by (smt (verit) le_ereal_less of_nat_add) end lemma q_omega_1: "$q_($\-1) = 1" using q_1_equiv_nat by (metis diff_less dual_order.refl le_diff_conv of_nat_1 omega_pos zero_less_one) end end diff --git a/thys/Green/Integrals.thy b/thys/Green/Integrals.thy --- a/thys/Green/Integrals.thy +++ b/thys/Green/Integrals.thy @@ -1,946 +1,946 @@ theory Integrals imports "HOL-Analysis.Analysis" General_Utils begin lemma gauge_integral_Fubini_universe_x: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" assumes fun_lesbegue_integrable: "integrable lborel f" and x_axis_integral_measurable: "(\x. integral UNIV (\y. f(x, y))) \ borel_measurable lborel" shows "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" proof - have f_is_measurable: "f \ borel_measurable lborel" using fun_lesbegue_integrable and borel_measurable_integrable by auto have fun_lborel_prod_integrable: "integrable (lborel \\<^sub>M lborel) f" using fun_lesbegue_integrable by (simp add: lborel_prod) then have region_integral_is_one_twoD_integral: - "LBINT x. LBINT y. f (x, y) = integral\<^sup>L (lborel \\<^sub>M lborel) f" + "(LBINT x. (LBINT y. f (x, y))) = integral\<^sup>L (lborel \\<^sub>M lborel) f" using lborel_pair.integral_fst' by auto - then have AE_one_D_integrals_eq: "AE x in lborel. LBINT y. f (x, y) = integral UNIV (\y. f(x,y))" + then have AE_one_D_integrals_eq: "AE x in lborel. (LBINT y. f (x, y)) = integral UNIV (\y. f(x,y))" proof - have "AE x in lborel. integrable lborel (\y. f(x,y))" using lborel_pair.AE_integrable_fst' and fun_lborel_prod_integrable by blast then show ?thesis using integral_lborel and always_eventually and AE_mp by fastforce qed have one_D_integral_measurable: "(\x. LBINT y. f (x, y)) \ borel_measurable lborel" using f_is_measurable and lborel.borel_measurable_lebesgue_integral by auto then have second_lesbegue_integral_eq: - "LBINT x. LBINT y. f (x, y) = LBINT x. (integral UNIV (\y. f(x,y)))" + "(LBINT x. (LBINT y. f (x, y))) = (LBINT x. (integral UNIV (\y. f(x,y))))" using x_axis_integral_measurable and integral_cong_AE and AE_one_D_integrals_eq by blast have "integrable lborel (\x. LBINT y. f (x, y))" using fun_lborel_prod_integrable and lborel_pair.integrable_fst' by auto then have oneD_gauge_integral_lesbegue_integrable: "integrable lborel (\x. integral UNIV (\y. f(x,y)))" using x_axis_integral_measurable and AE_one_D_integrals_eq and integrable_cong_AE_imp by blast then show one_D_gauge_integral_integrable: "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" using integrable_on_lborel by auto - have "LBINT x. (integral UNIV (\y. f(x,y))) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" + have "(LBINT x. (integral UNIV (\y. f(x,y)))) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" using integral_lborel oneD_gauge_integral_lesbegue_integrable by fastforce then have twoD_lesbeuge_eq_twoD_gauge: - "LBINT x. LBINT y. f (x, y) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" + "(LBINT x. (LBINT y. f (x, y))) = integral UNIV (\x. integral UNIV (\y. f(x, y)))" using second_lesbegue_integral_eq by auto then show "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" using fun_lesbegue_integrable and integral_lborel and region_integral_is_one_twoD_integral by (metis lborel_prod) qed lemma gauge_integral_Fubini_universe_y: fixes f :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" assumes fun_lesbegue_integrable: "integrable lborel f" and y_axis_integral_measurable: "(\x. integral UNIV (\y. f(y, x))) \ borel_measurable lborel" shows "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" proof - have f_is_measurable: "f \ borel_measurable lborel" using fun_lesbegue_integrable and borel_measurable_integrable by auto have fun_lborel_prod_integrable: "integrable (lborel \\<^sub>M lborel) f" using fun_lesbegue_integrable by (simp add: lborel_prod) then have region_integral_is_one_twoD_integral: - "LBINT x. LBINT y. f (y, x) = integral\<^sup>L (lborel \\<^sub>M lborel) f" + "(LBINT x. (LBINT y. f (y, x))) = integral\<^sup>L (lborel \\<^sub>M lborel) f" using lborel_pair.integral_fst' f_is_measurable lborel_pair.integrable_product_swap lborel_pair.integral_fst lborel_pair.integral_product_swap lborel_prod by force - then have AE_one_D_integrals_eq: "AE x in lborel. LBINT y. f (y, x) = integral UNIV (\y. f(y,x))" + then have AE_one_D_integrals_eq: "AE x in lborel. (LBINT y. f (y, x)) = integral UNIV (\y. f(y,x))" proof - have "AE x in lborel. integrable lborel (\y. f(y,x))" using lborel_pair.AE_integrable_fst' and fun_lborel_prod_integrable lborel_pair.AE_integrable_fst lborel_pair.integrable_product_swap by blast then show ?thesis using integral_lborel and always_eventually and AE_mp by fastforce qed have one_D_integral_measurable: "(\x. LBINT y. f (y, x)) \ borel_measurable lborel" using f_is_measurable and lborel.borel_measurable_lebesgue_integral by auto then have second_lesbegue_integral_eq: - "LBINT x. LBINT y. f (y, x) = LBINT x. (integral UNIV (\y. f(y, x)))" + "(LBINT x. (LBINT y. f (y, x))) = (LBINT x. (integral UNIV (\y. f(y, x))))" using y_axis_integral_measurable and integral_cong_AE and AE_one_D_integrals_eq by blast have "integrable lborel (\x. LBINT y. f (y, x))" using fun_lborel_prod_integrable and lborel_pair.integrable_fst' by (simp add: lborel_pair.integrable_fst lborel_pair.integrable_product_swap) then have oneD_gauge_integral_lesbegue_integrable: "integrable lborel (\x. integral UNIV (\y. f(y, x)))" using y_axis_integral_measurable and AE_one_D_integrals_eq and integrable_cong_AE_imp by blast then show one_D_gauge_integral_integrable: "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" using integrable_on_lborel by auto - have "LBINT x. (integral UNIV (\y. f(y, x))) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" + have "(LBINT x. (integral UNIV (\y. f(y, x)))) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using integral_lborel oneD_gauge_integral_lesbegue_integrable by fastforce then have twoD_lesbeuge_eq_twoD_gauge: - "LBINT x. LBINT y. f (y, x) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" + "(LBINT x. (LBINT y. f (y, x))) = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using second_lesbegue_integral_eq by auto then show "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using fun_lesbegue_integrable and integral_lborel and region_integral_is_one_twoD_integral by (metis lborel_prod) qed lemma gauge_integral_Fubini_curve_bounded_region_x: fixes f g :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" and g1 g2:: "'a \ 'b" and s:: "('a * 'b) set" assumes fun_lesbegue_integrable: "integrable lborel f" and x_axis_gauge_integrable: "\x. (\y. f(x, y)) integrable_on UNIV" and (*IS THIS redundant? NO IT IS NOT*) x_axis_integral_measurable: "(\x. integral UNIV (\y. f(x, y))) \ borel_measurable lborel" and f_is_g_indicator: "f = (\x. if x \ s then g x else 0)" and s_is_bounded_by_g1_and_g2: "s = {(x,y). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. (g1 x) \ i \ y \ i \ y \ i \ (g2 x) \ i)}" shows "integral s g = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" proof - have two_D_integral_to_one_D: "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(x,y)))" using gauge_integral_Fubini_universe_x and fun_lesbegue_integrable and x_axis_integral_measurable by auto have one_d_integral_integrable: "(\x. integral UNIV (\y. f(x,y))) integrable_on UNIV" using gauge_integral_Fubini_universe_x(2) and assms by blast have case_x_in_range: "\ x \ cbox a b. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)) = integral UNIV (\y. f(x,y))" proof fix x:: 'a assume within_range: "x \ (cbox a b)" let ?f_one_D_spec = "(\y. if y \ (cbox (g1 x) (g2 x)) then g(x,y) else 0)" have f_one_D_region: "(\y. f(x,y)) = (\y. if y \ cbox (g1 x) (g2 x) then g(x,y) else 0)" proof fix y::'b show "f (x, y) = (if y \ (cbox (g1 x) (g2 x)) then g (x, y) else 0)" apply (simp add: f_is_g_indicator s_is_bounded_by_g1_and_g2) using within_range apply (simp add: cbox_def) by smt qed have zero_out_of_bound: "\ y. y \ cbox (g1 x) (g2 x) \ f (x,y) = 0" using f_is_g_indicator and s_is_bounded_by_g1_and_g2 by (auto simp add: cbox_def) have "(\y. f(x,y)) integrable_on cbox (g1 x) (g2 x)" proof - have "?f_one_D_spec integrable_on UNIV" using f_one_D_region and x_axis_gauge_integrable by metis then have "?f_one_D_spec integrable_on cbox(g1 x) (g2 x)" using integrable_on_subcbox by blast then show ?thesis using f_one_D_region by auto qed then have f_integrale_x: "((\y. f(x,y)) has_integral (integral (cbox (g1 x) (g2 x)) (\y. f(x,y)))) (cbox (g1 x) (g2 x))" using integrable_integral and within_range and x_axis_gauge_integrable by auto have "integral (cbox (g1 x) (g2 x)) (\y. f (x, y)) = integral UNIV (\y. f (x, y))" using has_integral_on_superset[OF f_integrale_x _ Set.subset_UNIV] zero_out_of_bound by (simp add: integral_unique) then have "((\y. f(x, y)) has_integral integral UNIV (\y. f (x, y))) (cbox (g1 x) (g2 x))" using f_integrale_x by simp then have "((\y. g(x, y)) has_integral integral UNIV (\y. f (x, y))) (cbox (g1 x)(g2 x))" using Henstock_Kurzweil_Integration.has_integral_restrict [OF subset_refl ] and f_one_D_region by (smt has_integral_eq) then show "integral (cbox (g1 x) (g2 x)) (\y. g (x, y)) = integral UNIV (\y. f (x, y))" by auto qed have case_x_not_in_range: "\ x. x \ cbox a b \ integral UNIV (\y. f(x,y)) = 0" proof fix x::'a have "x \ (cbox a b) \ (\y. f(x,y) = 0)" apply (simp add: s_is_bounded_by_g1_and_g2 f_is_g_indicator cbox_def) by auto then show "x \ cbox a b \ integral UNIV (\y. f (x, y)) = 0" by (simp) qed have RHS: "integral UNIV (\x. integral UNIV (\y. f(x,y))) = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" proof - let ?first_integral = "(\x. integral (cbox (g1 x) (g2 x)) (\y. g(x,y)))" let ?x_integral_cases = "(\x. if x \ cbox a b then ?first_integral x else 0)" have x_integral_cases_integral: "(\x. integral UNIV (\y. f(x,y))) = ?x_integral_cases" using case_x_in_range and case_x_not_in_range by auto have "((\x. integral UNIV (\y. f(x,y))) has_integral (integral UNIV f)) UNIV" using two_D_integral_to_one_D one_d_integral_integrable by auto then have "(?x_integral_cases has_integral (integral UNIV f)) UNIV" using x_integral_cases_integral by auto then have "(?first_integral has_integral (integral UNIV f)) (cbox a b)" using has_integral_restrict_UNIV[of "cbox a b" "?first_integral" "integral UNIV f"] by auto then show ?thesis using two_D_integral_to_one_D by (simp add: integral_unique) qed have f_integrable:"f integrable_on UNIV" using fun_lesbegue_integrable and integrable_on_lborel by auto then have LHS: "integral UNIV f = integral s g" apply (simp add: f_is_g_indicator) using integrable_restrict_UNIV integral_restrict_UNIV by auto then show ?thesis using RHS and two_D_integral_to_one_D by auto qed lemma gauge_integral_Fubini_curve_bounded_region_y: fixes f g :: "('a::euclidean_space * 'b::euclidean_space) \ 'c::euclidean_space" and g1 g2:: "'b \ 'a" and s:: "('a * 'b) set" assumes fun_lesbegue_integrable: "integrable lborel f" and y_axis_gauge_integrable: "\x. (\y. f(y, x)) integrable_on UNIV" and (*IS THIS redundant? NO IT IS NOT*) y_axis_integral_measurable: "(\x. integral UNIV (\y. f(y, x))) \ borel_measurable lborel" and f_is_g_indicator: "f = (\x. if x \ s then g x else 0)" and s_is_bounded_by_g1_and_g2: "s = {(y, x). (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i) \ (\i\Basis. (g1 x) \ i \ y \ i \ y \ i \ (g2 x) \ i)}" shows "integral s g = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" proof - have two_D_integral_to_one_D: "integral UNIV f = integral UNIV (\x. integral UNIV (\y. f(y, x)))" using gauge_integral_Fubini_universe_y and fun_lesbegue_integrable and y_axis_integral_measurable by auto have one_d_integral_integrable: "(\x. integral UNIV (\y. f(y, x))) integrable_on UNIV" using gauge_integral_Fubini_universe_y(2) and assms by blast have case_y_in_range: "\ x \ cbox a b. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)) = integral UNIV (\y. f(y, x))" proof fix x:: 'b assume within_range: "x \ (cbox a b)" let ?f_one_D_spec = "(\y. if y \ (cbox (g1 x) (g2 x)) then g(y, x) else 0)" have f_one_D_region: "(\y. f(y, x)) = (\y. if y \ cbox (g1 x) (g2 x) then g(y, x) else 0)" proof fix y::'a show "f (y, x) = (if y \ (cbox (g1 x) (g2 x)) then g (y, x) else 0)" apply (simp add: f_is_g_indicator s_is_bounded_by_g1_and_g2) using within_range apply (simp add: cbox_def) by smt qed have zero_out_of_bound: "\ y. y \ cbox (g1 x) (g2 x) \ f (y, x) = 0" using f_is_g_indicator and s_is_bounded_by_g1_and_g2 by (auto simp add: cbox_def) have "(\y. f(y, x)) integrable_on cbox (g1 x) (g2 x)" proof - have "?f_one_D_spec integrable_on UNIV" using f_one_D_region and y_axis_gauge_integrable by metis then have "?f_one_D_spec integrable_on cbox(g1 x) (g2 x)" using integrable_on_subcbox by blast then show ?thesis using f_one_D_region by auto qed then have f_integrale_y: "((\y. f(y, x)) has_integral (integral (cbox (g1 x) (g2 x)) (\y. f(y,x)))) (cbox (g1 x) (g2 x))" using integrable_integral and within_range and y_axis_gauge_integrable by auto have "integral (cbox (g1 x) (g2 x)) (\y. f (y, x)) = integral UNIV (\y. f (y, x))" using has_integral_on_superset[OF f_integrale_y _ Set.subset_UNIV] zero_out_of_bound by (simp add: integral_unique) then have "((\y. f(y, x)) has_integral integral UNIV (\y. f (y, x))) (cbox (g1 x) (g2 x))" using f_integrale_y by simp then have "((\y. g(y, x)) has_integral integral UNIV (\y. f (y, x))) (cbox (g1 x)(g2 x))" using Henstock_Kurzweil_Integration.has_integral_restrict [OF subset_refl ] and f_one_D_region by (smt has_integral_eq) then show "integral (cbox (g1 x) (g2 x)) (\y. g (y, x)) = integral UNIV (\y. f (y, x))" by auto qed have case_y_not_in_range: "\ x. x \ cbox a b \ integral UNIV (\y. f(y, x)) = 0" proof fix x::'b have "x \ (cbox a b) \ (\y. f(y, x) = 0)" apply (simp add: s_is_bounded_by_g1_and_g2 f_is_g_indicator cbox_def) by auto then show "x \ cbox a b \ integral UNIV (\y. f (y, x)) = 0" by (simp) qed have RHS: "integral UNIV (\x. integral UNIV (\y. f(y, x))) = integral (cbox a b) (\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" proof - let ?first_integral = "(\x. integral (cbox (g1 x) (g2 x)) (\y. g(y, x)))" let ?x_integral_cases = "(\x. if x \ cbox a b then ?first_integral x else 0)" have y_integral_cases_integral: "(\x. integral UNIV (\y. f(y, x))) = ?x_integral_cases" using case_y_in_range and case_y_not_in_range by auto have "((\x. integral UNIV (\y. f(y, x))) has_integral (integral UNIV f)) UNIV" using two_D_integral_to_one_D one_d_integral_integrable by auto then have "(?x_integral_cases has_integral (integral UNIV f)) UNIV" using y_integral_cases_integral by auto then have "(?first_integral has_integral (integral UNIV f)) (cbox a b)" using has_integral_restrict_UNIV[of "cbox a b" "?first_integral" "integral UNIV f"] by auto then show ?thesis using two_D_integral_to_one_D by (simp add: integral_unique) qed have f_integrable:"f integrable_on UNIV" using fun_lesbegue_integrable and integrable_on_lborel by auto then have LHS: "integral UNIV f = integral s g" apply (simp add: f_is_g_indicator) using integrable_restrict_UNIV integral_restrict_UNIV by auto then show ?thesis using RHS and two_D_integral_to_one_D by auto qed lemma gauge_integral_by_substitution: fixes f::"(real \ real)" and g::"(real \ real)" and g'::"real \ real" and a::"real" and b::"real" assumes a_le_b: "a \ b" and ga_le_gb: "g a \ g b" and g'_derivative: "\x \ {a..b}. (g has_vector_derivative (g' x)) (at x within {a..b})" and g'_continuous: "continuous_on {a..b} g'" and f_continuous: "continuous_on (g ` {a..b}) f" shows "integral {g a..g b} (f) = integral {a..b} (\x. f(g x) * (g' x))" proof - have "\x \ {a..b}. (g has_real_derivative (g' x)) (at x within {a..b})" using has_real_derivative_iff_has_vector_derivative[of "g"] and g'_derivative by auto then have 2: "interval_lebesgue_integral lborel (ereal (a)) (ereal (b)) (\x. g' x *\<^sub>R f (g x)) = interval_lebesgue_integral lborel (ereal (g a)) (ereal (g b)) f" using interval_integral_substitution_finite[of "a" "b" "g" "g'" "f"] and g'_continuous and a_le_b and f_continuous by auto have g_continuous: "continuous_on {a .. b} g" using Derivative.differentiable_imp_continuous_on apply (simp add: differentiable_on_def differentiable_def) by (metis continuous_on_vector_derivative g'_derivative) have "set_integrable lborel {a .. b} (\x. g' x *\<^sub>R f (g x))" proof - have "continuous_on {a .. b} (\x. g' x *\<^sub>R f (g x))" proof - have "continuous_on {a .. b} (\x. f (g x))" proof - show ?thesis using Topological_Spaces.continuous_on_compose f_continuous g_continuous by auto qed then show ?thesis using Limits.continuous_on_mult g'_continuous by auto qed then show ?thesis using borel_integrable_atLeastAtMost' by blast qed then have 0: "interval_lebesgue_integral lborel (ereal (a)) (ereal (b)) (\x. g' x *\<^sub>R f (g x)) = integral {a .. b} (\x. g' x *\<^sub>R f (g x))" using a_le_b and interval_integral_eq_integral by (metis (no_types)) have "set_integrable lborel {g a .. g b} f" proof - have "continuous_on {g a .. g b} f" proof - have "{g a .. g b} \ g ` {a .. b}" using g_continuous by (metis a_le_b atLeastAtMost_iff atLeastatMost_subset_iff continuous_image_closed_interval imageI order_refl) then show "continuous_on {g a .. g b} f" using f_continuous continuous_on_subset by blast qed then show ?thesis using borel_integrable_atLeastAtMost' by blast qed then have 1: "interval_lebesgue_integral lborel (ereal (g a)) (ereal (g b)) f = integral {g a .. g b} f" using ga_le_gb and interval_integral_eq_integral by (metis (no_types)) then show ?thesis using 0 and 1 and 2 by (metis (no_types, lifting) Henstock_Kurzweil_Integration.integral_cong mult.commute real_scaleR_def) qed lemma frontier_ic: assumes "a < (b::real)" shows "frontier {a<..b} = {a,b}" apply(simp add: frontier_def) using assms by auto lemma frontier_ci: assumes "a < (b::real)" shows "frontier {a<.. closed {a<..b}" using assms frontier_subset_eq frontier_ic greaterThanAtMost_iff by blast lemma closure_ic_union_ci: assumes "a < (b::real)" "b < c" shows "closure ({a.. {b<..c}) = {a .. c}" using frontier_ic[OF assms(1)] frontier_ci[OF assms(2)] closure_Un assms apply(simp add: frontier_def) by auto lemma interior_ic_ci_union: assumes "a < (b::real)" "b < c" shows "b \ (interior ({a.. {b<..c}))" proof- have "b \ ({a.. {b<..c})" by auto then show ?thesis using interior_subset by blast qed lemma frontier_ic_union_ci: assumes "a < (b::real)" "b < c" shows "b \ frontier ({a.. {b<..c})" using closure_ic_union_ci assms interior_ic_ci_union by(simp add: frontier_def) lemma ic_union_ci_not_closed: assumes "a < (b::real)" "b < c" shows "\ closed ({a.. {b<..c})" proof- have "b \ ({a.. {b<..c})" by auto then show ?thesis using assms frontier_subset_eq frontier_ic_union_ci[OF assms] by (auto simp only: subset_iff) qed lemma integrable_continuous_: fixes f :: "'b::euclidean_space \ 'a::banach" assumes "continuous_on (cbox a b) f" shows "f integrable_on cbox a b" by (simp add: assms integrable_continuous) lemma removing_singletons_from_div: assumes "\t\S. \c d::real. c < d \ {c..d} = t" "{x} \ \S = {a..b}" "a < x" "x < b" "finite S" shows "\t\S. x \ t" proof(rule ccontr) assume "\(\t\S. x \ t)" then have "\t\S. x \ t" by auto then have "x \ \S" by auto then have i: "\S = {a..b} - {x}" using assms (2) by auto have "x \ {a..b}" using assms by auto then have "{a .. b} - {x} = {a.. {x<..b}" by auto then have 0: "\S = {a.. {x<..b}" using i by auto have 1:"closed (\S)" apply(rule closed_Union) proof- show "finite S" using assms by auto show "\T\S. closed T" using assms by auto qed show False using 0 1 ic_union_ci_not_closed assms by auto qed lemma remove_singleton_from_division_of:(*By Manuel Eberl*) assumes "A division_of {a::real..b}" "a < b" assumes "x \ {a..b}" shows "\c d. c < d \ {c..d} \ A \ x \ {c..d}" proof - from assms have "x islimpt {a..b}" by (intro connected_imp_perfect) auto also have "{a..b} = {x. {x..x} \ A} \ ({a..b} - {x. {x..x} \ A})" using assms by auto also have "x islimpt \ \ x islimpt {a..b} - {x. {x..x} \ A}" proof (intro islimpt_Un_finite) have "{x. {x..x} \ A} \ Inf ` A" proof safe fix x assume "{x..x} \ A" thus "x \ Inf ` A" by (auto intro!: bexI[of _ "{x}"] simp: image_iff) qed moreover from assms have "finite A" by (auto simp: division_of_def) hence "finite (Inf ` A)" by auto ultimately show "finite {x. {x..x} \ A}" by (rule finite_subset) qed also have "{a..b} = \A" using assms by (auto simp: division_of_def) finally have "x islimpt \(A - range (\x. {x..x}))" by (rule islimpt_subset) auto moreover have "closed (\(A - range (\x. {x..x})))" using assms by (intro closed_Union) auto ultimately have "x \ (\(A - range (\x. {x..x})))" by (auto simp: closed_limpt) then obtain X where "x \ X" "X \ A" "X \ range (\x. {x..x})" by blast moreover from division_ofD(2)[OF assms(1) this(2)] division_ofD(3)[OF assms(1) this(2)] division_ofD(4)[OF assms(1) this(2)] obtain c d where "X = cbox c d" "X \ {a..b}" "X \ {}" by blast ultimately have "c \ d" by auto have "c \ d" proof assume "c = d" with \X = cbox c d\ have "X = {c..c}" by auto hence "X \ range (\x. {x..x})" by blast with \X \ range (\x. {x..x})\ show False by contradiction qed with \c \ d\ have "c < d" by simp with \X = cbox c d\ and \x \ X\ and \X \ A\ show ?thesis by auto qed lemma remove_singleton_from_tagged_division_of: assumes "A tagged_division_of {a::real..b}" "a < b" assumes "x \ {a..b}" shows "\k c d. c < d \ (k, {c..d}) \ A \ x \ {c..d}" using remove_singleton_from_division_of[OF division_of_tagged_division[OF assms(1)] assms(2)] (*sledgehammer*) using assms(3) by fastforce lemma tagged_div_wo_singlestons: assumes "p tagged_division_of {a::real..b}" "a < b" shows "(p - {xk. \x y. xk = (x,{y})}) tagged_division_of cbox a b" using remove_singleton_from_tagged_division_of[OF assms] assms apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def) apply blast apply blast apply blast by fastforce lemma tagged_div_wo_empty: assumes "p tagged_division_of {a::real..b}" "a < b" shows "(p - {xk. \x. xk = (x,{})}) tagged_division_of cbox a b" using remove_singleton_from_tagged_division_of[OF assms] assms apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def) apply blast apply blast apply blast by fastforce lemma fine_diff: assumes "\ fine p" shows "\ fine (p - s)" apply (auto simp add: fine_def) using assms by auto lemma tagged_div_tage_notin_set: assumes "finite (s::real set)" "p tagged_division_of {a..b}" "\ fine p" "(\(x, K)\p. \c d::real. c < d \ K = {c..d})" "gauge \" shows "\p' \'. p' tagged_division_of {a..b} \ \' fine p' \ (\(x, K)\p'. x \ s) \ gauge \'" proof- have "(\(x::real, K)\p. \x'. x' \ s \ x'\ interior K)" proof- {fix x::real fix K assume ass: "(x::real,K) \ p" have "(\(x, K)\p. infinite (interior K))" using assms(4) infinite_Ioo interior_atLeastAtMost_real by (smt split_beta) then have i: "infinite (interior K)" using ass by auto have "\x'. x' \ s \ x'\ interior K" using infinite_imp_nonempty[OF Diff_infinite_finite[OF assms(1) i]] by auto} then show ?thesis by auto qed then obtain f where f: "(\(x::real, K)\p. (f (x,K)) \ s \ (f (x,K)) \ interior K)" using choice_iff[where ?Q = "\(x,K) x'. (x::real, K)\p \ x' \ s \ x' \ interior K"] apply (auto simp add: case_prod_beta) by metis have f': "(\(x::real, K)\p. (f (x,K)) \ s \ (f (x,K)) \ K)" using f interior_subset by (auto simp add: case_prod_beta subset_iff) let ?p' = "{m. (\xK. m = ((f xK), snd xK) \ xK \ p)}" have 0: "(\(x, K)\?p'. x \ s)" using f by (auto simp add: case_prod_beta) have i: "finite {(f (a, b), b) |a b. (a, b) \ p}" proof- have a: "{(f (a, b), b) |a b. (a, b) \ p} = (%(a,b). (f(a,b),b)) ` p" by auto have b: "finite p" using assms(2) by auto show ?thesis using a b by auto qed have 1: "?p' tagged_division_of {a..b}" using assms(2) f' apply(auto simp add: tagged_division_of_def tagged_partial_division_of_def case_prod_beta) apply(metis i) apply blast apply blast by fastforce (*f is injective becuase interiors of different K's are disjoint and f is in interior*) have f_inj: "inj_on f p" apply(simp add: inj_on_def) proof- {fix x y assume "x \ p" "y \ p" "f x = f y" then have "x = y" using f tagged_division_ofD(5)[OF assms(2)] (*sledgehammer*) by (smt case_prodE disjoint_insert(2) mk_disjoint_insert)}note * = this show "\x\p. \y\p. f x = f y \ x = y" using * by auto qed let ?\' = "\x. if (\xK \ p. f xK = x) then (\ o fst o the_inv_into p f) x else \ x" have 2: "?\' fine ?p'" using assms(3) apply(auto simp add: fine_def case_prod_beta the_inv_into_f_f[OF f_inj]) by force have 3: "gauge ?\'" using assms(5) assms(3) f' apply(auto simp add: fine_def gauge_def case_prod_beta the_inv_into_f_f[OF f_inj]) by force have "?p' tagged_division_of {a..b} \ ?\' fine ?p' \ (\(x, K)\?p'. x \ s) \ gauge ?\'" using 0 1 2 3 by auto then show ?thesis by smt qed lemma has_integral_bound_spike_finite: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "0 \ B" and "finite S" and f: "(f has_integral i) (cbox a b)" and leB: "\x. x \ cbox a b - S \ norm (f x) \ B" shows "norm i \ B * content (cbox a b)" proof - define g where "g \ (\x. if x \ S then 0 else f x)" then have "\x. x \ cbox a b - S \ norm (g x) \ B" using leB by simp moreover have "(g has_integral i) (cbox a b)" using has_integral_spike_finite [OF \finite S\ _ f] by (simp add: g_def) ultimately show ?thesis by (simp add: \0 \ B\ g_def has_integral_bound) qed lemma has_integral_bound_: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and "0 \ B" and f: "(f has_integral i) (cbox a b)" and "finite s" and "\x\(cbox a b)-s. norm (f x) \ B" shows "norm i \ B * content (cbox a b)" using has_integral_bound_spike_finite assms by blast corollary has_integral_bound_real': fixes f :: "real \ 'b::real_normed_vector" assumes "0 \ B" and f: "(f has_integral i) (cbox a b)" and "finite s" and "\x\(cbox a b)-s. norm (f x) \ B" shows "norm i \ B * content {a..b}" (*sledgehammer*) by (metis assms(1) assms(3) assms(4) box_real(2) f has_integral_bound_spike_finite) lemma integral_has_vector_derivative_continuous_at': fixes f :: "real \ 'a::banach" assumes "finite s" and f: "f integrable_on {a..b}" and x: "x \ {a..b} - s" and fx: "continuous (at x within ({a..b} - s)) f" shows "((\u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - s))" proof - let ?I = "\a b. integral {a..b} f" { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {a..b} - s; \x' - x\ < d\ \ norm(f x' - f x) \ e" using \e>0\ fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" if y: "y \ {a..b} - s" and yx: "\y - x\ < d" for y proof (cases "y < x") case False have "f integrable_on {a..y}" using f y by (simp add: integrable_subinterval_real) then have Idiff: "?I a y - ?I a x = ?I x y" using False x by (simp add: algebra_simps Henstock_Kurzweil_Integration.integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" x y] False apply simp done show ?thesis using False apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real'[where f="(\u. f u - f x)"]) using yx False d x y \e>0\ apply (auto simp add: Idiff fux_int) proof- let ?M48= "mset_set s" show "\z. y - x < d \ (\x'. a \ x' \ x' \ b \ x' \ s \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ z \# ?M48 \ a \ x \ x \ s \ y \ b \ y \ s \ x \ z \ z \ y \ norm (f z - f x) \ e" using assms by auto qed next case True have "f integrable_on {a..x}" using f x by (simp add: integrable_subinterval_real) then have Idiff: "?I a x - ?I a y = ?I y x" using True x y by (simp add: algebra_simps Henstock_Kurzweil_Integration.integral_combine) have fux_int: "((\u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}" apply (rule has_integral_diff) using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) using has_integral_const_real [of "f x" y x] True apply simp done have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \ e * \y - x\" using True apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) apply (rule has_integral_bound_real'[where f="(\u. f u - f x)"]) using yx True d x y \e>0\ apply (auto simp add: Idiff fux_int) proof- let ?M44= "mset_set s" show " \xa. x - y < d \ y < x \ (\x'. a \ x' \ x' \ b \ x' \ s \ \x' - x\ < d \ norm (f x' - f x) \ e) \ 0 < e \ xa \# ?M44 \ x \ b \ x \ s \ a \ y \ y \ s \ y \ xa \ xa \ x \ norm (f xa - f x) \ e" using assms by auto qed then show ?thesis by (simp add: algebra_simps norm_minus_commute) qed then have "\d>0. \y\{a..b} - s. \y - x\ < d \ norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \ e * \y - x\" using \d>0\ by blast } then show ?thesis by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) qed lemma integral_has_vector_derivative': fixes f :: "real \ 'a::banach" assumes "finite s" "f integrable_on {a..b}" "x \ {a..b} - s" "continuous (at x within {a..b} - s) f" shows "((\u. integral {a .. u} f) has_vector_derivative f(x)) (at x within {a .. b} - s)" apply (rule integral_has_vector_derivative_continuous_at') using assms apply (auto simp: continuous_on_eq_continuous_within) done lemma fundamental_theorem_of_calculus_interior_stronger: fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms proof (induction arbitrary: a b) case empty then show ?case using fundamental_theorem_of_calculus_interior by force next case (insert x S) show ?case proof (cases "x \ {a<..a < x\ \x < b\ insert.prems continuous_on_subset by force+ moreover have "(f' has_integral f b - f x) {x..b}" apply (rule insert) using \x < b\ \a < x\ insert.prems continuous_on_subset by force+ ultimately show ?thesis by (meson finite_insert fundamental_theorem_of_calculus_interior_strong insert.hyps(1) insert.prems(1) insert.prems(2) insert.prems(3)) qed qed lemma at_within_closed_interval_finite: fixes x::real assumes "a < x" "x < b" "x \ S" "finite S" shows "(at x within {a..b} - S) = at x" proof - have "interior ({a..b} - S) = {a<..finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma at_within_cbox_finite: assumes "x \ box a b" "x \ S" "finite S" shows "(at x within cbox a b - S) = at x" proof - have "interior (cbox a b - S) = box a b - S" using \finite S\ by (simp add: interior_diff finite_imp_closed) then show ?thesis using at_within_interior assms by fastforce qed lemma fundamental_theorem_of_calculus_interior_stronger': fixes f :: "real \ 'a::banach" assumes "finite S" and "a \ b" "\x. x \ {a <..< b} - S \ (f has_vector_derivative f'(x)) (at x within {a..b} - S)" and "continuous_on {a .. b} f" shows "(f' has_integral (f b - f a)) {a .. b}" using assms fundamental_theorem_of_calculus_interior_strong at_within_cbox_finite (*sledgehammer*) by (metis DiffD1 DiffD2 interior_atLeastAtMost_real interior_cbox interval_cbox) lemma has_integral_substitution_general_: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - (g ` s)) f" (*and f [continuous_intros]: "continuous_on {c..d} f"*) and g : "continuous_on {a..b} g" "inj_on g ({a..b} \ s)" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof - let ?F = "\x. integral {c..g x} f" have cont_int: "continuous_on {a..b} ?F" by (rule continuous_on_compose2[OF _ g(1) subset] indefinite_integral_continuous_1 f)+ have deriv: "\x. x \ {a..b} - s \ (((\x. integral {c..x} f) \ g) has_vector_derivative g' x *\<^sub>R f (g x)) (at x within ({a..b} - s))" apply (rule has_vector_derivative_eq_rhs) apply (rule vector_diff_chain_within) apply (subst has_real_derivative_iff_has_vector_derivative [symmetric]) proof- fix x::real assume ass: "x \ {a..b} - s" let ?f'3 = "g' x" have i:"{a..b} - s \ {a..b}" by auto have ii: " (g has_vector_derivative g' x) (at x within {a..b})" using deriv[OF ass] by (simp only: has_real_derivative_iff_has_vector_derivative) show "(g has_real_derivative ?f'3) (at x within {a..b} - s)" using has_vector_derivative_within_subset[OF ii i] by (simp only: has_real_derivative_iff_has_vector_derivative) next let ?g'3 = "f o g" show "\x. x \ {a..b} - s \ ((\x. integral {c..x} f) has_vector_derivative ?g'3 x) (at (g x) within g ` ({a..b} - s))" proof- fix x::real assume ass: "x \ {a..b} - s" have "finite (g ` s)" using s by auto then have i: "((\x. integral {c..x} f) has_vector_derivative f(g x)) (at (g x) within ({c..d} - g ` s))" apply (rule integral_has_vector_derivative') proof- show " f integrable_on {c..d}" using f by auto show "g x \ {c..d} - g ` s" using ass subset (*sledgehammer*) by (smt Diff_iff Un_upper1 Un_upper2 g(2) imageE image_subset_iff inj_onD subsetCE) show "continuous (at (g x) within {c..d} - g ` s) f" (*sledgehammer*) using \g x \ {c..d} - g ` s\ continuous_on_eq_continuous_within f(2) by blast qed have ii: "g ` ({a..b} - s) \ ({c..d} - g ` s)" using subset g(2) (*sledgehammer*) by (smt Diff_subset Un_Diff Un_commute Un_upper2 inj_on_image_set_diff subset_trans sup.order_iff) then show "((\x. integral {c..x} f) has_vector_derivative ?g'3 x) (at (g x) within g ` ({a..b} - s))" (*sledgehammer*) by (smt Diff_subset has_vector_derivative_weaken Un_upper1 Un_upper2 \finite (g ` s)\ ass comp_def continuous_on_eq_continuous_within f(1) f(2) g(2) image_diff_subset image_subset_iff inj_on_image_set_diff integral_has_vector_derivative_continuous_at' subset_trans) qed show "\x. x \ {a..b} - s \ g' x *\<^sub>R ?g'3 x = g' x *\<^sub>R f (g x)" by auto qed have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) (at x within {a..b} - s)" if "x \ {a<..x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}" using cont_int using fundamental_theorem_of_calculus_interior_stronger'[OF s le deriv] by blast also from subset have "g x \ {c..d}" if "x \ {a..b}" for x using that by blast from this[of a] this[of b] le have cd: "c \ g a" "g b \ d" "c \ g b" "g a \ d" by auto have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f" proof cases assume "g a \ g b" note le = le this from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f" by (meson Henstock_Kurzweil_Integration.integral_combine atLeastatMost_subset_iff f(1) integrable_on_subinterval le(2) order_refl) with le show ?thesis by (cases "g a = g b") (simp_all add: algebra_simps) next assume less: "\g a \ g b" then have "g a \ g b" by simp note le = le this from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f" by (meson Henstock_Kurzweil_Integration.integral_combine atLeastatMost_subset_iff f(1) integrable_on_subinterval le(2) order_refl) with less show ?thesis by (simp_all add: algebra_simps) qed finally show ?thesis . qed lemma has_integral_substitution_general__: fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and s_subset: "s \ {a..b}" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - (g ` s)) f" (*and f [continuous_intros]: "continuous_on {c..d} f"*) and g : "continuous_on {a..b} g" "inj_on g {a..b}" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" using s_subset has_integral_substitution_general_[OF s le subset f g(1) _ deriv] by (simp add: g(2) sup_absorb1) lemma has_integral_substitution_general_': fixes f :: "real \ 'a::euclidean_space" and g :: "real \ real" assumes s: "finite s" and le: "a \ b" and s': "finite s'" and subset: "g ` {a..b} \ {c..d}" and f: "f integrable_on {c..d}" "continuous_on ({c..d} - s') f" and g : "continuous_on {a..b} g" "\x\s'. finite (g -` {x})" "surj_on s' g" "inj_on g ({a..b} \ ((s \ g -` s')))" and deriv [derivative_intros]: "\x. x \ {a..b} - s \ (g has_field_derivative g' x) (at x within {a..b})" shows "((\x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}" proof- have a: "g -` s' = \{t. \x. t = g -` {x} \ x \ s'}" using s s' by blast have "finite (\{t. \x. t = g -` {x} \ x \ s'})" using s' by (metis (no_types, lifting) \g -` s' = \{g -` {x} |x. x \ s'}\ finite_UN_I g(2) vimage_eq_UN) then have 0: "finite (s \ (g -` s'))" using a s by simp have 1: "continuous_on ({c..d} - g ` (s \ g -` s')) f" using f(2) surj_on_image_vimage_eq by (metis Diff_mono Un_upper2 continuous_on_subset equalityE g(3) image_Un) have 2: " (\x. x \ {a..b} - (s \ g -` s') \ (g has_real_derivative g' x) (at x within {a..b}))" using deriv by auto show ?thesis using has_integral_substitution_general_[OF 0 assms(2) subset f(1) 1 g(1) g(4) 2] by auto qed end diff --git a/thys/Martingales/Bochner_Integration_Supplement.thy b/thys/Martingales/Bochner_Integration_Supplement.thy --- a/thys/Martingales/Bochner_Integration_Supplement.thy +++ b/thys/Martingales/Bochner_Integration_Supplement.thy @@ -1,722 +1,722 @@ (* Author: Ata Keskin, TU München *) theory Bochner_Integration_Supplement imports "HOL-Analysis.Bochner_Integration" "HOL-Analysis.Set_Integral" Elementary_Metric_Spaces_Supplement begin section \Supplementary Lemmas for Bochner Integration\ subsection \Integrable Simple Functions\ text \We restate some basic results concerning Bochner-integrable functions.\ lemma integrable_implies_simple_function_sequence: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "integrable M f" obtains s where "\i. simple_function M (s i)" and "\i. emeasure M {y \ space M. s i y \ 0} \ \" and "\x. x \ space M \ (\i. s i x) \ f x" and "\x i. x \ space M \ norm (s i x) \ 2 * norm (f x)" proof- have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" using assms unfolding integrable_iff_bounded by auto obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" using borel_measurable_implies_sequence_metric[OF f(1)] unfolding norm_conv_dist by metis { fix i have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ennreal (2 * norm (f x)) \M)" using s by (intro nn_integral_mono, auto) also have "\ < \" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) finally have sbi: "Bochner_Integration.simple_bochner_integrable M (s i)" using s by (intro simple_bochner_integrableI_bounded) auto hence "emeasure M {y \ space M. s i y \ 0} \ \" by (auto intro: integrableI_simple_bochner_integrable simple_bochner_integrable.cases) } thus ?thesis using that s by blast qed text \Simple functions can be represented by sums of indicator functions.\ lemma simple_function_indicator_representation: fixes f ::"'a \ 'b :: {second_countable_topology, banach}" assumes "simple_function M f" "x \ space M" shows "f x = (\y \ f ` space M. indicator (f -` {y} \ space M) x *\<^sub>R y)" (is "?l = ?r") proof - have "?r = (\y \ f ` space M. (if y = f x then indicator (f -` {y} \ space M) x *\<^sub>R y else 0))" by (auto intro!: sum.cong) also have "... = indicator (f -` {f x} \ space M) x *\<^sub>R f x" using assms by (auto dest: simple_functionD) also have "... = f x" using assms by (auto simp: indicator_def) finally show ?thesis by auto qed lemma simple_function_indicator_representation_AE: fixes f ::"'a \ 'b :: {second_countable_topology, banach}" assumes "simple_function M f" shows "AE x in M. f x = (\y \ f ` space M. indicator (f -` {y} \ space M) x *\<^sub>R y)" by (metis (mono_tags, lifting) AE_I2 simple_function_indicator_representation assms) lemmas simple_function_scaleR[intro] = simple_function_compose2[where h="(*\<^sub>R)"] lemmas integrable_simple_function = simple_bochner_integrable.intros[THEN has_bochner_integral_simple_bochner_integrable, THEN integrable.intros] text \Induction rule for simple integrable functions.\ lemma\<^marker>\tag important\ integrable_simple_function_induct[consumes 2, case_names cong indicator add, induct set: simple_function]: fixes f :: "'a \ 'b :: {second_countable_topology, banach}" assumes f: "simple_function M f" "emeasure M {y \ space M. f y \ 0} \ \" assumes cong: "\f g. simple_function M f \ emeasure M {y \ space M. f y \ 0} \ \ \ simple_function M g \ emeasure M {y \ space M. g y \ 0} \ \ \ (\x. x \ space M \ f x = g x) \ P f \ P g" assumes indicator: "\A y. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R y)" assumes add: "\f g. simple_function M f \ emeasure M {y \ space M. f y \ 0} \ \ \ simple_function M g \ emeasure M {y \ space M. g y \ 0} \ \ \ (\z. z \ space M \ norm (f z + g z) = norm (f z) + norm (g z)) \ P f \ P g \ P (\x. f x + g x)" shows "P f" proof- let ?f = "\x. (\y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" have f_ae_eq: "f x = ?f x" if "x \ space M" for x using simple_function_indicator_representation[OF f(1) that] . moreover have "emeasure M {y \ space M. ?f y \ 0} \ \" by (metis (no_types, lifting) Collect_cong calculation f(2)) moreover have "P (\x. \y\S. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" "simple_function M (\x. \y\S. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" "emeasure M {y \ space M. (\x\S. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ \" if "S \ f ` space M" for S using simple_functionD(1)[OF assms(1), THEN rev_finite_subset, OF that] that proof (induction rule: finite_induct) case empty { case 1 then show ?case using indicator[of "{}"] by force next case 2 then show ?case by force next case 3 then show ?case by force } next case (insert x F) have "(f -` {x} \ space M) \ {y \ space M. f y \ 0}" if "x \ 0" using that by blast moreover have "{y \ space M. f y \ 0} = space M - (f -` {0} \ space M)" by blast moreover have "space M - (f -` {0} \ space M) \ sets M" using simple_functionD(2)[OF f(1)] by blast ultimately have fin_0: "emeasure M (f -` {x} \ space M) < \" if "x \ 0" using that by (metis emeasure_mono f(2) infinity_ennreal_def top.not_eq_extremum top_unique) hence fin_1: "emeasure M {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0} \ \" if "x \ 0" by (metis (mono_tags, lifting) emeasure_mono f(1) indicator_simps(2) linorder_not_less mem_Collect_eq scaleR_eq_0_iff simple_functionD(2) subsetI that) have *: "(\y\insert x F. indicat_real (f -` {y} \ space M) xa *\<^sub>R y) = (\y\F. indicat_real (f -` {y} \ space M) xa *\<^sub>R y) + indicat_real (f -` {x} \ space M) xa *\<^sub>R x" for xa by (metis (no_types, lifting) Diff_empty Diff_insert0 add.commute insert.hyps(1) insert.hyps(2) sum.insert_remove) have **: "{y \ space M. (\x\insert x F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0}" unfolding * by fastforce { case 1 hence x: "x \ f ` space M" and F: "F \ f ` space M" by auto show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert(3)[OF F] by simp next case False have norm_argument: "norm ((\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) + indicat_real (f -` {x} \ space M) z *\<^sub>R x) = norm (\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) + norm (indicat_real (f -` {x} \ space M) z *\<^sub>R x)" if z: "z \ space M" for z proof (cases "f z = x") case True have "indicat_real (f -` {y} \ space M) z *\<^sub>R y = 0" if "y \ F" for y using True insert(2) z that 1 unfolding indicator_def by force hence "(\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) = 0" by (meson sum.neutral) then show ?thesis by force next case False then show ?thesis by force qed show ?thesis using False simple_functionD(2)[OF f(1)] insert(3,5)[OF F] simple_function_scaleR fin_0 fin_1 by (subst *, subst add, subst simple_function_sum) (blast intro: norm_argument indicator)+ qed next case 2 hence x: "x \ f ` space M" and F: "F \ f ` space M" by auto show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert(4)[OF F] by simp next case False then show ?thesis unfolding * using insert(4)[OF F] simple_functionD(2)[OF f(1)] by fast qed next case 3 hence x: "x \ f ` space M" and F: "F \ f ` space M" by auto show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert(5)[OF F] by simp next case False have "emeasure M {y \ space M. (\x\insert x F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ emeasure M ({y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0})" using ** simple_functionD(2)[OF insert(4)[OF F]] simple_functionD(2)[OF f(1)] by (intro emeasure_mono, force+) also have "... \ emeasure M {y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} + emeasure M {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0}" using simple_functionD(2)[OF insert(4)[OF F]] simple_functionD(2)[OF f(1)] by (intro emeasure_subadditive, force+) also have "... < \" using insert(5)[OF F] fin_1[OF False] by (simp add: less_top) finally show ?thesis by simp qed } qed moreover have "simple_function M (\x. \y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" using calculation by blast moreover have "P (\x. \y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" using calculation by blast ultimately show ?thesis by (intro cong[OF _ _ f(1,2)], blast, presburger+) qed text \Induction rule for non-negative simple integrable functions\ lemma\<^marker>\tag important\ integrable_simple_function_induct_nn[consumes 3, case_names cong indicator add, induct set: simple_function]: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes f: "simple_function M f" "emeasure M {y \ space M. f y \ 0} \ \" "\x. x \ space M \ f x \ 0" assumes cong: "\f g. simple_function M f \ emeasure M {y \ space M. f y \ 0} \ \ \ (\x. x \ space M \ f x \ 0) \ simple_function M g \ emeasure M {y \ space M. g y \ 0} \ \ \ (\x. x \ space M \ g x \ 0) \ (\x. x \ space M \ f x = g x) \ P f \ P g" assumes indicator: "\A y. y \ 0 \ A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R y)" assumes add: "\f g. (\x. x \ space M \ f x \ 0) \ simple_function M f \ emeasure M {y \ space M. f y \ 0} \ \ \ (\x. x \ space M \ g x \ 0) \ simple_function M g \ emeasure M {y \ space M. g y \ 0} \ \ \ (\z. z \ space M \ norm (f z + g z) = norm (f z) + norm (g z)) \ P f \ P g \ P (\x. f x + g x)" shows "P f" proof- let ?f = "\x. (\y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" have f_ae_eq: "f x = ?f x" if "x \ space M" for x using simple_function_indicator_representation[OF f(1) that] . moreover have "emeasure M {y \ space M. ?f y \ 0} \ \" by (metis (no_types, lifting) Collect_cong calculation f(2)) moreover have "P (\x. \y\S. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" "simple_function M (\x. \y\S. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" "emeasure M {y \ space M. (\x\S. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ \" "\x. x \ space M \ 0 \ (\y\S. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" if "S \ f ` space M" for S using simple_functionD(1)[OF assms(1), THEN rev_finite_subset, OF that] that proof (induction rule: finite_subset_induct') case empty { case 1 then show ?case using indicator[of 0 "{}"] by force next case 2 then show ?case by force next case 3 then show ?case by force next case 4 then show ?case by force } next case (insert x F) have "(f -` {x} \ space M) \ {y \ space M. f y \ 0}" if "x \ 0" using that by blast moreover have "{y \ space M. f y \ 0} = space M - (f -` {0} \ space M)" by blast moreover have "space M - (f -` {0} \ space M) \ sets M" using simple_functionD(2)[OF f(1)] by blast ultimately have fin_0: "emeasure M (f -` {x} \ space M) < \" if "x \ 0" using that by (metis emeasure_mono f(2) infinity_ennreal_def top.not_eq_extremum top_unique) hence fin_1: "emeasure M {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0} \ \" if "x \ 0" by (metis (mono_tags, lifting) emeasure_mono f(1) indicator_simps(2) linorder_not_less mem_Collect_eq scaleR_eq_0_iff simple_functionD(2) subsetI that) have nonneg_x: "x \ 0" using insert f by blast have *: "(\y\insert x F. indicat_real (f -` {y} \ space M) xa *\<^sub>R y) = (\y\F. indicat_real (f -` {y} \ space M) xa *\<^sub>R y) + indicat_real (f -` {x} \ space M) xa *\<^sub>R x" for xa by (metis (no_types, lifting) add.commute insert.hyps(1) insert.hyps(4) sum.insert) have **: "{y \ space M. (\x\insert x F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0}" unfolding * by fastforce { case 1 show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert by simp next case False have norm_argument: "norm ((\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) + indicat_real (f -` {x} \ space M) z *\<^sub>R x) = norm (\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) + norm (indicat_real (f -` {x} \ space M) z *\<^sub>R x)" if z: "z \ space M" for z proof (cases "f z = x") case True have "indicat_real (f -` {y} \ space M) z *\<^sub>R y = 0" if "y \ F" for y using True insert z that 1 unfolding indicator_def by force hence "(\y\F. indicat_real (f -` {y} \ space M) z *\<^sub>R y) = 0" by (meson sum.neutral) thus ?thesis by force qed (force) show ?thesis using False fin_0 fin_1 f norm_argument by (subst *, subst add, presburger add: insert, intro insert, intro insert, fastforce simp add: indicator_def intro!: insert(2) f(3), auto intro!: indicator insert nonneg_x) qed next case 2 show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert by simp next case False then show ?thesis unfolding * using insert simple_functionD(2)[OF f(1)] by fast qed next case 3 show ?case proof (cases "x = 0") case True then show ?thesis unfolding * using insert by simp next case False have "emeasure M {y \ space M. (\x\insert x F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ emeasure M ({y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} \ {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0})" using ** simple_functionD(2)[OF insert(6)] simple_functionD(2)[OF f(1)] insert.IH(2) by (intro emeasure_mono, blast, simp) also have "... \ emeasure M {y \ space M. (\x\F. indicat_real (f -` {x} \ space M) y *\<^sub>R x) \ 0} + emeasure M {y \ space M. indicat_real (f -` {x} \ space M) y *\<^sub>R x \ 0}" using simple_functionD(2)[OF insert(6)] simple_functionD(2)[OF f(1)] by (intro emeasure_subadditive, force+) also have "... < \" using insert(7) fin_1[OF False] by (simp add: less_top) finally show ?thesis by simp qed next case (4 \) thus ?case using insert nonneg_x f(3) by (auto simp add: scaleR_nonneg_nonneg intro: sum_nonneg) } qed moreover have "simple_function M (\x. \y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" using calculation by blast moreover have "P (\x. \y\f ` space M. indicat_real (f -` {y} \ space M) x *\<^sub>R y)" using calculation by blast moreover have "\x. x \ space M \ 0 \ f x" using f(3) by simp ultimately show ?thesis by (intro cong[OF _ _ _ f(1,2)], blast, blast, fast) presburger+ qed lemma finite_nn_integral_imp_ae_finite: fixes f :: "'a \ ennreal" assumes "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" shows "AE x in M. f x < \" proof (rule ccontr, goal_cases) case 1 let ?A = "space M \ {x. f x = \}" have *: "emeasure M ?A > 0" using 1 assms(1) by (metis (mono_tags, lifting) assms(2) eventually_mono infinity_ennreal_def nn_integral_noteq_infinite top.not_eq_extremum) have "(\\<^sup>+x. f x * indicator ?A x \M) = (\\<^sup>+x. \ * indicator ?A x \M)" by (metis (mono_tags, lifting) indicator_inter_arith indicator_simps(2) mem_Collect_eq mult_eq_0_iff) also have "... = \ * emeasure M ?A" using assms(1) by (intro nn_integral_cmult_indicator, simp) also have "... = \" using * by fastforce finally have "(\\<^sup>+x. f x * indicator ?A x \M) = \" . moreover have "(\\<^sup>+x. f x * indicator ?A x \M) \ (\\<^sup>+x. f x \M)" by (intro nn_integral_mono, simp add: indicator_def) ultimately show ?case using assms(2) by simp qed text \Convergence in L1-Norm implies existence of a subsequence which convergences almost everywhere. This lemma is easier to use than the existing one in \<^theory>\HOL-Analysis.Bochner_Integration\\ lemma cauchy_L1_AE_cauchy_subseq: fixes s :: "nat \ 'a \ 'b::{banach, second_countable_topology}" assumes [measurable]: "\n. integrable M (s n)" and "\e. e > 0 \ \N. \i\N. \j\N. LINT x|M. norm (s i x - s j x) < e" obtains r where "strict_mono r" "AE x in M. Cauchy (\i. s (r i) x)" proof- have "\r. \n. (\i\r n. \j\ r n. LINT x|M. norm (s i x - s j x) < (1 / 2) ^ n) \ (r (Suc n) > r n)" proof (intro dependent_nat_choice, goal_cases) case 1 then show ?case using assms(2) by presburger next case (2 x n) obtain N where *: "LINT x|M. norm (s i x - s j x) < (1 / 2) ^ Suc n" if "i \ N" "j \ N" for i j using assms(2)[of "(1 / 2) ^ Suc n"] by auto { fix i j assume "i \ max N (Suc x)" "j \ max N (Suc x)" hence "integral\<^sup>L M (\x. norm (s i x - s j x)) < (1 / 2) ^ Suc n" using * by fastforce } then show ?case by fastforce qed then obtain r where strict_mono: "strict_mono r" and "\i\r n. \j\ r n. LINT x|M. norm (s i x - s j x) < (1 / 2) ^ n" for n using strict_mono_Suc_iff by blast hence r_is: "LINT x|M. norm (s (r (Suc n)) x - s (r n) x) < (1 / 2) ^ n" for n by (simp add: strict_mono_leD) define g where "g = (\n x. (\i\n. ennreal (norm (s (r (Suc i)) x - s (r i) x))))" define g' where "g' = (\x. \i. ennreal (norm (s (r (Suc i)) x - s (r i) x)))" have integrable_g: "(\\<^sup>+ x. g n x \M) < 2" for n proof - have "(\\<^sup>+ x. g n x \M) = (\\<^sup>+ x. (\i\n. ennreal (norm (s (r (Suc i)) x - s (r i) x))) \M)" using g_def by simp also have "... = (\i\n. (\\<^sup>+ x. ennreal (norm (s (r (Suc i)) x - s (r i) x)) \M))" by (intro nn_integral_sum, simp) also have "... = (\i\n. LINT x|M. norm (s (r (Suc i)) x - s (r i) x))" unfolding dist_norm using assms(1) by (subst nn_integral_eq_integral[OF integrable_norm], auto) also have "... < ennreal (\i\n. (1 / 2) ^ i)" by (intro ennreal_lessI[OF sum_pos sum_strict_mono[OF finite_atMost _ r_is]], auto) also have "... \ ennreal 2" unfolding sum_gp0[of "1 / 2" n] by (intro ennreal_leI, auto) finally show "(\\<^sup>+ x. g n x \M) < 2" by simp qed have integrable_g': "(\\<^sup>+ x. g' x \M) \ 2" proof - have "incseq (\n. g n x)" for x by (intro incseq_SucI, auto simp add: g_def ennreal_leI) hence "convergent (\n. g n x)" for x unfolding convergent_def using LIMSEQ_incseq_SUP by fast hence "(\n. g n x) \ g' x" for x unfolding g_def g'_def by (intro summable_iff_convergent'[THEN iffD2, THEN summable_LIMSEQ'], blast) hence "(\\<^sup>+ x. g' x \M) = (\\<^sup>+ x. liminf (\n. g n x) \M)" by (metis lim_imp_Liminf trivial_limit_sequentially) also have "... \ liminf (\n. \\<^sup>+ x. g n x \M)" by (intro nn_integral_liminf, simp add: g_def) also have "... \ liminf (\n. 2)" using integrable_g by (intro Liminf_mono) (simp add: order_le_less) also have "... = 2" using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast finally show ?thesis . qed hence "AE x in M. g' x < \" by (intro finite_nn_integral_imp_ae_finite) (auto simp add: order_le_less_trans g'_def) moreover have "summable (\i. norm (s (r (Suc i)) x - s (r i) x))" if "g' x \ \" for x using that unfolding g'_def by (intro summable_suminf_not_top) fastforce+ ultimately have ae_summable: "AE x in M. summable (\i. s (r (Suc i)) x - s (r i) x)" using summable_norm_cancel unfolding dist_norm by force { fix x assume "summable (\i. s (r (Suc i)) x - s (r i) x)" hence "(\n. \i (\i. s (r (Suc i)) x - s (r i) x)" using summable_LIMSEQ by blast moreover have "(\n. (\in. s (r n) x - s (r 0) x)" using sum_lessThan_telescope by fastforce ultimately have "(\n. s (r n) x - s (r 0) x) \ (\i. s (r (Suc i)) x - s (r i) x)" by argo hence "(\n. s (r n) x - s (r 0) x + s (r 0) x) \ (\i. s (r (Suc i)) x - s (r i) x) + s (r 0) x" by (intro isCont_tendsto_compose[of _ "\z. z + s (r 0) x"], auto) hence "Cauchy (\n. s (r n) x)" by (simp add: LIMSEQ_imp_Cauchy) } hence "AE x in M. Cauchy (\i. s (r i) x)" using ae_summable by fast thus ?thesis by (rule that[OF strict_mono(1)]) qed subsection \Totally Ordered Banach Spaces\ lemma integrable_max[simp, intro]: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology}" assumes fg[measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. max (f x) (g x))" proof (rule Bochner_Integration.integrable_bound) { fix x y :: 'b have "norm (max x y) \ max (norm x) (norm y)" by linarith also have "... \ norm x + norm y" by simp finally have "norm (max x y) \ norm (norm x + norm y)" by simp } thus "AE x in M. norm (max (f x) (g x)) \ norm (norm (f x) + norm (g x))" by simp qed (auto intro: Bochner_Integration.integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]) lemma integrable_min[simp, intro]: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology}" assumes [measurable]: "integrable M f" "integrable M g" shows "integrable M (\x. min (f x) (g x))" proof - have "norm (min (f x) (g x)) \ norm (f x) \ norm (min (f x) (g x)) \ norm (g x)" for x by linarith thus ?thesis by (intro integrable_bound[OF integrable_max[OF integrable_norm(1,1), OF assms]], auto) qed text \Restatement of \integral_nonneg_AE\ for functions taking values in a Banach space.\ lemma integral_nonneg_AE_banach: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes [measurable]: "f \ borel_measurable M" and nonneg: "AE x in M. 0 \ f x" shows "0 \ integral\<^sup>L M f" proof cases assume integrable: "integrable M f" hence max: "(\x. max 0 (f x)) \ borel_measurable M" "\x. 0 \ max 0 (f x)" "integrable M (\x. max 0 (f x))" by auto hence "0 \ integral\<^sup>L M (\x. max 0 (f x))" proof - obtain s where *: "\i. simple_function M (s i)" "\i. emeasure M {y \ space M. s i y \ 0} \ \" "\x. x \ space M \ (\i. s i x) \ f x" "\x i. x \ space M \ norm (s i x) \ 2 * norm (f x)" using integrable_implies_simple_function_sequence[OF integrable] by blast have simple: "\i. simple_function M (\x. max 0 (s i x))" using * by fast have "\i. {y \ space M. max 0 (s i y) \ 0} \ {y \ space M. s i y \ 0}" unfolding max_def by force moreover have "\i. {y \ space M. s i y \ 0} \ sets M" using * by measurable ultimately have "\i. emeasure M {y \ space M. max 0 (s i y) \ 0} \ emeasure M {y \ space M. s i y \ 0}" by (rule emeasure_mono) hence **:"\i. emeasure M {y \ space M. max 0 (s i y) \ 0} \ \" using *(2) by (auto intro: order.strict_trans1 simp add: top.not_eq_extremum) have "\x. x \ space M \ (\i. max 0 (s i x)) \ max 0 (f x)" using *(3) tendsto_max by blast moreover have "\x i. x \ space M \ norm (max 0 (s i x)) \ norm (2 *\<^sub>R f x)" using *(4) unfolding max_def by auto ultimately have tendsto: "(\i. integral\<^sup>L M (\x. max 0 (s i x))) \ integral\<^sup>L M (\x. max 0 (f x))" using borel_measurable_simple_function simple integrable by (intro integral_dominated_convergence[OF max(1) _ integrable_norm[OF integrable_scaleR_right], of _ 2 f], blast+) { fix h :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assume "simple_function M h" "emeasure M {y \ space M. h y \ 0} \ \" "\x. x \ space M \ h x \ 0" hence *: "integral\<^sup>L M h \ 0" proof (induct rule: integrable_simple_function_induct_nn) case (cong f g) then show ?case using Bochner_Integration.integral_cong by force next case (indicator A y) hence "A \ {} \ y \ 0" using sets.sets_into_space by fastforce then show ?case using indicator by (cases "A = {}", auto simp add: scaleR_nonneg_nonneg) next case (add f g) then show ?case by (simp add: integrable_simple_function) qed } thus ?thesis using LIMSEQ_le_const[OF tendsto, of 0] ** simple by fastforce qed also have "\ = integral\<^sup>L M f" using nonneg by (auto intro: integral_cong_AE) finally show ?thesis . qed (simp add: not_integrable_integral_eq) lemma integral_mono_AE_banach: fixes f g :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" shows "integral\<^sup>L M f \ integral\<^sup>L M g" using integral_nonneg_AE_banach[of "\x. g x - f x"] assms Bochner_Integration.integral_diff[OF assms(1,2)] by force lemma integral_mono_banach: fixes f g :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ f x \ g x" shows "integral\<^sup>L M f \ integral\<^sup>L M g" using integral_mono_AE_banach assms by blast subsection \Integrability and Measurability of the Diameter\ context fixes s :: "nat \ 'a \ 'b :: {second_countable_topology, banach}" and M assumes bounded: "\x. x \ space M \ bounded (range (\i. s i x))" begin lemma borel_measurable_diameter: assumes [measurable]: "\i. (s i) \ borel_measurable M" shows "(\x. diameter {s i x |i. n \ i}) \ borel_measurable M" proof - have "{s i x |i. N \ i} \ {}" for x N by blast hence diameter_SUP: "diameter {s i x |i. N \ i} = (SUP (i, j) \ {N..} \ {N..}. dist (s i x) (s j x))" for x N unfolding diameter_def by (auto intro!: arg_cong[of _ _ Sup]) have "case_prod dist ` ({s i x |i. n \ i} \ {s i x |i. n \ i}) = ((\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..}))" for x by fast hence *: "(\x. diameter {s i x |i. n \ i}) = (\x. Sup ((\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..})))" using diameter_SUP by (simp add: case_prod_beta') have "bounded ((\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..}))" if "x \ space M" for x by (rule bounded_imp_dist_bounded[OF bounded, OF that]) hence bdd: "bdd_above ((\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..}))" if "x \ space M" for x using that bounded_imp_bdd_above by presburger have "fst p \ borel_measurable M" "snd p \ borel_measurable M" if "p \ s ` {n..} \ s ` {n..}" for p using that by fastforce+ hence "(\x. fst p x - snd p x) \ borel_measurable M" if "p \ s ` {n..} \ s ` {n..}" for p using that borel_measurable_diff by simp hence "(\x. case p of (f, g) \ dist (f x) (g x)) \ borel_measurable M" if "p \ s ` {n..} \ s ` {n..}" for p unfolding dist_norm using that by measurable moreover have "countable (s ` {n..} \ s ` {n..})" by (intro countable_SIGMA countable_image, auto) ultimately show ?thesis unfolding * by (auto intro!: borel_measurable_cSUP bdd) qed lemma integrable_bound_diameter: fixes f :: "'a \ real" assumes "integrable M f" and [measurable]: "\i. (s i) \ borel_measurable M" and "\x i. x \ space M \ norm (s i x) \ f x" shows "integrable M (\x. diameter {s i x |i. n \ i})" proof - have "{s i x |i. N \ i} \ {}" for x N by blast hence diameter_SUP: "diameter {s i x |i. N \ i} = (SUP (i, j) \ {N..} \ {N..}. dist (s i x) (s j x))" for x N unfolding diameter_def by (auto intro!: arg_cong[of _ _ Sup]) { fix x assume x: "x \ space M" let ?S = "(\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..})" have "case_prod dist ` ({s i x |i. n \ i} \ {s i x |i. n \ i}) = (\(i, j). dist (s i x) (s j x)) ` ({n..} \ {n..})" by fast hence *: "diameter {s i x |i. n \ i} = Sup ?S" using diameter_SUP by (simp add: case_prod_beta') have "bounded ?S" by (rule bounded_imp_dist_bounded[OF bounded[OF x]]) hence Sup_S_nonneg:"0 \ Sup ?S" by (auto intro!: cSup_upper2 x bounded_imp_bdd_above) have "dist (s i x) (s j x) \ 2 * f x" for i j by (intro dist_triangle2[THEN order_trans, of _ 0]) (metis norm_conv_dist assms(3) x add_mono mult_2) hence "\c \ ?S. c \ 2 * f x" by force hence "Sup ?S \ 2 * f x" by (intro cSup_least, auto) hence "norm (Sup ?S) \ 2 * norm (f x)" using Sup_S_nonneg by auto also have "... = norm (2 *\<^sub>R f x)" by simp finally have "norm (diameter {s i x |i. n \ i}) \ norm (2 *\<^sub>R f x)" unfolding * . } hence "AE x in M. norm (diameter {s i x |i. n \ i}) \ norm (2 *\<^sub>R f x)" by blast thus "integrable M (\x. diameter {s i x |i. n \ i})" using borel_measurable_diameter by (intro Bochner_Integration.integrable_bound[OF assms(1)[THEN integrable_scaleR_right[of 2]]], measurable) qed end subsection \Auxiliary Lemmas for Set Integrals\ lemma set_integral_scaleR_left: assumes "A \ sets M" "c \ 0 \ integrable M f" - shows "LINT t:A|M. f t *\<^sub>R c = (LINT t:A|M. f t) *\<^sub>R c" + shows "(LINT t:A|M. f t *\<^sub>R c) = (LINT t:A|M. f t) *\<^sub>R c" unfolding set_lebesgue_integral_def using integrable_mult_indicator[OF assms] by (subst integral_scaleR_left[symmetric], auto) lemma nn_set_integral_eq_set_integral: assumes [measurable]:"integrable M f" and "AE x \ A in M. 0 \ f x" "A \ sets M" shows "(\\<^sup>+x\A. f x \M) = (\ x \ A. f x \M)" proof- have "(\\<^sup>+x. indicator A x *\<^sub>R f x \M) = (\ x \ A. f x \M)" unfolding set_lebesgue_integral_def using assms(2) by (intro nn_integral_eq_integral[of _ "\x. indicat_real A x *\<^sub>R f x"], blast intro: assms integrable_mult_indicator, fastforce) moreover have "(\\<^sup>+x. indicator A x *\<^sub>R f x \M) = (\\<^sup>+x\A. f x \M)" by (metis ennreal_0 indicator_simps(1) indicator_simps(2) mult.commute mult_1 mult_zero_left real_scaleR_def) ultimately show ?thesis by argo qed lemma set_integral_restrict_space: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "\ \ space M \ sets M" shows "set_lebesgue_integral (restrict_space M \) A f = set_lebesgue_integral M A (\x. indicator \ x *\<^sub>R f x)" unfolding set_lebesgue_integral_def by (subst integral_restrict_space, auto intro!: integrable_mult_indicator assms simp: mult.commute) lemma set_integral_const: fixes c :: "'b::{banach, second_countable_topology}" assumes "A \ sets M" "emeasure M A \ \" shows "set_lebesgue_integral M A (\_. c) = measure M A *\<^sub>R c" unfolding set_lebesgue_integral_def using assms by (metis has_bochner_integral_indicator has_bochner_integral_integral_eq infinity_ennreal_def less_top) lemma set_integral_mono_banach: fixes f g :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "set_integrable M A f" "set_integrable M A g" "\x. x \ A \ f x \ g x" shows "(LINT x:A|M. f x) \ (LINT x:A|M. g x)" using assms unfolding set_integrable_def set_lebesgue_integral_def by (auto intro: integral_mono_banach split: split_indicator) lemma set_integral_mono_AE_banach: fixes f g :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "set_integrable M A f" "set_integrable M A g" "AE x\A in M. f x \ g x" shows "set_lebesgue_integral M A f \ set_lebesgue_integral M A g" using assms unfolding set_lebesgue_integral_def by (auto simp add: set_integrable_def intro!: integral_mono_AE_banach[of M "\x. indicator A x *\<^sub>R f x" "\x. indicator A x *\<^sub>R g x"], simp add: indicator_def) subsection \Averaging Theorem\ text \We aim to lift results from the real case to arbitrary Banach spaces. Our fundamental tool in this regard will be the averaging theorem. The proof of this theorem is due to Serge Lang (Real and Functional Analysis) \cite{Lang_1993}. The theorem allows us to make statements about a function’s value almost everywhere, depending on the value its integral takes on various sets of the measure space.\ text \Before we introduce and prove the averaging theorem, we will first show the following lemma which is crucial for our proof. While not stated exactly in this manner, our proof makes use of the characterization of second-countable topological spaces given in the book General Topology by Ryszard Engelking (Theorem 4.1.15) \cite{engelking_1989}.\ (* Engelking's book General Topology *) lemma balls_countable_basis: obtains D :: "'a :: {metric_space, second_countable_topology} set" where "topological_basis (case_prod ball ` (D \ (\ \ {0<..})))" and "countable D" and "D \ {}" proof - obtain D :: "'a set" where dense_subset: "countable D" "D \ {}" "\open U; U \ {}\ \ \y \ D. y \ U" for U using countable_dense_exists by blast have "topological_basis (case_prod ball ` (D \ (\ \ {0<..})))" proof (intro topological_basis_iff[THEN iffD2], fast, clarify) fix U and x :: 'a assume asm: "open U" "x \ U" obtain e where e: "e > 0" "ball x e \ U" using asm openE by blast obtain y where y: "y \ D" "y \ ball x (e / 3)" using dense_subset(3)[OF open_ball, of x "e / 3"] centre_in_ball[THEN iffD2, OF divide_pos_pos[OF e(1), of 3]] by force obtain r where r: "r \ \ \ {e/3<.. ball y r" using r y by (simp add: dist_commute) hence "ball y r \ U" using r by (intro order_trans[OF _ e(2)], simp, metric) moreover have "ball y r \ (case_prod ball ` (D \ (\ \ {0<..})))" using y(1) r by force ultimately show "\B'\(case_prod ball ` (D \ (\ \ {0<..}))). x \ B' \ B' \ U" using * by meson qed thus ?thesis using that dense_subset by blast qed context sigma_finite_measure begin text \To show statements concerning \\\-finite measure spaces, one usually shows the statement for finite measure spaces and uses a limiting argument to show it for the \\\-finite case. The following induction scheme formalizes this.\ lemma sigma_finite_measure_induct[case_names finite_measure, consumes 0]: assumes "\(N :: 'a measure) \. finite_measure N \ N = restrict_space M \ \ \ \ sets M \ emeasure N \ \ \ \ emeasure N \ \ 0 \ almost_everywhere N Q" and [measurable]: "Measurable.pred M Q" shows "almost_everywhere M Q" proof - have *: "almost_everywhere N Q" if "finite_measure N" "N = restrict_space M \" "\ \ sets M" "emeasure N \ \ \" for N \ using that by (cases "emeasure N \ = 0", auto intro: emeasure_0_AE assms(1)) obtain A :: "nat \ 'a set" where A: "range A \ sets M" "(\i. A i) = space M" and emeasure_finite: "emeasure M (A i) \ \" for i using sigma_finite by metis note A(1)[measurable] have space_restr: "space (restrict_space M (A i)) = A i" for i unfolding space_restrict_space by simp { fix i have *: "{x \ A i \ space M. Q x} = {x \ space M. Q x} \ (A i)" by fast have "Measurable.pred (restrict_space M (A i)) Q" using A by (intro measurableI, auto simp add: space_restr intro!: sets_restrict_space_iff[THEN iffD2], measurable, auto) } note this[measurable] { fix i have "finite_measure (restrict_space M (A i))" using emeasure_finite by (intro finite_measureI, subst space_restr, subst emeasure_restrict_space, auto) hence "emeasure (restrict_space M (A i)) {x \ A i. \Q x} = 0" using emeasure_finite by (intro AE_iff_measurable[THEN iffD1, OF _ _ *], measurable, subst space_restr[symmetric], intro sets.top, auto simp add: emeasure_restrict_space) hence "emeasure M {x \ A i. \ Q x} = 0" by (subst emeasure_restrict_space[symmetric], auto) } hence "emeasure M (\i. {x \ A i. \ Q x}) = 0" by (intro emeasure_UN_eq_0, auto) moreover have "(\i. {x \ A i. \ Q x}) = {x \ space M. \ Q x}" using A by auto ultimately show ?thesis by (intro AE_iff_measurable[THEN iffD2], auto) qed (* Real Functional Analysis - Lang*) text \The Averaging Theorem allows us to make statements concerning how a function behaves almost everywhere, depending on its behaviour on average.\ lemma averaging_theorem: fixes f::"_ \ 'b::{second_countable_topology, banach}" assumes [measurable]:"integrable M f" and closed: "closed S" and "\A. A \ sets M \ measure M A > 0 \ (1 / measure M A) *\<^sub>R set_lebesgue_integral M A f \ S" shows "AE x in M. f x \ S" proof (induct rule: sigma_finite_measure_induct) case (finite_measure N \) interpret finite_measure N by (rule finite_measure) have integrable[measurable]: "integrable N f" using assms finite_measure by (auto simp: integrable_restrict_space integrable_mult_indicator) have average: "(1 / Sigma_Algebra.measure N A) *\<^sub>R set_lebesgue_integral N A f \ S" if "A \ sets N" "measure N A > 0" for A proof - have *: "A \ sets M" using that by (simp add: sets_restrict_space_iff finite_measure) have "A = A \ \" by (metis finite_measure(2,3) inf.orderE sets.sets_into_space space_restrict_space that(1)) hence "set_lebesgue_integral N A f = set_lebesgue_integral M A f" unfolding finite_measure by (subst set_integral_restrict_space, auto simp add: finite_measure set_lebesgue_integral_def indicator_inter_arith[symmetric]) moreover have "measure N A = measure M A" using that by (auto intro!: measure_restrict_space simp add: finite_measure sets_restrict_space_iff) ultimately show ?thesis using that * assms(3) by presburger qed obtain D :: "'b set" where balls_basis: "topological_basis (case_prod ball ` (D \ (\ \ {0<..})))" and countable_D: "countable D" using balls_countable_basis by blast have countable_balls: "countable (case_prod ball ` (D \ (\ \ {0<..})))" using countable_rat countable_D by blast obtain B where B_balls: "B \ case_prod ball ` (D \ (\ \ {0<..}))" "\B = -S" using topological_basis[THEN iffD1, OF balls_basis] open_Compl[OF assms(2)] by meson hence countable_B: "countable B" using countable_balls countable_subset by fast define b where "b = from_nat_into (B \ {{}})" have "B \ {{}} \ {}" by simp have range_b: "range b = B \ {{}}" using countable_B by (auto simp add: b_def intro!: range_from_nat_into) have open_b: "open (b i)" for i unfolding b_def using B_balls open_ball from_nat_into[of "B \ {{}}" i] by force have Union_range_b: "\(range b) = -S" using B_balls range_b by simp { fix v r assume ball_in_Compl: "ball v r \ -S" define A where "A = f -` ball v r \ space N" have dist_less: "dist (f x) v < r" if "x \ A" for x using that unfolding A_def vimage_def by (simp add: dist_commute) hence AE_less: "AE x \ A in N. norm (f x - v) < r" by (auto simp add: dist_norm) have *: "A \ sets N" unfolding A_def by simp have "emeasure N A = 0" proof - { assume asm: "emeasure N A > 0" hence measure_pos: "measure N A > 0" unfolding emeasure_eq_measure by simp hence "(1 / measure N A) *\<^sub>R set_lebesgue_integral N A f - v = (1 / measure N A) *\<^sub>R set_lebesgue_integral N A (\x. f x - v)" using integrable integrable_const * by (subst set_integral_diff(2), auto simp add: set_integrable_def set_integral_const[OF *] algebra_simps intro!: integrable_mult_indicator) moreover have "norm (\x\A. (f x - v)\N) \ (\x\A. norm (f x - v)\N)" using * by (auto intro!: integral_norm_bound[of N "\x. indicator A x *\<^sub>R (f x - v)", THEN order_trans] integrable_mult_indicator integrable simp add: set_lebesgue_integral_def) ultimately have "norm ((1 / measure N A) *\<^sub>R set_lebesgue_integral N A f - v) \ set_lebesgue_integral N A (\x. norm (f x - v)) / measure N A" using asm by (auto intro: divide_right_mono) also have "... < set_lebesgue_integral N A (\x. r) / measure N A" unfolding set_lebesgue_integral_def using asm * integrable integrable_const AE_less measure_pos by (intro divide_strict_right_mono integral_less_AE[of _ _ A] integrable_mult_indicator) (fastforce simp add: dist_less dist_norm indicator_def)+ also have "... = r" using * measure_pos by (simp add: set_integral_const) finally have "dist ((1 / measure N A) *\<^sub>R set_lebesgue_integral N A f) v < r" by (subst dist_norm) hence "False" using average[OF * measure_pos] by (metis ComplD dist_commute in_mono mem_ball ball_in_Compl) } thus ?thesis by fastforce qed } note * = this { fix b' assume "b' \ B" hence ball_subset_Compl: "b' \ -S" and ball_radius_pos: "\v \ D. \r>0. b' = ball v r" using B_balls by (blast, fast) } note ** = this hence "emeasure N (f -` b i \ space N) = 0" for i by (cases "b i = {}", simp) (metis UnE singletonD * range_b[THEN eq_refl, THEN range_subsetD]) hence "emeasure N (\i. f -` b i \ space N) = 0" using open_b by (intro emeasure_UN_eq_0) fastforce+ moreover have "(\i. f -` b i \ space N) = f -` (\(range b)) \ space N" by blast ultimately have "emeasure N (f -` (-S) \ space N) = 0" using Union_range_b by argo hence "AE x in N. f x \ -S" using open_Compl[OF assms(2)] by (intro AE_iff_measurable[THEN iffD2], auto) thus ?case by force qed (simp add: pred_sets2[OF borel_closed] assms(2)) lemma density_zero: fixes f::"'a \ 'b::{second_countable_topology, banach}" assumes "integrable M f" and density_0: "\A. A \ sets M \ set_lebesgue_integral M A f = 0" shows "AE x in M. f x = 0" using averaging_theorem[OF assms(1), of "{0}"] assms(2) by (simp add: scaleR_nonneg_nonneg) text \The following lemma shows that densities are unique in Banach spaces.\ lemma density_unique_banach: fixes f f'::"'a \ 'b::{second_countable_topology, banach}" assumes "integrable M f" "integrable M f'" and density_eq: "\A. A \ sets M \ set_lebesgue_integral M A f = set_lebesgue_integral M A f'" shows "AE x in M. f x = f' x" proof- { fix A assume asm: "A \ sets M" hence "LINT x|M. indicat_real A x *\<^sub>R (f x - f' x) = 0" using density_eq assms(1,2) by (simp add: set_lebesgue_integral_def algebra_simps Bochner_Integration.integral_diff[OF integrable_mult_indicator(1,1)]) } thus ?thesis using density_zero[OF Bochner_Integration.integrable_diff[OF assms(1,2)]] by (simp add: set_lebesgue_integral_def) qed lemma density_nonneg: fixes f::"_ \ 'b::{second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" and "\A. A \ sets M \ set_lebesgue_integral M A f \ 0" shows "AE x in M. f x \ 0" using averaging_theorem[OF assms(1), of "{0..}", OF closed_atLeast] assms(2) by (simp add: scaleR_nonneg_nonneg) corollary integral_nonneg_eq_0_iff_AE_banach: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \ f x" shows "integral\<^sup>L M f = 0 \ (AE x in M. f x = 0)" proof assume *: "integral\<^sup>L M f = 0" { fix A assume asm: "A \ sets M" have "0 \ integral\<^sup>L M (\x. indicator A x *\<^sub>R f x)" using nonneg by (subst integral_zero[of M, symmetric], intro integral_mono_AE_banach integrable_mult_indicator asm f integrable_zero, auto simp add: indicator_def) moreover have "... \ integral\<^sup>L M f" using nonneg by (intro integral_mono_AE_banach integrable_mult_indicator asm f, auto simp add: indicator_def) ultimately have "set_lebesgue_integral M A f = 0" unfolding set_lebesgue_integral_def using * by force } thus "AE x in M. f x = 0" by (intro density_zero f, blast) qed (auto simp add: integral_eq_zero_AE) corollary integral_eq_mono_AE_eq_AE: fixes f g :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" "integral\<^sup>L M f = integral\<^sup>L M g" "AE x in M. f x \ g x" shows "AE x in M. f x = g x" proof - define h where "h = (\x. g x - f x)" have "AE x in M. h x = 0" unfolding h_def using assms by (subst integral_nonneg_eq_0_iff_AE_banach[symmetric]) auto then show ?thesis unfolding h_def by auto qed end end \ No newline at end of file diff --git a/thys/Martingales/Conditional_Expectation_Banach.thy b/thys/Martingales/Conditional_Expectation_Banach.thy --- a/thys/Martingales/Conditional_Expectation_Banach.thy +++ b/thys/Martingales/Conditional_Expectation_Banach.thy @@ -1,1048 +1,1048 @@ (* Author: Ata Keskin, TU München *) theory Conditional_Expectation_Banach imports "HOL-Probability.Conditional_Expectation" "HOL-Probability.Independent_Family" Bochner_Integration_Supplement begin section \Conditional Expectation in Banach Spaces\ text \While constructing the conditional expectation operator, we have come up with the following approach, which is based on the construction in \cite{Hytoenen_2016}. Both our approach, and the one in \cite{Hytoenen_2016} are based on showing that the conditional expectation is a contraction on some dense subspace of the space of functions \L\<^sup>1(E)\. In our approach, we start by constructing the conditional expectation explicitly for simple functions. Then we show that the conditional expectation is a contraction on simple functions, i.e. \\E(s|F)(x)\ \ E(\s(x)\|F)\ for \\\-almost all \x \ \\ with \s : \ \ E\ simple and integrable. Using this, we can show that the conditional expectation of a convergent sequence of simple functions is again convergent. Finally, we show that this limit exhibits the properties of a conditional expectation. This approach has the benefit of being straightforward and easy to implement, since we could make use of the existing formalization for real-valued functions. To use the construction in \cite{Hytoenen_2016} we need more tools from functional analysis, which Isabelle/HOL currently does not have.\ text \Before we can talk about 'the' conditional expectation, we must define what it means for a function to have a conditional expectation.\ definition has_cond_exp :: "'a measure \ 'a measure \ ('a \ 'b) \ ('a \ 'b::{real_normed_vector, second_countable_topology}) \ bool" where "has_cond_exp M F f g = ((\A \ sets F. (\ x \ A. f x \M) = (\ x \ A. g x \M)) \ integrable M f \ integrable M g \ g \ borel_measurable F)" text \This predicate precisely characterizes what it means for a function \<^term>\f\ to have a conditional expectation \<^term>\g\, with respect to the measure \<^term>\M\ and the sub-\\\-algebra \<^term>\F\.\ lemma has_cond_expI': assumes "\A. A \ sets F \ (\ x \ A. f x \M) = (\ x \ A. g x \M)" "integrable M f" "integrable M g" "g \ borel_measurable F" shows "has_cond_exp M F f g" using assms unfolding has_cond_exp_def by simp lemma has_cond_expD: assumes "has_cond_exp M F f g" shows "\A. A \ sets F \ (\ x \ A. f x \M) = (\ x \ A. g x \M)" "integrable M f" "integrable M g" "g \ borel_measurable F" using assms unfolding has_cond_exp_def by simp+ text \Now we can use Hilbert’s \\\-operator to define the conditional expectation, if it exists.\ definition cond_exp :: "'a measure \ 'a measure \ ('a \ 'b) \ ('a \ 'b::{banach, second_countable_topology})" where "cond_exp M F f = (if \g. has_cond_exp M F f g then (SOME g. has_cond_exp M F f g) else (\_. 0))" lemma borel_measurable_cond_exp[measurable]: "cond_exp M F f \ borel_measurable F" by (metis cond_exp_def someI has_cond_exp_def borel_measurable_const) lemma integrable_cond_exp[intro]: "integrable M (cond_exp M F f)" by (metis cond_exp_def has_cond_expD(3) integrable_zero someI) lemma set_integrable_cond_exp[intro]: assumes "A \ sets M" shows "set_integrable M A (cond_exp M F f)" using integrable_mult_indicator[OF assms integrable_cond_exp, of F f] by (auto simp add: set_integrable_def intro!: integrable_mult_indicator[OF assms integrable_cond_exp]) lemma has_cond_exp_self: assumes "integrable M f" shows "has_cond_exp M (vimage_algebra (space M) f borel) f f" using assms by (auto intro!: has_cond_expI' measurable_vimage_algebra1) lemma has_cond_exp_sets_cong: assumes "sets F = sets G" shows "has_cond_exp M F = has_cond_exp M G" using assms unfolding has_cond_exp_def by force lemma cond_exp_sets_cong: assumes "sets F = sets G" shows "AE x in M. cond_exp M F f x = cond_exp M G f x" by (intro AE_I2, simp add: cond_exp_def has_cond_exp_sets_cong[OF assms, of M]) context sigma_finite_subalgebra begin lemma borel_measurable_cond_exp'[measurable]: "cond_exp M F f \ borel_measurable M" by (metis cond_exp_def someI has_cond_exp_def borel_measurable_const subalg measurable_from_subalg) lemma cond_exp_null: assumes "\g. has_cond_exp M F f g" shows "cond_exp M F f = (\_. 0)" unfolding cond_exp_def using assms by argo text \We state the tower property of the conditional expectation in terms of the predicate \<^term>\has_cond_exp\.\ lemma has_cond_exp_nested_subalg: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "subalgebra G F" "has_cond_exp M F f h" "has_cond_exp M G f h'" shows "has_cond_exp M F h' h" by (intro has_cond_expI') (metis assms has_cond_expD in_mono subalgebra_def)+ text \The following lemma shows that the conditional expectation is unique as an element of L1, given that it exists.\ lemma has_cond_exp_charact: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "has_cond_exp M F f g" shows "has_cond_exp M F f (cond_exp M F f)" "AE x in M. cond_exp M F f x = g x" proof - show cond_exp: "has_cond_exp M F f (cond_exp M F f)" using assms someI cond_exp_def by metis let ?MF = "restr_to_subalg M F" interpret sigma_finite_measure ?MF by (rule sigma_fin_subalg) { fix A assume "A \ sets ?MF" then have [measurable]: "A \ sets F" using sets_restr_to_subalg[OF subalg] by simp have "(\x \ A. g x \?MF) = (\x \ A. g x \M)" using assms subalg by (auto simp add: integral_subalgebra2 set_lebesgue_integral_def dest!: has_cond_expD) also have "... = (\x \ A. cond_exp M F f x \M)" using assms cond_exp by (simp add: has_cond_exp_def) also have "... = (\x \ A. cond_exp M F f x \?MF)" using subalg by (auto simp add: integral_subalgebra2 set_lebesgue_integral_def) finally have "(\x \ A. g x \?MF) = (\x \ A. cond_exp M F f x \?MF)" by simp } hence "AE x in ?MF. cond_exp M F f x = g x" using cond_exp assms subalg by (intro density_unique_banach, auto dest: has_cond_expD intro!: integrable_in_subalg) then show "AE x in M. cond_exp M F f x = g x" using AE_restr_to_subalg[OF subalg] by simp qed corollary cond_exp_charact: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "\A. A \ sets F \ (\ x \ A. f x \M) = (\ x \ A. g x \M)" "integrable M f" "integrable M g" "g \ borel_measurable F" shows "AE x in M. cond_exp M F f x = g x" by (intro has_cond_exp_charact has_cond_expI' assms) auto text \Identity on F-measurable functions:\ text \If an integrable function \<^term>\f\ is already \<^term>\F\-measurable, then \<^term>\cond_exp M F f = f\ \\\-a.e. This is a corollary of the lemma on the characterization of \<^term>\cond_exp\.\ corollary cond_exp_F_meas[intro, simp]: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "integrable M f" "f \ borel_measurable F" shows "AE x in M. cond_exp M F f x = f x" by (rule cond_exp_charact, auto intro: assms) text \Congruence\ lemma has_cond_exp_cong: assumes "integrable M f" "\x. x \ space M \ f x = g x" "has_cond_exp M F g h" shows "has_cond_exp M F f h" proof (intro has_cond_expI'[OF _ assms(1)]) fix A assume asm: "A \ sets F" hence "set_lebesgue_integral M A f = set_lebesgue_integral M A g" by (intro set_lebesgue_integral_cong) (meson assms(2) subalg in_mono subalgebra_def sets.sets_into_space subalgebra_def subsetD)+ thus "set_lebesgue_integral M A f = set_lebesgue_integral M A h" using asm assms(3) by (simp add: has_cond_exp_def) qed (auto simp add: has_cond_expD[OF assms(3)]) lemma cond_exp_cong: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" "integrable M g" "\x. x \ space M \ f x = g x" shows "AE x in M. cond_exp M F f x = cond_exp M F g x" proof (cases "\h. has_cond_exp M F f h") case True then obtain h where h: "has_cond_exp M F f h" "has_cond_exp M F g h" using has_cond_exp_cong assms by metis show ?thesis using h[THEN has_cond_exp_charact(2)] by fastforce next case False moreover have "\h. has_cond_exp M F g h" using False has_cond_exp_cong assms by auto ultimately show ?thesis unfolding cond_exp_def by auto qed lemma has_cond_exp_cong_AE: assumes "integrable M f" "AE x in M. f x = g x" "has_cond_exp M F g h" shows "has_cond_exp M F f h" using assms(1,2) subalg subalgebra_def subset_iff by (intro has_cond_expI', subst set_lebesgue_integral_cong_AE[OF _ assms(1)[THEN borel_measurable_integrable] borel_measurable_integrable(1)[OF has_cond_expD(2)[OF assms(3)]]]) (fast intro: has_cond_expD[OF assms(3)] integrable_cong_AE_imp[OF _ _ AE_symmetric])+ lemma has_cond_exp_cong_AE': assumes "h \ borel_measurable F" "AE x in M. h x = h' x" "has_cond_exp M F f h'" shows "has_cond_exp M F f h" using assms(1, 2) subalg subalgebra_def subset_iff using AE_restr_to_subalg2[OF subalg assms(2)] measurable_from_subalg by (intro has_cond_expI' , subst set_lebesgue_integral_cong_AE[OF _ measurable_from_subalg(1,1)[OF subalg], OF _ assms(1) has_cond_expD(4)[OF assms(3)]]) (fast intro: has_cond_expD[OF assms(3)] integrable_cong_AE_imp[OF _ _ AE_symmetric])+ lemma cond_exp_cong_AE: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" "integrable M g" "AE x in M. f x = g x" shows "AE x in M. cond_exp M F f x = cond_exp M F g x" proof (cases "\h. has_cond_exp M F f h") case True then obtain h where h: "has_cond_exp M F f h" "has_cond_exp M F g h" using has_cond_exp_cong_AE assms by (metis (mono_tags, lifting) eventually_mono) show ?thesis using h[THEN has_cond_exp_charact(2)] by fastforce next case False moreover have "\h. has_cond_exp M F g h" using False has_cond_exp_cong_AE assms by auto ultimately show ?thesis unfolding cond_exp_def by auto qed text \The conditional expectation operator on the reals, \<^term>\real_cond_exp\, satisfies the conditions of the conditional expectation as we have defined it.\ lemma has_cond_exp_real: fixes f :: "'a \ real" assumes "integrable M f" shows "has_cond_exp M F f (real_cond_exp M F f)" by (intro has_cond_expI', auto intro!: real_cond_exp_intA assms) lemma cond_exp_real[intro]: fixes f :: "'a \ real" assumes "integrable M f" shows "AE x in M. cond_exp M F f x = real_cond_exp M F f x" using has_cond_exp_charact has_cond_exp_real assms by blast lemma cond_exp_cmult: fixes f :: "'a \ real" assumes "integrable M f" shows "AE x in M. cond_exp M F (\x. c * f x) x = c * cond_exp M F f x" using real_cond_exp_cmult[OF assms(1), of c] assms(1)[THEN cond_exp_real] assms(1)[THEN integrable_mult_right, THEN cond_exp_real, of c] by fastforce subsection \Existence\ text \Showing the existence is a bit involved. Specifically, what we aim to show is that \<^term>\has_cond_exp M F f (cond_exp M F f)\ holds for any Bochner-integrable \<^term>\f\. We will employ the standard machinery of measure theory. First, we will prove existence for indicator functions. Then we will extend our proof by linearity to simple functions. Finally we use a limiting argument to show that the conditional expectation exists for all Bochner-integrable functions.\ text \Indicator functions\ lemma has_cond_exp_indicator: assumes "A \ sets M" "emeasure M A < \" shows "has_cond_exp M F (\x. indicat_real A x *\<^sub>R y) (\x. real_cond_exp M F (indicator A) x *\<^sub>R y)" proof (intro has_cond_expI', goal_cases) case (1 B) - have "\x\B. (indicat_real A x *\<^sub>R y) \M = (\x\B. indicat_real A x \M) *\<^sub>R y" using assms by (intro set_integral_scaleR_left, meson 1 in_mono subalg subalgebra_def, blast) + have "(\x\B. (indicat_real A x *\<^sub>R y) \M) = (\x\B. indicat_real A x \M) *\<^sub>R y" using assms by (intro set_integral_scaleR_left, meson 1 in_mono subalg subalgebra_def, blast) also have "... = (\x\B. real_cond_exp M F (indicator A) x \M) *\<^sub>R y" using 1 assms by (subst real_cond_exp_intA, auto) - also have "... = \x\B. (real_cond_exp M F (indicator A) x *\<^sub>R y) \M" using assms by (intro set_integral_scaleR_left[symmetric], meson 1 in_mono subalg subalgebra_def, blast) + also have "... = (\x\B. (real_cond_exp M F (indicator A) x *\<^sub>R y) \M)" using assms by (intro set_integral_scaleR_left[symmetric], meson 1 in_mono subalg subalgebra_def, blast) finally show ?case . next case 2 show ?case using integrable_scaleR_left integrable_real_indicator assms by blast next case 3 show ?case using assms by (intro integrable_scaleR_left, intro real_cond_exp_int, blast+) next case 4 show ?case by (intro borel_measurable_scaleR, intro Conditional_Expectation.borel_measurable_cond_exp, simp) qed lemma cond_exp_indicator[intro]: fixes y :: "'b::{second_countable_topology,banach}" assumes [measurable]: "A \ sets M" "emeasure M A < \" shows "AE x in M. cond_exp M F (\x. indicat_real A x *\<^sub>R y) x = cond_exp M F (indicator A) x *\<^sub>R y" proof - have "AE x in M. cond_exp M F (\x. indicat_real A x *\<^sub>R y) x = real_cond_exp M F (indicator A) x *\<^sub>R y" using has_cond_exp_indicator[OF assms] has_cond_exp_charact by blast thus ?thesis using cond_exp_real[OF integrable_real_indicator, OF assms] by fastforce qed text \Addition\ lemma has_cond_exp_add: fixes f g :: "'a \ 'b::{second_countable_topology,banach}" assumes "has_cond_exp M F f f'" "has_cond_exp M F g g'" shows "has_cond_exp M F (\x. f x + g x) (\x. f' x + g' x)" proof (intro has_cond_expI', goal_cases) case (1 A) - have "\x\A. (f x + g x)\M = (\x\A. f x \M) + (\x\A. g x \M)" using assms[THEN has_cond_expD(2)] subalg 1 by (intro set_integral_add(2), auto simp add: subalgebra_def set_integrable_def intro: integrable_mult_indicator) + have "(\x\A. (f x + g x)\M) = (\x\A. f x \M) + (\x\A. g x \M)" using assms[THEN has_cond_expD(2)] subalg 1 by (intro set_integral_add(2), auto simp add: subalgebra_def set_integrable_def intro: integrable_mult_indicator) also have "... = (\x\A. f' x \M) + (\x\A. g' x \M)" using assms[THEN has_cond_expD(1)[OF _ 1]] by argo - also have "... = \x\A. (f' x + g' x)\M" using assms[THEN has_cond_expD(3)] subalg 1 by (intro set_integral_add(2)[symmetric], auto simp add: subalgebra_def set_integrable_def intro: integrable_mult_indicator) + also have "... = (\x\A. (f' x + g' x)\M)" using assms[THEN has_cond_expD(3)] subalg 1 by (intro set_integral_add(2)[symmetric], auto simp add: subalgebra_def set_integrable_def intro: integrable_mult_indicator) finally show ?case . next case 2 show ?case by (metis Bochner_Integration.integrable_add assms has_cond_expD(2)) next case 3 show ?case by (metis Bochner_Integration.integrable_add assms has_cond_expD(3)) next case 4 show ?case using assms borel_measurable_add has_cond_expD(4) by blast qed lemma has_cond_exp_scaleR_right: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "has_cond_exp M F f f'" shows "has_cond_exp M F (\x. c *\<^sub>R f x) (\x. c *\<^sub>R f' x)" using has_cond_expD[OF assms] by (intro has_cond_expI', auto) lemma cond_exp_scaleR_right: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" shows "AE x in M. cond_exp M F (\x. c *\<^sub>R f x) x = c *\<^sub>R cond_exp M F f x" proof (cases "\f'. has_cond_exp M F f f'") case True then show ?thesis using assms has_cond_exp_charact has_cond_exp_scaleR_right by metis next case False show ?thesis proof (cases "c = 0") case True then show ?thesis by simp next case c_nonzero: False have "\f'. has_cond_exp M F (\x. c *\<^sub>R f x) f'" proof (standard, goal_cases) case 1 then obtain f' where f': "has_cond_exp M F (\x. c *\<^sub>R f x) f'" by blast have "has_cond_exp M F f (\x. inverse c *\<^sub>R f' x)" using has_cond_expD[OF f'] divideR_right[OF c_nonzero] assms by (intro has_cond_expI', auto) then show ?case using False by blast qed then show ?thesis using cond_exp_null[OF False] cond_exp_null by force qed qed lemma cond_exp_uminus: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" shows "AE x in M. cond_exp M F (\x. - f x) x = - cond_exp M F f x" using cond_exp_scaleR_right[OF assms, of "-1"] by force text \Together with the induction scheme \integrable_simple_function_induct\, we can show that the conditional expectation of an integrable simple function exists.\ corollary has_cond_exp_simple: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "simple_function M f" "emeasure M {y \ space M. f y \ 0} \ \" shows "has_cond_exp M F f (cond_exp M F f)" using assms proof (induction rule: integrable_simple_function_induct) case (cong f g) then show ?case using has_cond_exp_cong by (metis (no_types, opaque_lifting) Bochner_Integration.integrable_cong has_cond_expD(2) has_cond_exp_charact(1)) next case (indicator A y) then show ?case using has_cond_exp_charact[OF has_cond_exp_indicator] by fast next case (add u v) then show ?case using has_cond_exp_add has_cond_exp_charact(1) by blast qed text \Now comes the most difficult part. Given a convergent sequence of integrable simple functions \<^term>\\n. s n\, we must show that the sequence \<^term>\\n. cond_exp M F (s n)\ is also convergent. Furthermore, we must show that this limit satisfies the properties of a conditional expectation. Unfortunately, we will only be able to show that this sequence convergences in the L1-norm. Luckily, this is enough to show that the operator \<^term>\cond_exp M F\ preserves limits as a function from L1 to L1.\ text \In anticipation of this result, we show that the conditional expectation operator is a contraction for simple functions. We first reformulate the lemma \real_cond_exp_abs\, which shows the statement for real-valued functions, using our definitions. Then we show the statement for simple functions via induction.\ lemma cond_exp_contraction_real: fixes f :: "'a \ real" assumes integrable[measurable]: "integrable M f" shows "AE x in M. norm (cond_exp M F f x) \ cond_exp M F (\x. norm (f x)) x" proof- have int: "integrable M (\x. norm (f x))" using assms by blast have *: "AE x in M. 0 \ cond_exp M F (\x. norm (f x)) x" using cond_exp_real[THEN AE_symmetric, OF integrable_norm[OF integrable]] real_cond_exp_ge_c[OF integrable_norm[OF integrable], of 0] norm_ge_zero by fastforce - have **: "A \ sets F \ \x\A. \f x\ \M = \x\A. real_cond_exp M F (\x. norm (f x)) x \M" for A unfolding real_norm_def using assms integrable_abs real_cond_exp_intA by blast + have **: "A \ sets F \ (\x\A. \f x\ \M) = (\x\A. real_cond_exp M F (\x. norm (f x)) x \M)" for A unfolding real_norm_def using assms integrable_abs real_cond_exp_intA by blast have norm_int: "A \ sets F \ (\x\A. \f x\ \M) = (\\<^sup>+x\A. \f x\ \M)" for A using assms by (intro nn_set_integral_eq_set_integral[symmetric], blast, fastforce) (meson subalg subalgebra_def subsetD) have "AE x in M. real_cond_exp M F (\x. norm (f x)) x \ 0" using int real_cond_exp_ge_c by force hence cond_exp_norm_int: "A \ sets F \ (\x\A. real_cond_exp M F (\x. norm (f x)) x \M) = (\\<^sup>+x\A. real_cond_exp M F (\x. norm (f x)) x \M)" for A using assms by (intro nn_set_integral_eq_set_integral[symmetric], blast, fastforce) (meson subalg subalgebra_def subsetD) - have "A \ sets F \ \\<^sup>+x\A. \f x\\M = \\<^sup>+x\A. real_cond_exp M F (\x. norm (f x)) x \M" for A using ** norm_int cond_exp_norm_int by (auto simp add: nn_integral_set_ennreal) + have "A \ sets F \ (\\<^sup>+x\A. \f x\\M) = (\\<^sup>+x\A. real_cond_exp M F (\x. norm (f x)) x \M)" for A using ** norm_int cond_exp_norm_int by (auto simp add: nn_integral_set_ennreal) moreover have "(\x. ennreal \f x\) \ borel_measurable M" by measurable moreover have "(\x. ennreal (real_cond_exp M F (\x. norm (f x)) x)) \ borel_measurable F" by measurable ultimately have "AE x in M. nn_cond_exp M F (\x. ennreal \f x\) x = real_cond_exp M F (\x. norm (f x)) x" by (intro nn_cond_exp_charact[THEN AE_symmetric], auto) hence "AE x in M. nn_cond_exp M F (\x. ennreal \f x\) x \ cond_exp M F (\x. norm (f x)) x" using cond_exp_real[OF int] by force moreover have "AE x in M. \real_cond_exp M F f x\ = norm (cond_exp M F f x)" unfolding real_norm_def using cond_exp_real[OF assms] * by force ultimately have "AE x in M. ennreal (norm (cond_exp M F f x)) \ cond_exp M F (\x. norm (f x)) x" using real_cond_exp_abs[OF assms[THEN borel_measurable_integrable]] by fastforce hence "AE x in M. enn2real (ennreal (norm (cond_exp M F f x))) \ enn2real (cond_exp M F (\x. norm (f x)) x)" using ennreal_le_iff2 by force thus ?thesis using * by fastforce qed lemma cond_exp_contraction_simple: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "simple_function M f" "emeasure M {y \ space M. f y \ 0} \ \" shows "AE x in M. norm (cond_exp M F f x) \ cond_exp M F (\x. norm (f x)) x" using assms proof (induction rule: integrable_simple_function_induct) case (cong f g) hence ae: "AE x in M. f x = g x" by blast hence "AE x in M. cond_exp M F f x = cond_exp M F g x" using cong has_cond_exp_simple by (subst cond_exp_cong_AE) (auto intro!: has_cond_expD(2)) hence "AE x in M. norm (cond_exp M F f x) = norm (cond_exp M F g x)" by force moreover have "AE x in M. cond_exp M F (\x. norm (f x)) x = cond_exp M F (\x. norm (g x)) x" using ae cong has_cond_exp_simple by (subst cond_exp_cong_AE) (auto dest: has_cond_expD) ultimately show ?case using cong(6) by fastforce next case (indicator A y) hence "AE x in M. cond_exp M F (\a. indicator A a *\<^sub>R y) x = cond_exp M F (indicator A) x *\<^sub>R y" by blast hence *: "AE x in M. norm (cond_exp M F (\a. indicat_real A a *\<^sub>R y) x) \ norm y * cond_exp M F (\x. norm (indicat_real A x)) x" using cond_exp_contraction_real[OF integrable_real_indicator, OF indicator] by fastforce have "AE x in M. norm y * cond_exp M F (\x. norm (indicat_real A x)) x = norm y * real_cond_exp M F (\x. norm (indicat_real A x)) x" using cond_exp_real[OF integrable_real_indicator, OF indicator] by fastforce moreover have "AE x in M. cond_exp M F (\x. norm y * norm (indicat_real A x)) x = real_cond_exp M F (\x. norm y * norm (indicat_real A x)) x" using indicator by (intro cond_exp_real, auto) ultimately have "AE x in M. norm y * cond_exp M F (\x. norm (indicat_real A x)) x = cond_exp M F (\x. norm y * norm (indicat_real A x)) x" using real_cond_exp_cmult[of "\x. norm (indicat_real A x)" "norm y"] indicator by fastforce moreover have "(\x. norm y * norm (indicat_real A x)) = (\x. norm (indicat_real A x *\<^sub>R y))" by force ultimately show ?case using * by force next case (add u v) have "AE x in M. norm (cond_exp M F (\a. u a + v a) x) = norm (cond_exp M F u x + cond_exp M F v x)" using has_cond_exp_charact(2)[OF has_cond_exp_add, OF has_cond_exp_simple(1,1), OF add(1,2,3,4)] by fastforce moreover have "AE x in M. norm (cond_exp M F u x + cond_exp M F v x) \ norm (cond_exp M F u x) + norm (cond_exp M F v x)" using norm_triangle_ineq by blast moreover have "AE x in M. norm (cond_exp M F u x) + norm (cond_exp M F v x) \ cond_exp M F (\x. norm (u x)) x + cond_exp M F (\x. norm (v x)) x" using add(6,7) by fastforce moreover have "AE x in M. cond_exp M F (\x. norm (u x)) x + cond_exp M F (\x. norm (v x)) x = cond_exp M F (\x. norm (u x) + norm (v x)) x" using integrable_simple_function[OF add(1,2)] integrable_simple_function[OF add(3,4)] by (intro has_cond_exp_charact(2)[OF has_cond_exp_add[OF has_cond_exp_charact(1,1)], THEN AE_symmetric], auto intro: has_cond_exp_real) moreover have "AE x in M. cond_exp M F (\x. norm (u x) + norm (v x)) x = cond_exp M F (\x. norm (u x + v x)) x" using add(5) integrable_simple_function[OF add(1,2)] integrable_simple_function[OF add(3,4)] by (intro cond_exp_cong, auto) ultimately show ?case by force qed lemma has_cond_exp_simple_lim: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes integrable[measurable]: "integrable M f" and "\i. simple_function M (s i)" and "\i. emeasure M {y \ space M. s i y \ 0} \ \" and "\x. x \ space M \ (\i. s i x) \ f x" and "\x i. x \ space M \ norm (s i x) \ 2 * norm (f x)" obtains r where "strict_mono r" "has_cond_exp M F f (\x. lim (\i. cond_exp M F (s (r i)) x))" "AE x in M. convergent (\i. cond_exp M F (s (r i)) x)" proof - have [measurable]: "(s i) \ borel_measurable M" for i using assms(2) by (simp add: borel_measurable_simple_function) have integrable_s: "integrable M (\x. s i x)" for i using assms integrable_simple_function by blast have integrable_4f: "integrable M (\x. 4 * norm (f x))" using assms(1) by simp have integrable_2f: "integrable M (\x. 2 * norm (f x))" using assms(1) by simp have integrable_2_cond_exp_norm_f: "integrable M (\x. 2 * cond_exp M F (\x. norm (f x)) x)" by fast have "emeasure M {y \ space M. s i y - s j y \ 0} \ emeasure M {y \ space M. s i y \ 0} + emeasure M {y \ space M. s j y \ 0}" for i j using simple_functionD(2)[OF assms(2)] by (intro order_trans[OF emeasure_mono emeasure_subadditive], auto) hence fin_sup: "emeasure M {y \ space M. s i y - s j y \ 0} \ \" for i j using assms(3) by (metis (mono_tags) ennreal_add_eq_top linorder_not_less top.not_eq_extremum infinity_ennreal_def) have "emeasure M {y \ space M. norm (s i y - s j y) \ 0} \ emeasure M {y \ space M. s i y \ 0} + emeasure M {y \ space M. s j y \ 0}" for i j using simple_functionD(2)[OF assms(2)] by (intro order_trans[OF emeasure_mono emeasure_subadditive], auto) hence fin_sup_norm: "emeasure M {y \ space M. norm (s i y - s j y) \ 0} \ \" for i j using assms(3) by (metis (mono_tags) ennreal_add_eq_top linorder_not_less top.not_eq_extremum infinity_ennreal_def) have Cauchy: "Cauchy (\n. s n x)" if "x \ space M" for x using assms(4) LIMSEQ_imp_Cauchy that by blast hence bounded_range_s: "bounded (range (\n. s n x))" if "x \ space M" for x using that cauchy_imp_bounded by fast text \Since the sequence \<^term>\(\n. s n x)\ is Cauchy for almost all \<^term>\x\, we know that the diameter tends to zero almost everywhere.\ text \Dominated convergence tells us that the integral of the diameter also converges to zero.\ have "AE x in M. (\n. diameter {s i x | i. n \ i}) \ 0" using Cauchy cauchy_iff_diameter_tends_to_zero_and_bounded by fast moreover have "(\x. diameter {s i x |i. n \ i}) \ borel_measurable M" for n using bounded_range_s borel_measurable_diameter by measurable moreover have "AE x in M. norm (diameter {s i x |i. n \ i}) \ 4 * norm (f x)" for n proof - { fix x assume x: "x \ space M" have "diameter {s i x |i. n \ i} \ 2 * norm (f x) + 2 * norm (f x)" by (intro diameter_le, blast, subst dist_norm[symmetric], intro dist_triangle3[THEN order_trans, of 0], intro add_mono) (auto intro: assms(5)[OF x]) hence "norm (diameter {s i x |i. n \ i}) \ 4 * norm (f x)" using diameter_ge_0[OF bounded_subset[OF bounded_range_s], OF x, of "{s i x |i. n \ i}"] by force } thus ?thesis by fast qed ultimately have diameter_tendsto_zero: "(\n. LINT x|M. diameter {s i x | i. n \ i}) \ 0" by (intro integral_dominated_convergence[OF borel_measurable_const[of 0] _ integrable_4f, simplified]) (fast+) have diameter_integrable: "integrable M (\x. diameter {s i x | i. n \ i})" for n using assms(1,5) by (intro integrable_bound_diameter[OF bounded_range_s integrable_2f], auto) have dist_integrable: "integrable M (\x. dist (s i x) (s j x))" for i j using assms(5) dist_triangle3[of "s i _" _ 0, THEN order_trans, OF add_mono, of _ "2 * norm (f _)"] by (intro Bochner_Integration.integrable_bound[OF integrable_4f]) fastforce+ text \Since \<^term>\cond_exp M F\ is a contraction for simple functions, the following sequence of integral values is also Cauchy.\ text \This follows, since the distance between the terms of this sequence are always less than or equal to the diameter, which itself converges to zero.\ text \Hence, we obtain a subsequence which is Cauchy almost everywhere.\ have "\N. \i\N. \j\N. LINT x|M. norm (cond_exp M F (s i) x - cond_exp M F (s j) x) < e" if e_pos: "e > 0" for e proof - obtain N where *: "LINT x|M. diameter {s i x | i. n \ i} < e" if "n \ N" for n using that order_tendsto_iff[THEN iffD1, OF diameter_tendsto_zero, unfolded eventually_sequentially] e_pos by presburger { fix i j x assume asm: "i \ N" "j \ N" "x \ space M" have "case_prod dist ` ({s i x |i. N \ i} \ {s i x |i. N \ i}) = case_prod (\i j. dist (s i x) (s j x)) ` ({N..} \ {N..})" by fast hence "diameter {s i x | i. N \ i} = (SUP (i, j) \ {N..} \ {N..}. dist (s i x) (s j x))" unfolding diameter_def by auto moreover have "(SUP (i, j) \ {N..} \ {N..}. dist (s i x) (s j x)) \ dist (s i x) (s j x)" using asm bounded_imp_bdd_above[OF bounded_imp_dist_bounded, OF bounded_range_s] by (intro cSup_upper, auto) ultimately have "diameter {s i x | i. N \ i} \ dist (s i x) (s j x)" by presburger } hence "LINT x|M. dist (s i x) (s j x) < e" if "i \ N" "j \ N" for i j using that * by (intro integral_mono[OF dist_integrable diameter_integrable, THEN order.strict_trans1], blast+) moreover have "LINT x|M. norm (cond_exp M F (s i) x - cond_exp M F (s j) x) \ LINT x|M. dist (s i x) (s j x)" for i j proof - have "LINT x|M. norm (cond_exp M F (s i) x - cond_exp M F (s j) x) = LINT x|M. norm (cond_exp M F (s i) x + - 1 *\<^sub>R cond_exp M F (s j) x)" unfolding dist_norm by simp also have "... = LINT x|M. norm (cond_exp M F (\x. s i x - s j x) x)" using has_cond_exp_charact(2)[OF has_cond_exp_add[OF _ has_cond_exp_scaleR_right, OF has_cond_exp_charact(1,1), OF has_cond_exp_simple(1,1)[OF assms(2,3)]], THEN AE_symmetric, of i "-1" j] by (intro integral_cong_AE) force+ also have "... \ LINT x|M. cond_exp M F (\x. norm (s i x - s j x)) x" using cond_exp_contraction_simple[OF _ fin_sup, of i j] integrable_cond_exp assms(2) by (intro integral_mono_AE, fast+) also have "... = LINT x|M. norm (s i x - s j x)" unfolding set_integral_space(1)[OF integrable_cond_exp, symmetric] set_integral_space[OF dist_integrable[unfolded dist_norm], symmetric] by (intro has_cond_expD(1)[OF has_cond_exp_simple[OF _ fin_sup_norm], symmetric]) (metis assms(2) simple_function_compose1 simple_function_diff, metis sets.top subalg subalgebra_def) finally show ?thesis unfolding dist_norm . qed ultimately show ?thesis using order.strict_trans1 by meson qed then obtain r where strict_mono_r: "strict_mono r" and AE_Cauchy: "AE x in M. Cauchy (\i. cond_exp M F (s (r i)) x)" by (rule cauchy_L1_AE_cauchy_subseq[OF integrable_cond_exp], auto) hence ae_lim_cond_exp: "AE x in M. (\n. cond_exp M F (s (r n)) x) \ lim (\n. cond_exp M F (s (r n)) x)" using Cauchy_convergent_iff convergent_LIMSEQ_iff by fastforce text \Now that we have a candidate for the conditional expectation, we must show that it actually has the required properties.\ text \Dominated convergence shows that this limit is indeed integrable.\ text \Here, we again use the fact that conditional expectation is a contraction on simple functions.\ have cond_exp_bounded: "AE x in M. norm (cond_exp M F (s (r n)) x) \ cond_exp M F (\x. 2 * norm (f x)) x" for n proof - have "AE x in M. norm (cond_exp M F (s (r n)) x) \ cond_exp M F (\x. norm (s (r n) x)) x" by (rule cond_exp_contraction_simple[OF assms(2,3)]) moreover have "AE x in M. real_cond_exp M F (\x. norm (s (r n) x)) x \ real_cond_exp M F (\x. 2 * norm (f x)) x" using integrable_s integrable_2f assms(5) by (intro real_cond_exp_mono, auto) ultimately show ?thesis using cond_exp_real[OF integrable_norm, OF integrable_s, of "r n"] cond_exp_real[OF integrable_2f] by force qed have lim_integrable: "integrable M (\x. lim (\i. cond_exp M F (s (r i)) x))" by (intro integrable_dominated_convergence[OF _ borel_measurable_cond_exp' integrable_cond_exp ae_lim_cond_exp cond_exp_bounded], simp) text \Moreover, we can use the DCT twice to show that the conditional expectation property holds, i.e. the value of the integral of the candidate, agrees with \<^term>\f\ on sets \<^term>\A \ F\.\ { fix A assume A_in_sets_F: "A \ sets F" have "AE x in M. norm (indicator A x *\<^sub>R cond_exp M F (s (r n)) x) \ cond_exp M F (\x. 2 * norm (f x)) x" for n proof - have "AE x in M. norm (indicator A x *\<^sub>R cond_exp M F (s (r n)) x) \ norm (cond_exp M F (s (r n)) x)" unfolding indicator_def by simp thus ?thesis using cond_exp_bounded[of n] by force qed - hence lim_cond_exp_int: "(\n. LINT x:A|M. cond_exp M F (s (r n)) x) \ LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x)" + hence lim_cond_exp_int: "(\n. LINT x:A|M. cond_exp M F (s (r n)) x) \ (LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x))" using ae_lim_cond_exp measurable_from_subalg[OF subalg borel_measurable_indicator, OF A_in_sets_F] cond_exp_bounded unfolding set_lebesgue_integral_def by (intro integral_dominated_convergence[OF borel_measurable_scaleR borel_measurable_scaleR integrable_cond_exp]) (fastforce simp add: tendsto_scaleR)+ have "AE x in M. norm (indicator A x *\<^sub>R s (r n) x) \ 2 * norm (f x)" for n proof - have "AE x in M. norm (indicator A x *\<^sub>R s (r n) x) \ norm (s (r n) x)" unfolding indicator_def by simp thus ?thesis using assms(5)[of _ "r n"] by fastforce qed - hence lim_s_int: "(\n. LINT x:A|M. s (r n) x) \ LINT x:A|M. f x" + hence lim_s_int: "(\n. LINT x:A|M. s (r n) x) \ (LINT x:A|M. f x)" using measurable_from_subalg[OF subalg borel_measurable_indicator, OF A_in_sets_F] LIMSEQ_subseq_LIMSEQ[OF assms(4) strict_mono_r] assms(5) unfolding set_lebesgue_integral_def comp_def by (intro integral_dominated_convergence[OF borel_measurable_scaleR borel_measurable_scaleR integrable_2f]) (fastforce simp add: tendsto_scaleR)+ - have "LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x) = lim (\n. LINT x:A|M. cond_exp M F (s (r n)) x)" using limI[OF lim_cond_exp_int] by argo + have "(LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x)) = lim (\n. LINT x:A|M. cond_exp M F (s (r n)) x)" using limI[OF lim_cond_exp_int] by argo also have "... = lim (\n. LINT x:A|M. s (r n) x)" using has_cond_expD(1)[OF has_cond_exp_simple[OF assms(2,3)] A_in_sets_F, symmetric] by presburger - also have "... = LINT x:A|M. f x" using limI[OF lim_s_int] by argo - finally have "LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x) = LINT x:A|M. f x" . + also have "... = (LINT x:A|M. f x)" using limI[OF lim_s_int] by argo + finally have "(LINT x:A|M. lim (\n. cond_exp M F (s (r n)) x)) = (LINT x:A|M. f x)" . } text \Putting it all together, we have the statement we are looking for.\ hence "has_cond_exp M F f (\x. lim (\i. cond_exp M F (s (r i)) x))" using assms(1) lim_integrable by (intro has_cond_expI', auto) thus thesis using AE_Cauchy Cauchy_convergent strict_mono_r by (auto intro!: that) qed text \Now, we can show that the conditional expectation is well-defined for all integrable functions.\ corollary has_cond_expI: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" shows "has_cond_exp M F f (cond_exp M F f)" proof - obtain s where s_is: "\i. simple_function M (s i)" "\i. emeasure M {y \ space M. s i y \ 0} \ \" "\x. x \ space M \ (\i. s i x) \ f x" "\x i. x \ space M \ norm (s i x) \ 2 * norm (f x)" using integrable_implies_simple_function_sequence[OF assms] by blast show ?thesis using has_cond_exp_simple_lim[OF assms s_is] has_cond_exp_charact(1) by metis qed subsection \Properties\ text \The defining property of the conditional expectation now always holds, given that the function \<^term>\f\ is integrable.\ lemma cond_exp_set_integral: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" "A \ sets F" shows "(\ x \ A. f x \M) = (\ x \ A. cond_exp M F f x \M)" using has_cond_expD(1)[OF has_cond_expI, OF assms] by argo (* Tower Property *) text \The following property of the conditional expectation is called the "Tower Property".\ lemma cond_exp_nested_subalg: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" "subalgebra M G" "subalgebra G F" shows "AE \ in M. cond_exp M F f \ = cond_exp M F (cond_exp M G f) \" using has_cond_expI assms sigma_finite_subalgebra_def by (auto intro!: has_cond_exp_nested_subalg[THEN has_cond_exp_charact(2), THEN AE_symmetric] sigma_finite_subalgebra.has_cond_expI[OF sigma_finite_subalgebra.intro[OF assms(2)]] nested_subalg_is_sigma_finite) (* Linearity *) text \The conditional expectation is linear.\ lemma cond_exp_add: fixes f :: "'a \ 'b::{second_countable_topology,banach}" assumes "integrable M f" "integrable M g" shows "AE x in M. cond_exp M F (\x. f x + g x) x = cond_exp M F f x + cond_exp M F g x" using has_cond_exp_add[OF has_cond_expI(1,1), OF assms, THEN has_cond_exp_charact(2)] . lemma cond_exp_diff: fixes f :: "'a \ 'b :: {second_countable_topology, banach}" assumes "integrable M f" "integrable M g" shows "AE x in M. cond_exp M F (\x. f x - g x) x = cond_exp M F f x - cond_exp M F g x" using has_cond_exp_add[OF _ has_cond_exp_scaleR_right, OF has_cond_expI(1,1), OF assms, THEN has_cond_exp_charact(2), of "-1"] by simp lemma cond_exp_diff': fixes f :: "'a \ 'b :: {second_countable_topology, banach}" assumes "integrable M f" "integrable M g" shows "AE x in M. cond_exp M F (f - g) x = cond_exp M F f x - cond_exp M F g x" unfolding fun_diff_def using assms by (rule cond_exp_diff) lemma cond_exp_scaleR_left: fixes f :: "'a \ real" assumes "integrable M f" shows "AE x in M. cond_exp M F (\x. f x *\<^sub>R c) x = cond_exp M F f x *\<^sub>R c" using cond_exp_set_integral[OF assms] subalg assms unfolding subalgebra_def by (intro cond_exp_charact, subst set_integral_scaleR_left, blast, intro assms, subst set_integral_scaleR_left, blast, intro integrable_cond_exp) auto text \The conditional expectation operator is a contraction, i.e. a bounded linear operator with operator norm less than or equal to 1.\ text \To show this we first obtain a subsequence \<^term>\\x. (\i. s (r i) x)\, such that \<^term>\(\i. cond_exp M F (s (r i)) x)\ converges to \<^term>\cond_exp M F f x\ a.e. Afterwards, we obtain a sub-subsequence \<^term>\\x. (\i. s (r (r' i)) x)\, such that \<^term>\(\i. cond_exp M F (\x. norm (s (r i))) x)\ converges to \<^term>\cond_exp M F (\x. norm (f x)) x\ a.e. Finally, we show that the inequality holds by showing that the terms of the subsequences obey the inequality and the fact that a subsequence of a convergent sequence converges to the same limit.\ lemma cond_exp_contraction: fixes f :: "'a \ 'b::{second_countable_topology, banach}" assumes "integrable M f" shows "AE x in M. norm (cond_exp M F f x) \ cond_exp M F (\x. norm (f x)) x" proof - obtain s where s: "\i. simple_function M (s i)" "\i. emeasure M {y \ space M. s i y \ 0} \ \" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" by (blast intro: integrable_implies_simple_function_sequence[OF assms]) obtain r where r: "strict_mono r" and "has_cond_exp M F f (\x. lim (\i. cond_exp M F (s (r i)) x))" "AE x in M. (\i. cond_exp M F (s (r i)) x) \ lim (\i. cond_exp M F (s (r i)) x)" using has_cond_exp_simple_lim[OF assms s] unfolding convergent_LIMSEQ_iff by blast hence r_tendsto: "AE x in M. (\i. cond_exp M F (s (r i)) x) \ cond_exp M F f x" using has_cond_exp_charact(2) by force have norm_s_r: "\i. simple_function M (\x. norm (s (r i) x))" "\i. emeasure M {y \ space M. norm (s (r i) y) \ 0} \ \" "\x. x \ space M \ (\i. norm (s (r i) x)) \ norm (f x)" "\i x. x \ space M \ norm (norm (s (r i) x)) \ 2 * norm (norm (f x))" using s by (auto intro: LIMSEQ_subseq_LIMSEQ[OF tendsto_norm r, unfolded comp_def] simple_function_compose1) obtain r' where r': "strict_mono r'" and "has_cond_exp M F (\x. norm (f x)) (\x. lim (\i. cond_exp M F (\x. norm (s (r (r' i)) x)) x))" "AE x in M. (\i. cond_exp M F (\x. norm (s (r (r' i)) x)) x) \ lim (\i. cond_exp M F (\x. norm (s (r (r' i)) x)) x)" using has_cond_exp_simple_lim[OF integrable_norm norm_s_r, OF assms] unfolding convergent_LIMSEQ_iff by blast hence r'_tendsto: "AE x in M. (\i. cond_exp M F (\x. norm (s (r (r' i)) x)) x) \ cond_exp M F (\x. norm (f x)) x" using has_cond_exp_charact(2) by force have "AE x in M. \i. norm (cond_exp M F (s (r (r' i))) x) \ cond_exp M F (\x. norm (s (r (r' i)) x)) x" using s by (auto intro: cond_exp_contraction_simple simp add: AE_all_countable) moreover have "AE x in M. (\i. norm (cond_exp M F (s (r (r' i))) x)) \ norm (cond_exp M F f x)" using r_tendsto LIMSEQ_subseq_LIMSEQ[OF tendsto_norm r', unfolded comp_def] by fast ultimately show ?thesis using LIMSEQ_le r'_tendsto by fast qed text \The following lemmas are called "pulling out whats known". We first show the statement for real-valued functions using the lemma \real_cond_exp_intg\, which is already present. We then show it for arbitrary \<^term>\g\ using the lecture notes of Gordan Zitkovic for the course "Theory of Probability I" \cite{Zitkovic_2015}.\ lemma cond_exp_measurable_mult: fixes f g :: "'a \ real" assumes [measurable]: "integrable M (\x. f x * g x)" "integrable M g" "f \ borel_measurable F" shows "integrable M (\x. f x * cond_exp M F g x)" "AE x in M. cond_exp M F (\x. f x * g x) x = f x * cond_exp M F g x" proof- show integrable: "integrable M (\x. f x * cond_exp M F g x)" using cond_exp_real[OF assms(2)] by (intro integrable_cong_AE_imp[OF real_cond_exp_intg(1), OF assms(1,3) assms(2)[THEN borel_measurable_integrable]] measurable_from_subalg[OF subalg]) auto interpret sigma_finite_measure "restr_to_subalg M F" by (rule sigma_fin_subalg) { fix A assume asm: "A \ sets F" hence asm': "A \ sets M" using subalg by (fastforce simp add: subalgebra_def) have "set_lebesgue_integral M A (cond_exp M F (\x. f x * g x)) = set_lebesgue_integral M A (\x. f x * g x)" by (simp add: cond_exp_set_integral[OF assms(1) asm]) also have "... = set_lebesgue_integral M A (\x. f x * real_cond_exp M F g x)" using borel_measurable_times[OF borel_measurable_indicator[OF asm] assms(3)] borel_measurable_integrable[OF assms(2)] integrable_mult_indicator[OF asm' assms(1)] by (fastforce simp add: set_lebesgue_integral_def mult.assoc[symmetric] intro: real_cond_exp_intg(2)[symmetric]) also have "... = set_lebesgue_integral M A (\x. f x * cond_exp M F g x)" using cond_exp_real[OF assms(2)] asm' borel_measurable_cond_exp' borel_measurable_cond_exp2 measurable_from_subalg[OF subalg assms(3)] by (auto simp add: set_lebesgue_integral_def intro: integral_cong_AE) - finally have "set_lebesgue_integral M A (cond_exp M F (\x. f x * g x)) = \x\A. (f x * cond_exp M F g x)\M" . + finally have "set_lebesgue_integral M A (cond_exp M F (\x. f x * g x)) = (\x\A. (f x * cond_exp M F g x)\M)" . } hence "AE x in restr_to_subalg M F. cond_exp M F (\x. f x * g x) x = f x * cond_exp M F g x" by (intro density_unique_banach integrable_cond_exp integrable integrable_in_subalg subalg, measurable, simp add: set_lebesgue_integral_def integral_subalgebra2[OF subalg] sets_restr_to_subalg[OF subalg]) thus "AE x in M. cond_exp M F (\x. f x * g x) x = f x * cond_exp M F g x" by (rule AE_restr_to_subalg[OF subalg]) qed lemma cond_exp_measurable_scaleR: fixes f :: "'a \ real" and g :: "'a \ 'b :: {second_countable_topology, banach}" assumes [measurable]: "integrable M (\x. f x *\<^sub>R g x)" "integrable M g" "f \ borel_measurable F" shows "integrable M (\x. f x *\<^sub>R cond_exp M F g x)" "AE x in M. cond_exp M F (\x. f x *\<^sub>R g x) x = f x *\<^sub>R cond_exp M F g x" proof - let ?F = "restr_to_subalg M F" have subalg': "subalgebra M (restr_to_subalg M F)" by (metis sets_eq_imp_space_eq sets_restr_to_subalg subalg subalgebra_def) { fix z assume asm[measurable]: "integrable M (\x. z x *\<^sub>R g x)" "z \ borel_measurable ?F" hence asm'[measurable]: "z \ borel_measurable F" using measurable_in_subalg' subalg by blast have "integrable M (\x. z x *\<^sub>R cond_exp M F g x)" "LINT x|M. z x *\<^sub>R g x = LINT x|M. z x *\<^sub>R cond_exp M F g x" proof - obtain s where s_is: "\i. simple_function ?F (s i)" "\x. x \ space ?F \ (\i. s i x) \ z x" "\i x. x \ space ?F \ norm (s i x) \ 2 * norm (z x)" using borel_measurable_implies_sequence_metric[OF asm(2), of 0] by force text \We need to apply the dominated convergence theorem twice, therefore we need to show the following prerequisites.\ have s_scaleR_g_tendsto: "AE x in M. (\i. s i x *\<^sub>R g x) \ z x *\<^sub>R g x" using s_is(2) by (simp add: space_restr_to_subalg tendsto_scaleR) have s_scaleR_cond_exp_g_tendsto: "AE x in ?F. (\i. s i x *\<^sub>R cond_exp M F g x) \ z x *\<^sub>R cond_exp M F g x" using s_is(2) by (simp add: tendsto_scaleR) have s_scaleR_g_meas: "(\x. s i x *\<^sub>R g x) \ borel_measurable M" for i using s_is(1)[THEN borel_measurable_simple_function, THEN subalg'[THEN measurable_from_subalg]] by simp have s_scaleR_cond_exp_g_meas: "(\x. s i x *\<^sub>R cond_exp M F g x) \ borel_measurable ?F" for i using s_is(1)[THEN borel_measurable_simple_function] measurable_in_subalg[OF subalg borel_measurable_cond_exp] by (fastforce intro: borel_measurable_scaleR) have s_scaleR_g_AE_bdd: "AE x in M. norm (s i x *\<^sub>R g x) \ 2 * norm (z x *\<^sub>R g x)" for i using s_is(3) by (fastforce simp add: space_restr_to_subalg mult.assoc[symmetric] mult_right_mono) { fix i have asm: "integrable M (\x. norm (z x) * norm (g x))" using asm(1)[THEN integrable_norm] by simp have "AE x in ?F. norm (s i x *\<^sub>R cond_exp M F g x) \ 2 * norm (z x) * norm (cond_exp M F g x)" using s_is(3) by (fastforce simp add: mult_mono) moreover have "AE x in ?F. norm (z x) * cond_exp M F (\x. norm (g x)) x = cond_exp M F (\x. norm (z x) * norm (g x)) x" by (rule cond_exp_measurable_mult(2)[THEN AE_symmetric, OF asm integrable_norm, OF assms(2), THEN AE_restr_to_subalg2[OF subalg]], auto) ultimately have "AE x in ?F. norm (s i x *\<^sub>R cond_exp M F g x) \ 2 * cond_exp M F (\x. norm (z x *\<^sub>R g x)) x" using cond_exp_contraction[OF assms(2), THEN AE_restr_to_subalg2[OF subalg]] order_trans[OF _ mult_mono] by fastforce } note s_scaleR_cond_exp_g_AE_bdd = this text \In the following section we need to pay attention to which measures we are using for integration. The rhs is F-measurable while the lhs is only M-measurable.\ { fix i have s_meas_M[measurable]: "s i \ borel_measurable M" by (meson borel_measurable_simple_function measurable_from_subalg s_is(1) subalg') have s_meas_F[measurable]: "s i \ borel_measurable F" by (meson borel_measurable_simple_function measurable_in_subalg' s_is(1) subalg) have s_scaleR_eq: "s i x *\<^sub>R h x = (\y\s i ` space M. (indicator (s i -` {y} \ space M) x *\<^sub>R y) *\<^sub>R h x)" if "x \ space M" for x and h :: "'a \ 'b" using simple_function_indicator_representation[OF s_is(1), of x i] that unfolding space_restr_to_subalg scaleR_left.sum[of _ _ "h x", symmetric] by presburger have "LINT x|M. s i x *\<^sub>R g x = LINT x|M. (\y\s i ` space M. indicator (s i -` {y} \ space M) x *\<^sub>R y *\<^sub>R g x)" using s_scaleR_eq by (intro Bochner_Integration.integral_cong) auto also have "... = (\y\s i ` space M. LINT x|M. indicator (s i -` {y} \ space M) x *\<^sub>R y *\<^sub>R g x)" by (intro Bochner_Integration.integral_sum integrable_mult_indicator[OF _ integrable_scaleR_right] assms(2)) simp also have "... = (\y\s i ` space M. y *\<^sub>R set_lebesgue_integral M (s i -` {y} \ space M) g)" by (simp only: set_lebesgue_integral_def[symmetric]) simp also have "... = (\y\s i ` space M. y *\<^sub>R set_lebesgue_integral M (s i -` {y} \ space M) (cond_exp M F g))" using assms(2) subalg borel_measurable_vimage[OF s_meas_F] by (subst cond_exp_set_integral, auto simp add: subalgebra_def) also have "... = (\y\s i ` space M. LINT x|M. indicator (s i -` {y} \ space M) x *\<^sub>R y *\<^sub>R cond_exp M F g x)" by (simp only: set_lebesgue_integral_def[symmetric]) simp also have "... = LINT x|M. (\y\s i ` space M. indicator (s i -` {y} \ space M) x *\<^sub>R y *\<^sub>R cond_exp M F g x)" by (intro Bochner_Integration.integral_sum[symmetric] integrable_mult_indicator[OF _ integrable_scaleR_right]) auto also have "... = LINT x|M. s i x *\<^sub>R cond_exp M F g x" using s_scaleR_eq by (intro Bochner_Integration.integral_cong) auto finally have "LINT x|M. s i x *\<^sub>R g x = LINT x|?F. s i x *\<^sub>R cond_exp M F g x" by (simp add: integral_subalgebra2[OF subalg]) } note integral_s_eq = this text \Now we just plug in the results we obtained into DCT, and use the fact that limits are unique.\ show "integrable M (\x. z x *\<^sub>R cond_exp M F g x)" using s_scaleR_cond_exp_g_meas asm(2) borel_measurable_cond_exp' by (intro integrable_from_subalg[OF subalg] integrable_cond_exp integrable_dominated_convergence[OF _ _ _ s_scaleR_cond_exp_g_tendsto s_scaleR_cond_exp_g_AE_bdd]) (auto intro: measurable_from_subalg[OF subalg] integrable_in_subalg measurable_in_subalg subalg) have "(\i. LINT x|M. s i x *\<^sub>R g x) \ LINT x|M. z x *\<^sub>R g x" using s_scaleR_g_meas asm(1)[THEN integrable_norm] asm' borel_measurable_cond_exp' by (intro integral_dominated_convergence[OF _ _ _ s_scaleR_g_tendsto s_scaleR_g_AE_bdd]) (auto intro: measurable_from_subalg[OF subalg]) moreover have "(\i. LINT x|?F. s i x *\<^sub>R cond_exp M F g x) \ LINT x|?F. z x *\<^sub>R cond_exp M F g x" using s_scaleR_cond_exp_g_meas asm(2) borel_measurable_cond_exp' by (intro integral_dominated_convergence[OF _ _ _ s_scaleR_cond_exp_g_tendsto s_scaleR_cond_exp_g_AE_bdd]) (auto intro: measurable_from_subalg[OF subalg] integrable_in_subalg measurable_in_subalg subalg) ultimately show "LINT x|M. z x *\<^sub>R g x = LINT x|M. z x *\<^sub>R cond_exp M F g x" using integral_s_eq using subalg by (simp add: LIMSEQ_unique integral_subalgebra2) qed } note * = this text \The main statement now follows with \<^term>\z = (\x. indicator A x * f x)\.\ show "integrable M (\x. f x *\<^sub>R cond_exp M F g x)" using * assms measurable_in_subalg[OF subalg] by blast { fix A assume asm: "A \ F" hence "integrable M (\x. indicat_real A x *\<^sub>R f x *\<^sub>R g x)" using subalg by (fastforce simp add: subalgebra_def intro!: integrable_mult_indicator assms(1)) hence "set_lebesgue_integral M A (\x. f x *\<^sub>R g x) = set_lebesgue_integral M A (\x. f x *\<^sub>R cond_exp M F g x)" unfolding set_lebesgue_integral_def using asm by (auto intro!: * measurable_in_subalg[OF subalg]) } thus "AE x in M. cond_exp M F (\x. f x *\<^sub>R g x) x = f x *\<^sub>R cond_exp M F g x" using borel_measurable_cond_exp by (intro cond_exp_charact, auto intro!: * assms measurable_in_subalg[OF subalg]) qed lemma cond_exp_sum [intro, simp]: fixes f :: "'t \ 'a \ 'b :: {second_countable_topology,banach}" assumes [measurable]: "\i. integrable M (f i)" shows "AE x in M. cond_exp M F (\x. \i\I. f i x) x = (\i\I. cond_exp M F (f i) x)" proof (rule has_cond_exp_charact, intro has_cond_expI') fix A assume [measurable]: "A \ sets F" then have A_meas [measurable]: "A \ sets M" by (meson subsetD subalg subalgebra_def) have "(\x\A. (\i\I. f i x)\M) = (\x. (\i\I. indicator A x *\<^sub>R f i x)\M)" unfolding set_lebesgue_integral_def by (simp add: scaleR_sum_right) also have "... = (\i\I. (\x. indicator A x *\<^sub>R f i x \M))" using assms by (auto intro!: Bochner_Integration.integral_sum integrable_mult_indicator) also have "... = (\i\I. (\x. indicator A x *\<^sub>R cond_exp M F (f i) x \M))" using cond_exp_set_integral[OF assms] by (simp add: set_lebesgue_integral_def) also have "... = (\x. (\i\I. indicator A x *\<^sub>R cond_exp M F (f i) x)\M)" using assms by (auto intro!: Bochner_Integration.integral_sum[symmetric] integrable_mult_indicator) also have "... = (\x\A. (\i\I. cond_exp M F (f i) x)\M)" unfolding set_lebesgue_integral_def by (simp add: scaleR_sum_right) finally show "(\x\A. (\i\I. f i x)\M) = (\x\A. (\i\I. cond_exp M F (f i) x)\M)" by auto qed (auto simp add: assms integrable_cond_exp) subsection \Linearly Ordered Banach Spaces\ text \In this subsection we show monotonicity results concerning the conditional expectation operator.\ lemma cond_exp_gr_c: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "AE x in M. f x > c" shows "AE x in M. cond_exp M F f x > c" proof - define X where "X = {x \ space M. cond_exp M F f x \ c}" have [measurable]: "X \ sets F" unfolding X_def by measurable (metis sets.top subalg subalgebra_def) hence X_in_M: "X \ sets M" using sets_restr_to_subalg subalg subalgebra_def by blast have "emeasure M X = 0" proof (rule ccontr) assume "emeasure M X \ 0" have "emeasure (restr_to_subalg M F) X = emeasure M X" by (simp add: emeasure_restr_to_subalg subalg) hence "emeasure (restr_to_subalg M F) X > 0" using \\(emeasure M X) = 0\ gr_zeroI by auto then obtain A where A: "A \ sets (restr_to_subalg M F)" "A \ X" "emeasure (restr_to_subalg M F) A > 0" "emeasure (restr_to_subalg M F) A < \" using sigma_fin_subalg by (metis emeasure_notin_sets ennreal_0 infinity_ennreal_def le_less_linear neq_top_trans not_gr_zero order_refl sigma_finite_measure.approx_PInf_emeasure_with_finite) hence [simp]: "A \ sets F" using subalg sets_restr_to_subalg by blast hence A_in_sets_M[simp]: "A \ sets M" using sets_restr_to_subalg subalg subalgebra_def by blast have [simp]: "set_integrable M A (\x. c)" using A subalg by (auto simp add: set_integrable_def emeasure_restr_to_subalg) have [simp]: "set_integrable M A f" unfolding set_integrable_def by (rule integrable_mult_indicator, auto simp add: assms(1)) have "AE x in M. indicator A x *\<^sub>R c = indicator A x *\<^sub>R f x" proof (rule integral_eq_mono_AE_eq_AE) have "(\x\A. c \M) \ (\x\A. f x \M)" using assms(2) by (intro set_integral_mono_AE_banach) auto moreover { have "(\x\A. f x \M) = (\x\A. cond_exp M F f x \M)" by (rule cond_exp_set_integral, auto simp add: assms) also have "... \ (\x\A. c \M)" using A by (auto intro!: set_integral_mono_banach simp add: X_def) finally have "(\x\A. f x \M) \ (\x\A. c \M)" by simp } ultimately show "LINT x|M. indicator A x *\<^sub>R c = LINT x|M. indicator A x *\<^sub>R f x" unfolding set_lebesgue_integral_def by simp show "AE x in M. indicator A x *\<^sub>R c \ indicator A x *\<^sub>R f x" using assms by (auto simp add: X_def indicator_def) qed (auto simp add: set_integrable_def[symmetric]) hence "AE x\A in M. c = f x" by auto hence "AE x\A in M. False" using assms(2) by auto hence "A \ null_sets M" using AE_iff_null_sets A_in_sets_M by metis thus False using A(3) by (simp add: emeasure_restr_to_subalg null_setsD1 subalg) qed thus ?thesis using AE_iff_null_sets[OF X_in_M] unfolding X_def by auto qed corollary cond_exp_less_c: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "AE x in M. f x < c" shows "AE x in M. cond_exp M F f x < c" proof - have "AE x in M. cond_exp M F f x = - cond_exp M F (\x. - f x) x" using cond_exp_uminus[OF assms(1)] by auto moreover have "AE x in M. cond_exp M F (\x. - f x) x > - c" using assms by (intro cond_exp_gr_c) auto ultimately show ?thesis by (force simp add: minus_less_iff) qed lemma cond_exp_mono_strict: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" "AE x in M. f x < g x" shows "AE x in M. cond_exp M F f x < cond_exp M F g x" using cond_exp_less_c[OF Bochner_Integration.integrable_diff, OF assms(1,2), of 0] cond_exp_diff[OF assms(1,2)] assms(3) by auto lemma cond_exp_ge_c: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes [measurable]: "integrable M f" and "AE x in M. f x \ c" shows "AE x in M. cond_exp M F f x \ c" proof - let ?F = "restr_to_subalg M F" interpret sigma_finite_measure "restr_to_subalg M F" using sigma_fin_subalg by auto { fix A assume asm: "A \ sets ?F" "0 < measure ?F A" have [simp]: "sets ?F = sets F" "measure ?F A = measure M A" using asm by (auto simp add: measure_def sets_restr_to_subalg[OF subalg] emeasure_restr_to_subalg[OF subalg]) have M_A: "emeasure M A < \" using measure_zero_top asm by (force simp add: top.not_eq_extremum) hence F_A: "emeasure ?F A < \" using asm(1) emeasure_restr_to_subalg subalg by fastforce have "set_lebesgue_integral M A (\_. c) \ set_lebesgue_integral M A f" using assms asm M_A subalg by (intro set_integral_mono_AE_banach, auto simp add: set_integrable_def integrable_mult_indicator subalgebra_def sets_restr_to_subalg) also have "... = set_lebesgue_integral M A (cond_exp M F f)" using cond_exp_set_integral[OF assms(1)] asm by auto also have "... = set_lebesgue_integral ?F A (cond_exp M F f)" unfolding set_lebesgue_integral_def using asm borel_measurable_cond_exp by (intro integral_subalgebra2[OF subalg, symmetric], simp) finally have "(1 / measure ?F A) *\<^sub>R set_lebesgue_integral ?F A (cond_exp M F f) \ {c..}" using asm subalg M_A by (auto simp add: set_integral_const subalgebra_def intro!: pos_divideR_le_eq[THEN iffD1]) } thus ?thesis using AE_restr_to_subalg[OF subalg] averaging_theorem[OF integrable_in_subalg closed_atLeast, OF subalg borel_measurable_cond_exp integrable_cond_exp] by auto qed corollary cond_exp_le_c: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" and "AE x in M. f x \ c" shows "AE x in M. cond_exp M F f x \ c" proof - have "AE x in M. cond_exp M F f x = - cond_exp M F (\x. - f x) x" using cond_exp_uminus[OF assms(1)] by force moreover have "AE x in M. cond_exp M F (\x. - f x) x \ - c" using assms by (intro cond_exp_ge_c) auto ultimately show ?thesis by (force simp add: minus_le_iff) qed corollary cond_exp_mono: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" shows "AE x in M. cond_exp M F f x \ cond_exp M F g x" using cond_exp_le_c[OF Bochner_Integration.integrable_diff, OF assms(1,2), of 0] cond_exp_diff[OF assms(1,2)] assms(3) by auto corollary cond_exp_min: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" shows "AE \ in M. cond_exp M F (\x. min (f x) (g x)) \ \ min (cond_exp M F f \) (cond_exp M F g \)" proof - have "AE \ in M. cond_exp M F (\x. min (f x) (g x)) \ \ cond_exp M F f \" by (intro cond_exp_mono integrable_min assms, simp) moreover have "AE \ in M. cond_exp M F (\x. min (f x) (g x)) \ \ cond_exp M F g \" by (intro cond_exp_mono integrable_min assms, simp) ultimately show "AE \ in M. cond_exp M F (\x. min (f x) (g x)) \ \ min (cond_exp M F f \) (cond_exp M F g \)" by fastforce qed corollary cond_exp_max: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" assumes "integrable M f" "integrable M g" shows "AE \ in M. cond_exp M F (\x. max (f x) (g x)) \ \ max (cond_exp M F f \) (cond_exp M F g \)" proof - have "AE \ in M. cond_exp M F (\x. max (f x) (g x)) \ \ cond_exp M F f \" by (intro cond_exp_mono integrable_max assms, simp) moreover have "AE \ in M. cond_exp M F (\x. max (f x) (g x)) \ \ cond_exp M F g \" by (intro cond_exp_mono integrable_max assms, simp) ultimately show "AE \ in M. cond_exp M F (\x. max (f x) (g x)) \ \ max (cond_exp M F f \) (cond_exp M F g \)" by fastforce qed corollary cond_exp_inf: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector, lattice}" assumes "integrable M f" "integrable M g" shows "AE \ in M. cond_exp M F (\x. inf (f x) (g x)) \ \ inf (cond_exp M F f \) (cond_exp M F g \)" unfolding inf_min using assms by (rule cond_exp_min) corollary cond_exp_sup: fixes f :: "'a \ 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector, lattice}" assumes "integrable M f" "integrable M g" shows "AE \ in M. cond_exp M F (\x. sup (f x) (g x)) \ \ sup (cond_exp M F f \) (cond_exp M F g \)" unfolding sup_max using assms by (rule cond_exp_max) end subsection \Probability Spaces\ lemma (in prob_space) sigma_finite_subalgebra_restr_to_subalg: assumes "subalgebra M F" shows "sigma_finite_subalgebra M F" proof (intro sigma_finite_subalgebra.intro) interpret F: prob_space "restr_to_subalg M F" using assms prob_space_restr_to_subalg prob_space_axioms by blast show "sigma_finite_measure (restr_to_subalg M F)" by (rule F.sigma_finite_measure_axioms) qed (rule assms) lemma (in prob_space) cond_exp_trivial: fixes f :: "'a \ 'b :: {second_countable_topology, banach}" assumes "integrable M f" shows "AE x in M. cond_exp M (sigma (space M) {}) f x = expectation f" proof - interpret sigma_finite_subalgebra M "sigma (space M) {}" by (auto intro: sigma_finite_subalgebra_restr_to_subalg simp add: subalgebra_def sigma_sets_empty_eq) show ?thesis using assms by (intro cond_exp_charact) (auto simp add: sigma_sets_empty_eq set_lebesgue_integral_def prob_space cong: Bochner_Integration.integral_cong) qed text \The following lemma shows that independent \\\-algebras don't matter for the conditional expectation. The proof is adapted from \cite{Zitkovic_2015}.\ lemma (in prob_space) cond_exp_indep_subalgebra: fixes f :: "'a \ 'b :: {second_countable_topology, banach, real_normed_field}" assumes subalgebra: "subalgebra M F" "subalgebra M G" and independent: "indep_set G (sigma (space M) (F \ vimage_algebra (space M) f borel))" assumes [measurable]: "integrable M f" shows "AE x in M. cond_exp M (sigma (space M) (F \ G)) f x = cond_exp M F f x" proof - interpret Un_sigma: sigma_finite_subalgebra M "sigma (space M) (F \ G)" using assms(1,2) by (auto intro!: sigma_finite_subalgebra_restr_to_subalg sets.sigma_sets_subset simp add: subalgebra_def space_measure_of_conv sets_measure_of_conv) interpret sigma_finite_subalgebra M F using assms by (auto intro: sigma_finite_subalgebra_restr_to_subalg) { fix A assume asm: "A \ sigma (space M) {a \ b | a b. a \ F \ b \ G}" have in_events: "sigma_sets (space M) {a \ b |a b. a \ sets F \ b \ sets G} \ events" using subalgebra by (intro sets.sigma_sets_subset, auto simp add: subalgebra_def) have "Int_stable {a \ b | a b. a \ F \ b \ G}" proof - { fix af bf ag bg assume F: "af \ F" "bf \ F" and G: "ag \ G" "bg \ G" have "af \ bf \ F" by (intro sets.Int F) moreover have "ag \ bg \ G" by (intro sets.Int G) ultimately have "\a b. af \ ag \ (bf \ bg) = a \ b \ a \ sets F \ b \ sets G" by (metis inf_assoc inf_left_commute) } thus ?thesis by (force intro!: Int_stableI) qed moreover have "{a \ b | a b. a \ F \ b \ G} \ Pow (space M)" using subalgebra by (force simp add: subalgebra_def dest: sets.sets_into_space) moreover have "A \ sigma_sets (space M) {a \ b | a b. a \ F \ b \ G}" using calculation asm by force ultimately have "set_lebesgue_integral M A f = set_lebesgue_integral M A (cond_exp M F f)" proof (induction rule: sigma_sets_induct_disjoint) case (basic A) then obtain a b where A: "A = a \ b" "a \ F" "b \ G" by blast hence events[measurable]: "a \ events" "b \ events" using subalgebra by (auto simp add: subalgebra_def) have [simp]: "sigma_sets (space M) {indicator b -` A \ space M |A. A \ borel} \ G" using borel_measurable_indicator[OF A(3), THEN measurable_sets] sets.top subalgebra by (intro sets.sigma_sets_subset') (fastforce simp add: subalgebra_def)+ have Un_in_sigma: "F \ vimage_algebra (space M) f borel \ sigma (space M) (F \ vimage_algebra (space M) f borel)" by (metis equalityE le_supI sets.space_closed sigma_le_sets space_vimage_algebra subalg subalgebra_def) have [intro]: "indep_var borel (indicator b) borel (\\. indicator a \ *\<^sub>R f \)" proof - have [simp]: "sigma_sets (space M) {(\\. indicator a \ *\<^sub>R f \) -` A \ space M |A. A \ borel} \ sigma (space M) (F \ vimage_algebra (space M) f borel)" proof - have *: "(\\. indicator a \ *\<^sub>R f \) \ borel_measurable (sigma (space M) (F \ vimage_algebra (space M) f borel))" using borel_measurable_indicator[OF A(2), THEN measurable_sets, OF borel_open] subalgebra by (intro borel_measurable_scaleR borel_measurableI Un_in_sigma[THEN subsetD]) (auto simp add: space_measure_of_conv subalgebra_def sets_vimage_algebra2) thus ?thesis using measurable_sets[OF *] by (intro sets.sigma_sets_subset', auto simp add: space_measure_of_conv) qed have "indep_set (sigma_sets (space M) {indicator b -` A \ space M |A. A \ borel}) (sigma_sets (space M) {(\\. indicator a \ *\<^sub>R f \) -` A \ space M |A. A \ borel})" using independent unfolding indep_set_def by (rule indep_sets_mono_sets, auto split: bool.split) thus ?thesis by (subst indep_var_eq, auto intro!: borel_measurable_scaleR) qed have [intro]: "indep_var borel (indicator b) borel (\\. indicat_real a \ *\<^sub>R cond_exp M F f \)" proof - have [simp]:"sigma_sets (space M) {(\\. indicator a \ *\<^sub>R cond_exp M F f \) -` A \ space M |A. A \ borel} \ sigma (space M) (F \ vimage_algebra (space M) f borel)" proof - have *: "(\\. indicator a \ *\<^sub>R cond_exp M F f \) \ borel_measurable (sigma (space M) (F \ vimage_algebra (space M) f borel))" using borel_measurable_indicator[OF A(2), THEN measurable_sets, OF borel_open] subalgebra borel_measurable_cond_exp[THEN measurable_sets, OF borel_open, of _ M F f] by (intro borel_measurable_scaleR borel_measurableI Un_in_sigma[THEN subsetD]) (auto simp add: space_measure_of_conv subalgebra_def) thus ?thesis using measurable_sets[OF *] by (intro sets.sigma_sets_subset', auto simp add: space_measure_of_conv) qed have "indep_set (sigma_sets (space M) {indicator b -` A \ space M |A. A \ borel}) (sigma_sets (space M) {(\\. indicator a \ *\<^sub>R cond_exp M F f \) -` A \ space M |A. A \ borel})" using independent unfolding indep_set_def by (rule indep_sets_mono_sets, auto split: bool.split) thus ?thesis by (subst indep_var_eq, auto intro!: borel_measurable_scaleR) qed have "set_lebesgue_integral M A f = (LINT x|M. indicator b x * (indicator a x *\<^sub>R f x))" unfolding set_lebesgue_integral_def A indicator_inter_arith by (intro Bochner_Integration.integral_cong, auto simp add: scaleR_scaleR[symmetric] indicator_times_eq_if(1)) also have "... = (LINT x|M. indicator b x) * (LINT x|M. indicator a x *\<^sub>R f x)" by (intro indep_var_lebesgue_integral Bochner_Integration.integrable_bound[OF integrable_const[of "1 :: 'b"] borel_measurable_indicator] integrable_mult_indicator[OF _ assms(4)], blast) (auto simp add: indicator_def) also have "... = (LINT x|M. indicator b x) * (LINT x|M. indicator a x *\<^sub>R cond_exp M F f x)" using cond_exp_set_integral[OF assms(4) A(2)] unfolding set_lebesgue_integral_def by argo also have "... = (LINT x|M. indicator b x * (indicator a x *\<^sub>R cond_exp M F f x))" by (intro indep_var_lebesgue_integral[symmetric] Bochner_Integration.integrable_bound[OF integrable_const[of "1 :: 'b"] borel_measurable_indicator] integrable_mult_indicator[OF _ integrable_cond_exp], blast) (auto simp add: indicator_def) also have "... = set_lebesgue_integral M A (cond_exp M F f)" unfolding set_lebesgue_integral_def A indicator_inter_arith by (intro Bochner_Integration.integral_cong, auto simp add: scaleR_scaleR[symmetric] indicator_times_eq_if(1)) finally show ?case . next case empty then show ?case unfolding set_lebesgue_integral_def by simp next case (compl A) have A_in_space: "A \ space M" using compl using in_events sets.sets_into_space by blast have "set_lebesgue_integral M (space M - A) f = set_lebesgue_integral M (space M - A \ A) f - set_lebesgue_integral M A f" using compl(1) in_events by (subst set_integral_Un[of "space M - A" A], blast) (simp | intro integrable_mult_indicator[folded set_integrable_def, OF _ assms(4)], fast)+ also have "... = set_lebesgue_integral M (space M - A \ A) (cond_exp M F f) - set_lebesgue_integral M A (cond_exp M F f)" using cond_exp_set_integral[OF assms(4) sets.top] compl subalgebra by (simp add: subalgebra_def Un_absorb2[OF A_in_space]) also have "... = set_lebesgue_integral M (space M - A) (cond_exp M F f)" using compl(1) in_events by (subst set_integral_Un[of "space M - A" A], blast) (simp | intro integrable_mult_indicator[folded set_integrable_def, OF _ integrable_cond_exp], fast)+ finally show ?case . next case (union A) have "set_lebesgue_integral M (\ (range A)) f = (\i. set_lebesgue_integral M (A i) f)" using union in_events by (intro lebesgue_integral_countable_add) (auto simp add: disjoint_family_onD intro!: integrable_mult_indicator[folded set_integrable_def, OF _ assms(4)]) also have "... = (\i. set_lebesgue_integral M (A i) (cond_exp M F f))" using union by presburger also have "... = set_lebesgue_integral M (\ (range A)) (cond_exp M F f)" using union in_events by (intro lebesgue_integral_countable_add[symmetric]) (auto simp add: disjoint_family_onD intro!: integrable_mult_indicator[folded set_integrable_def, OF _ integrable_cond_exp]) finally show ?case . qed } moreover have "sigma (space M) {a \ b | a b. a \ F \ b \ G} = sigma (space M) (F \ G)" proof - have "sigma_sets (space M) {a \ b |a b. a \ sets F \ b \ sets G} = sigma_sets (space M) (sets F \ sets G)" proof - { fix a b assume asm: "a \ F" "b \ G" hence "a \ b \ sigma_sets (space M) (F \ G)" using subalgebra unfolding Int_range_binary by (intro sigma_sets_Inter[OF _ binary_in_sigma_sets]) (force simp add: subalgebra_def dest: sets.sets_into_space)+ } moreover { fix a assume "a \ sets F" hence "a \ sigma_sets (space M) {a \ b |a b. a \ sets F \ b \ sets G}" using subalgebra sets.top[of G] sets.sets_into_space[of _ F] by (intro sigma_sets.Basic, auto simp add: subalgebra_def) } moreover { fix a assume "a \ sets F \ a \ sets G" "a \ sets F" hence "a \ sets G" by blast hence "a \ sigma_sets (space M) {a \ b |a b. a \ sets F \ b \ sets G}" using subalgebra sets.top[of F] sets.sets_into_space[of _ G] by (intro sigma_sets.Basic, auto simp add: subalgebra_def) } ultimately show ?thesis by (intro sigma_sets_eqI) auto qed thus ?thesis using subalgebra by (intro sigma_eqI) (force simp add: subalgebra_def dest: sets.sets_into_space)+ qed moreover have "(cond_exp M F f) \ borel_measurable (sigma (space M) (sets F \ sets G))" proof - have "F \ sigma (space M) (F \ G)" by (metis Un_least Un_upper1 measure_of_of_measure sets.space_closed sets_measure_of sigma_sets_subseteq subalg subalgebra(2) subalgebra_def) thus ?thesis using borel_measurable_cond_exp[THEN measurable_sets, OF borel_open, of _ M F f] subalgebra by (intro borel_measurableI, force simp only: space_measure_of_conv subalgebra_def) qed ultimately show ?thesis using assms(4) integrable_cond_exp by (intro Un_sigma.cond_exp_charact) presburger+ qed text \If a random variable is independent of a \\\-algebra \<^term>\F\, its conditional expectation \<^term>\cond_exp M F f\ is just its expectation.\ lemma (in prob_space) cond_exp_indep: fixes f :: "'a \ 'b :: {second_countable_topology, banach, real_normed_field}" assumes subalgebra: "subalgebra M F" and independent: "indep_set F (vimage_algebra (space M) f borel)" and integrable: "integrable M f" shows "AE x in M. cond_exp M F f x = expectation f" proof - have "indep_set F (sigma (space M) (sigma (space M) {} \ (vimage_algebra (space M) f borel)))" using independent unfolding indep_set_def by (rule indep_sets_mono_sets, simp add: bool.split) (metis bot.extremum dual_order.refl sets.sets_measure_of_eq sets.sigma_sets_subset' sets_vimage_algebra_space space_vimage_algebra sup.absorb_iff2) hence cond_exp_indep: "AE x in M. cond_exp M (sigma (space M) (sigma (space M) {} \ F)) f x = expectation f" using cond_exp_indep_subalgebra[OF _ subalgebra _ integrable, of "sigma (space M) {}"] cond_exp_trivial[OF integrable] by (auto simp add: subalgebra_def sigma_sets_empty_eq) have "sets (sigma (space M) (sigma (space M) {} \ F)) = F" using subalgebra sets.top[of F] unfolding subalgebra_def by (simp add: sigma_sets_empty_eq, subst insert_absorb[of "space M" F], blast) (metis insert_absorb[OF sets.empty_sets] sets.sets_measure_of_eq) hence "AE x in M. cond_exp M (sigma (space M) (sigma (space M) {} \ F)) f x = cond_exp M F f x" by (rule cond_exp_sets_cong) thus ?thesis using cond_exp_indep by force qed end \ No newline at end of file diff --git a/thys/Martingales/Martingale.thy b/thys/Martingales/Martingale.thy --- a/thys/Martingales/Martingale.thy +++ b/thys/Martingales/Martingale.thy @@ -1,738 +1,738 @@ (* Author: Ata Keskin, TU München *) theory Martingale imports Stochastic_Process Conditional_Expectation_Banach begin section \Martingales\ text \The following locales are necessary for defining martingales.\ subsection \Martingale\ text \A martingale is an adapted process where the expected value of the next observation, given all past observations, is equal to the current value.\ locale martingale = sigma_finite_filtered_measure + adapted_process + assumes integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and martingale_property: "\i j. t\<^sub>0 \ i \ i \ j \ AE \ in M. X i \ = cond_exp M (F i) (X j) \" locale martingale_order = martingale M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {order_topology, ordered_real_vector}" locale martingale_linorder = martingale M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {linorder_topology, ordered_real_vector}" sublocale martingale_linorder \ martingale_order .. lemma (in sigma_finite_filtered_measure) martingale_const_fun[intro]: assumes "integrable M f" "f \ borel_measurable (F t\<^sub>0)" shows "martingale M F t\<^sub>0 (\_. f)" using assms sigma_finite_subalgebra.cond_exp_F_meas[OF _ assms(1), THEN AE_symmetric] borel_measurable_mono by (unfold_locales) blast+ lemma (in sigma_finite_filtered_measure) martingale_cond_exp[intro]: assumes "integrable M f" shows "martingale M F t\<^sub>0 (\i. cond_exp M (F i) f)" using sigma_finite_subalgebra.borel_measurable_cond_exp' borel_measurable_cond_exp by (unfold_locales) (auto intro: sigma_finite_subalgebra.cond_exp_nested_subalg[OF _ assms] simp add: subalgebra_F subalgebras) corollary (in sigma_finite_filtered_measure) martingale_zero[intro]: "martingale M F t\<^sub>0 (\_ _. 0)" by fastforce corollary (in finite_filtered_measure) martingale_const[intro]: "martingale M F t\<^sub>0 (\_ _. c)" by fastforce subsection \Submartingale\ text \A submartingale is an adapted process where the expected value of the next observation, given all past observations, is greater than or equal to the current value.\ locale submartingale = sigma_finite_filtered_measure M F t\<^sub>0 + adapted_process M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {order_topology, ordered_real_vector}" + assumes integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and submartingale_property: "\i j. t\<^sub>0 \ i \ i \ j \ AE \ in M. X i \ \ cond_exp M (F i) (X j) \" locale submartingale_linorder = submartingale M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {linorder_topology}" lemma (in sigma_finite_filtered_measure) submartingale_const_fun[intro]: assumes "integrable M f" "f \ borel_measurable (F t\<^sub>0)" shows "submartingale M F t\<^sub>0 (\_. f)" proof - interpret martingale M F t\<^sub>0 "\_. f" using assms by (rule martingale_const_fun) show "submartingale M F t\<^sub>0 (\_. f)" using martingale_property by (unfold_locales) (force simp add: integrable)+ qed lemma (in sigma_finite_filtered_measure) submartingale_cond_exp[intro]: assumes "integrable M f" shows "submartingale M F t\<^sub>0 (\i. cond_exp M (F i) f)" proof - interpret martingale M F t\<^sub>0 "\i. cond_exp M (F i) f" using assms by (rule martingale_cond_exp) show "submartingale M F t\<^sub>0 (\i. cond_exp M (F i) f)" using martingale_property by (unfold_locales) (force simp add: integrable)+ qed corollary (in finite_filtered_measure) submartingale_const[intro]: "submartingale M F t\<^sub>0 (\_ _. c)" by fastforce sublocale martingale_order \ submartingale using martingale_property by (unfold_locales) (force simp add: integrable)+ sublocale martingale_linorder \ submartingale_linorder .. subsection \Supermartingale\ text \A supermartingale is an adapted process where the expected value of the next observation, given all past observations, is less than or equal to the current value.\ locale supermartingale = sigma_finite_filtered_measure M F t\<^sub>0 + adapted_process M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {order_topology, ordered_real_vector}" + assumes integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and supermartingale_property: "\i j. t\<^sub>0 \ i \ i \ j \ AE \ in M. X i \ \ cond_exp M (F i) (X j) \" locale supermartingale_linorder = supermartingale M F t\<^sub>0 X for M F t\<^sub>0 and X :: "_ \ _ \ _ :: {linorder_topology}" lemma (in sigma_finite_filtered_measure) supermartingale_const_fun[intro]: assumes "integrable M f" "f \ borel_measurable (F t\<^sub>0)" shows "supermartingale M F t\<^sub>0 (\_. f)" proof - interpret martingale M F t\<^sub>0 "\_. f" using assms by (rule martingale_const_fun) show "supermartingale M F t\<^sub>0 (\_. f)" using martingale_property by (unfold_locales) (force simp add: integrable)+ qed lemma (in sigma_finite_filtered_measure) supermartingale_cond_exp[intro]: assumes "integrable M f" shows "supermartingale M F t\<^sub>0 (\i. cond_exp M (F i) f)" proof - interpret martingale M F t\<^sub>0 "\i. cond_exp M (F i) f" using assms by (rule martingale_cond_exp) show "supermartingale M F t\<^sub>0 (\i. cond_exp M (F i) f)" using martingale_property by (unfold_locales) (force simp add: integrable)+ qed corollary (in finite_filtered_measure) supermartingale_const[intro]: "supermartingale M F t\<^sub>0 (\_ _. c)" by fastforce sublocale martingale_order \ supermartingale using martingale_property by (unfold_locales) (force simp add: integrable)+ sublocale martingale_linorder \ supermartingale_linorder .. text \A stochastic process is a martingale, if and only if it is both a submartingale and a supermartingale.\ lemma martingale_iff: shows "martingale M F t\<^sub>0 X \ submartingale M F t\<^sub>0 X \ supermartingale M F t\<^sub>0 X" proof (rule iffI) assume asm: "martingale M F t\<^sub>0 X" interpret martingale_order M F t\<^sub>0 X by (intro martingale_order.intro asm) show "submartingale M F t\<^sub>0 X \ supermartingale M F t\<^sub>0 X" using submartingale_axioms supermartingale_axioms by blast next assume asm: "submartingale M F t\<^sub>0 X \ supermartingale M F t\<^sub>0 X" interpret submartingale M F t\<^sub>0 X by (simp add: asm) interpret supermartingale M F t\<^sub>0 X by (simp add: asm) show "martingale M F t\<^sub>0 X" using submartingale_property supermartingale_property by (unfold_locales) (intro integrable, blast, force) qed subsection \Martingale Lemmas\ text \In the following segment, we cover basic properties of martingales.\ context martingale begin lemma cond_exp_diff_eq_zero: assumes "t\<^sub>0 \ i" "i \ j" shows "AE \ in M. cond_exp M (F i) (\\. X j \ - X i \) \ = 0" using martingale_property[OF assms] assms sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of "F i" j i] by fastforce lemma set_integral_eq: assumes "A \ F i" "t\<^sub>0 \ i" "i \ j" shows "set_lebesgue_integral M A (X i) = set_lebesgue_integral M A (X j)" proof - interpret sigma_finite_subalgebra M "F i" using assms(2) by blast - have "\x \ A. X i x \M = \x \ A. cond_exp M (F i) (X j) x \M" using martingale_property[OF assms(2,3)] borel_measurable_cond_exp' assms subalgebras subalgebra_def by (intro set_lebesgue_integral_cong_AE[OF _ random_variable]) fastforce+ - also have "... = \x \ A. X j x \M" using assms by (auto simp: integrable intro: cond_exp_set_integral[symmetric]) + have "(\x \ A. X i x \M) = (\x \ A. cond_exp M (F i) (X j) x \M)" using martingale_property[OF assms(2,3)] borel_measurable_cond_exp' assms subalgebras subalgebra_def by (intro set_lebesgue_integral_cong_AE[OF _ random_variable]) fastforce+ + also have "... = (\x \ A. X j x \M)" using assms by (auto simp: integrable intro: cond_exp_set_integral[symmetric]) finally show ?thesis . qed lemma scaleR_const[intro]: shows "martingale M F t\<^sub>0 (\i x. c *\<^sub>R X i x)" proof - { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" interpret sigma_finite_subalgebra M "F i" using asm by blast have "AE x in M. c *\<^sub>R X i x = cond_exp M (F i) (\x. c *\<^sub>R X j x) x" using asm cond_exp_scaleR_right[OF integrable, of j, THEN AE_symmetric] martingale_property[OF asm] by force } thus ?thesis by (unfold_locales) (auto simp add: integrable martingale.integrable) qed lemma uminus[intro]: shows "martingale M F t\<^sub>0 (- X)" using scaleR_const[of "-1"] by (force intro: back_subst[of "martingale M F t\<^sub>0"]) lemma add[intro]: assumes "martingale M F t\<^sub>0 Y" shows "martingale M F t\<^sub>0 (\i \. X i \ + Y i \)" proof - interpret Y: martingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ + Y i \ = cond_exp M (F i) (\x. X j x + Y j x) \" using sigma_finite_subalgebra.cond_exp_add[OF _ integrable martingale.integrable[OF assms], of "F i" j j, THEN AE_symmetric] martingale_property[OF asm] martingale.martingale_property[OF assms asm] by force } thus ?thesis using assms by (unfold_locales) (auto simp add: integrable martingale.integrable) qed lemma diff[intro]: assumes "martingale M F t\<^sub>0 Y" shows "martingale M F t\<^sub>0 (\i x. X i x - Y i x)" proof - interpret Y: martingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ - Y i \ = cond_exp M (F i) (\x. X j x - Y j x) \" using sigma_finite_subalgebra.cond_exp_diff[OF _ integrable martingale.integrable[OF assms], of "F i" j j, THEN AE_symmetric] martingale_property[OF asm] martingale.martingale_property[OF assms asm] by fastforce } thus ?thesis using assms by (unfold_locales) (auto simp add: integrable martingale.integrable) qed end text \Using properties of the conditional expectation, we present the following alternative characterizations of martingales.\ lemma (in sigma_finite_filtered_measure) martingale_of_cond_exp_diff_eq_zero: assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and diff_zero: "\i j. t\<^sub>0 \ i \ i \ j \ AE x in M. cond_exp M (F i) (\\. X j \ - X i \) x = 0" shows "martingale M F t\<^sub>0 X" proof interpret adapted_process M F t\<^sub>0 X by (rule adapted) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. X i \ = cond_exp M (F i) (X j) \" using diff_zero[OF asm] sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of "F i" j i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] by fastforce } qed (auto intro: integrable adapted[THEN adapted_process.adapted]) lemma (in sigma_finite_filtered_measure) martingale_of_set_integral_eq: assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and "\A i j. t\<^sub>0 \ i \ i \ j \ A \ F i \ set_lebesgue_integral M A (X i) = set_lebesgue_integral M A (X j)" shows "martingale M F t\<^sub>0 X" proof (unfold_locales) fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" interpret adapted_process M F t\<^sub>0 X by (rule adapted) interpret sigma_finite_subalgebra M "F i" using asm by blast interpret r: sigma_finite_measure "restr_to_subalg M (F i)" by (simp add: sigma_fin_subalg) { fix A assume "A \ restr_to_subalg M (F i)" hence *: "A \ F i" using sets_restr_to_subalg subalgebras asm by blast have "set_lebesgue_integral (restr_to_subalg M (F i)) A (X i) = set_lebesgue_integral M A (X i)" using * subalg asm by (auto simp: set_lebesgue_integral_def intro: integral_subalgebra2 borel_measurable_scaleR adapted borel_measurable_indicator) also have "... = set_lebesgue_integral M A (cond_exp M (F i) (X j))" using * assms(3)[OF asm] cond_exp_set_integral[OF integrable] asm by auto finally have "set_lebesgue_integral (restr_to_subalg M (F i)) A (X i) = set_lebesgue_integral (restr_to_subalg M (F i)) A (cond_exp M (F i) (X j))" using * subalg by (auto simp: set_lebesgue_integral_def intro!: integral_subalgebra2[symmetric] borel_measurable_scaleR borel_measurable_cond_exp borel_measurable_indicator) } hence "AE \ in restr_to_subalg M (F i). X i \ = cond_exp M (F i) (X j) \" using asm by (intro r.density_unique_banach, auto intro: integrable_in_subalg subalg borel_measurable_cond_exp integrable) thus "AE \ in M. X i \ = cond_exp M (F i) (X j) \" using AE_restr_to_subalg[OF subalg] by blast qed (auto intro: integrable adapted[THEN adapted_process.adapted]) subsection \Submartingale Lemmas\ context submartingale begin lemma cond_exp_diff_nonneg: assumes "t\<^sub>0 \ i" "i \ j" shows "AE x in M. cond_exp M (F i) (\\. X j \ - X i \) x \ 0" using submartingale_property[OF assms] assms sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of _ j i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] by fastforce lemma add[intro]: assumes "submartingale M F t\<^sub>0 Y" shows "submartingale M F t\<^sub>0 (\i \. X i \ + Y i \)" proof - interpret Y: submartingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ + Y i \ \ cond_exp M (F i) (\x. X j x + Y j x) \" using sigma_finite_subalgebra.cond_exp_add[OF _ integrable submartingale.integrable[OF assms], of "F i" j j] submartingale_property[OF asm] submartingale.submartingale_property[OF assms asm] add_mono[of "X i _" _ "Y i _"] by force } thus ?thesis using assms by (unfold_locales) (auto simp add: borel_measurable_add random_variable adapted integrable Y.random_variable Y.adapted submartingale.integrable) qed lemma diff[intro]: assumes "supermartingale M F t\<^sub>0 Y" shows "submartingale M F t\<^sub>0 (\i \. X i \ - Y i \)" proof - interpret Y: supermartingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ - Y i \ \ cond_exp M (F i) (\x. X j x - Y j x) \" using sigma_finite_subalgebra.cond_exp_diff[OF _ integrable supermartingale.integrable[OF assms], of "F i" j j] submartingale_property[OF asm] supermartingale.supermartingale_property[OF assms asm] diff_mono[of "X i _" _ _ "Y i _"] by force } thus ?thesis using assms by (unfold_locales) (auto simp add: borel_measurable_diff random_variable adapted integrable Y.random_variable Y.adapted supermartingale.integrable) qed lemma scaleR_nonneg: assumes "c \ 0" shows "submartingale M F t\<^sub>0 (\i \. c *\<^sub>R X i \)" proof { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. c *\<^sub>R X i \ \ cond_exp M (F i) (\\. c *\<^sub>R X j \) \" using sigma_finite_subalgebra.cond_exp_scaleR_right[OF _ integrable, of "F i" j c] submartingale_property[OF asm] by (fastforce intro!: scaleR_left_mono[OF _ assms]) } qed (auto simp add: borel_measurable_integrable borel_measurable_scaleR integrable random_variable adapted borel_measurable_const_scaleR) lemma scaleR_le_zero: assumes "c \ 0" shows "supermartingale M F t\<^sub>0 (\i \. c *\<^sub>R X i \)" proof { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. c *\<^sub>R X i \ \ cond_exp M (F i) (\\. c *\<^sub>R X j \) \" using sigma_finite_subalgebra.cond_exp_scaleR_right[OF _ integrable, of "F i" j c] submartingale_property[OF asm] by (fastforce intro!: scaleR_left_mono_neg[OF _ assms]) } qed (auto simp add: borel_measurable_integrable borel_measurable_scaleR integrable random_variable adapted borel_measurable_const_scaleR) lemma uminus[intro]: shows "supermartingale M F t\<^sub>0 (- X)" unfolding fun_Compl_def using scaleR_le_zero[of "-1"] by simp end context submartingale_linorder begin lemma set_integral_le: assumes "A \ F i" "t\<^sub>0 \ i" "i \ j" shows "set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X j)" using submartingale_property[OF assms(2), of j] assms subsetD[OF sets_F_subset] by (subst sigma_finite_subalgebra.cond_exp_set_integral[OF _ integrable assms(1), of j]) (auto intro!: scaleR_left_mono integral_mono_AE_banach integrable_mult_indicator integrable simp add: set_lebesgue_integral_def) lemma max: assumes "submartingale M F t\<^sub>0 Y" shows "submartingale M F t\<^sub>0 (\i \. max (X i \) (Y i \))" proof (unfold_locales) interpret Y: submartingale_linorder M F t\<^sub>0 Y by (intro submartingale_linorder.intro assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" have "AE \ in M. max (X i \) (Y i \) \ max (cond_exp M (F i) (X j) \) (cond_exp M (F i) (Y j) \)" using submartingale_property Y.submartingale_property asm unfolding max_def by fastforce thus "AE \ in M. max (X i \) (Y i \) \ cond_exp M (F i) (\\. max (X j \) (Y j \)) \" using sigma_finite_subalgebra.cond_exp_max[OF _ integrable Y.integrable, of "F i" j j] asm by (fast intro: order.trans) } show "\i. t\<^sub>0 \ i \ (\\. max (X i \) (Y i \)) \ borel_measurable (F i)" "\i. t\<^sub>0 \ i \ integrable M (\\. max (X i \) (Y i \))" by (force intro: Y.integrable integrable assms)+ qed lemma max_0: shows "submartingale M F t\<^sub>0 (\i \. max 0 (X i \))" proof - interpret zero: martingale_linorder M F t\<^sub>0 "\_ _. 0" by (force intro: martingale_linorder.intro martingale_order.intro) show ?thesis by (intro zero.max submartingale_linorder.intro submartingale_axioms) qed end lemma (in sigma_finite_filtered_measure) submartingale_of_cond_exp_diff_nonneg: assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and diff_nonneg: "\i j. t\<^sub>0 \ i \ i \ j \ AE x in M. cond_exp M (F i) (\\. X j \ - X i \) x \ 0" shows "submartingale M F t\<^sub>0 X" proof (unfold_locales) interpret adapted_process M F t\<^sub>0 X by (rule adapted) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. X i \ \ cond_exp M (F i) (X j) \" using diff_nonneg[OF asm] sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of "F i" j i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] by fastforce } qed (auto intro: integrable adapted[THEN adapted_process.adapted]) lemma (in sigma_finite_filtered_measure) submartingale_of_set_integral_le: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and "\A i j. t\<^sub>0 \ i \ i \ j \ A \ F i \ set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X j)" shows "submartingale M F t\<^sub>0 X" proof (unfold_locales) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" interpret adapted_process M F t\<^sub>0 X by (rule adapted) interpret r: sigma_finite_measure "restr_to_subalg M (F i)" using asm sigma_finite_subalgebra.sigma_fin_subalg by blast { fix A assume "A \ restr_to_subalg M (F i)" hence *: "A \ F i" using asm sets_restr_to_subalg subalgebras by blast have "set_lebesgue_integral (restr_to_subalg M (F i)) A (X i) = set_lebesgue_integral M A (X i)" using * asm subalgebras by (auto simp: set_lebesgue_integral_def intro: integral_subalgebra2 borel_measurable_scaleR adapted borel_measurable_indicator) also have "... \ set_lebesgue_integral M A (cond_exp M (F i) (X j))" using * assms(3)[OF asm] asm sigma_finite_subalgebra.cond_exp_set_integral[OF _ integrable] by fastforce also have "... = set_lebesgue_integral (restr_to_subalg M (F i)) A (cond_exp M (F i) (X j))" using * asm subalgebras by (auto simp: set_lebesgue_integral_def intro!: integral_subalgebra2[symmetric] borel_measurable_scaleR borel_measurable_cond_exp borel_measurable_indicator) finally have "0 \ set_lebesgue_integral (restr_to_subalg M (F i)) A (\\. cond_exp M (F i) (X j) \ - X i \)" using * asm subalgebras by (subst set_integral_diff, auto simp add: set_integrable_def sets_restr_to_subalg intro!: integrable adapted integrable_in_subalg borel_measurable_scaleR borel_measurable_indicator borel_measurable_cond_exp integrable_mult_indicator) } hence "AE \ in restr_to_subalg M (F i). 0 \ cond_exp M (F i) (X j) \ - X i \" by (intro r.density_nonneg integrable_in_subalg asm subalgebras borel_measurable_diff borel_measurable_cond_exp adapted Bochner_Integration.integrable_diff integrable_cond_exp integrable) thus "AE \ in M. X i \ \ cond_exp M (F i) (X j) \" using AE_restr_to_subalg[OF subalgebras] asm by simp } qed (auto intro: integrable adapted[THEN adapted_process.adapted]) subsection \Supermartingale Lemmas\ text \The following lemmas are exact duals of the ones for submartingales.\ context supermartingale begin lemma cond_exp_diff_nonneg: assumes "t\<^sub>0 \ i" "i \ j" shows "AE x in M. cond_exp M (F i) (\\. X i \ - X j \) x \ 0" using assms supermartingale_property[OF assms] sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of "F i" i j] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] by fastforce lemma add[intro]: assumes "supermartingale M F t\<^sub>0 Y" shows "supermartingale M F t\<^sub>0 (\i \. X i \ + Y i \)" proof - interpret Y: supermartingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ + Y i \ \ cond_exp M (F i) (\x. X j x + Y j x) \" using sigma_finite_subalgebra.cond_exp_add[OF _ integrable supermartingale.integrable[OF assms], of "F i" j j] supermartingale_property[OF asm] supermartingale.supermartingale_property[OF assms asm] add_mono[of _ "X i _" _ "Y i _"] by force } thus ?thesis using assms by (unfold_locales) (auto simp add: borel_measurable_add random_variable adapted integrable Y.random_variable Y.adapted supermartingale.integrable) qed lemma diff[intro]: assumes "submartingale M F t\<^sub>0 Y" shows "supermartingale M F t\<^sub>0 (\i \. X i \ - Y i \)" proof - interpret Y: submartingale M F t\<^sub>0 Y by (rule assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" hence "AE \ in M. X i \ - Y i \ \ cond_exp M (F i) (\x. X j x - Y j x) \" using sigma_finite_subalgebra.cond_exp_diff[OF _ integrable submartingale.integrable[OF assms], of "F i" j j, unfolded fun_diff_def] supermartingale_property[OF asm] submartingale.submartingale_property[OF assms asm] diff_mono[of _ "X i _" "Y i _"] by force } thus ?thesis using assms by (unfold_locales) (auto simp add: borel_measurable_diff random_variable adapted integrable Y.random_variable Y.adapted submartingale.integrable) qed lemma scaleR_nonneg: assumes "c \ 0" shows "supermartingale M F t\<^sub>0 (\i \. c *\<^sub>R X i \)" proof { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. c *\<^sub>R X i \ \ cond_exp M (F i) (\\. c *\<^sub>R X j \) \" using sigma_finite_subalgebra.cond_exp_scaleR_right[OF _ integrable, of "F i" j c] supermartingale_property[OF asm] by (fastforce intro!: scaleR_left_mono[OF _ assms]) } qed (auto simp add: borel_measurable_integrable borel_measurable_scaleR integrable random_variable adapted borel_measurable_const_scaleR) lemma scaleR_le_zero: assumes "c \ 0" shows "submartingale M F t\<^sub>0 (\i \. c *\<^sub>R X i \)" proof { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. c *\<^sub>R X i \ \ cond_exp M (F i) (\\. c *\<^sub>R X j \) \" using sigma_finite_subalgebra.cond_exp_scaleR_right[OF _ integrable, of "F i" j c] supermartingale_property[OF asm] by (fastforce intro!: scaleR_left_mono_neg[OF _ assms]) } qed (auto simp add: borel_measurable_integrable borel_measurable_scaleR integrable random_variable adapted borel_measurable_const_scaleR) lemma uminus[intro]: shows "submartingale M F t\<^sub>0 (- X)" unfolding fun_Compl_def using scaleR_le_zero[of "-1"] by simp end context supermartingale_linorder begin lemma set_integral_ge: assumes "A \ F i" "t\<^sub>0 \ i" "i \ j" shows "set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X j)" using supermartingale_property[OF assms(2), of j] assms subsetD[OF sets_F_subset] by (subst sigma_finite_subalgebra.cond_exp_set_integral[OF _ integrable assms(1), of j]) (auto intro!: scaleR_left_mono integral_mono_AE_banach integrable_mult_indicator integrable simp add: set_lebesgue_integral_def) lemma min: assumes "supermartingale M F t\<^sub>0 Y" shows "supermartingale M F t\<^sub>0 (\i \. min (X i \) (Y i \))" proof (unfold_locales) interpret Y: supermartingale_linorder M F t\<^sub>0 Y by (intro supermartingale_linorder.intro assms) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" have "AE \ in M. min (X i \) (Y i \) \ min (cond_exp M (F i) (X j) \) (cond_exp M (F i) (Y j) \)" using supermartingale_property Y.supermartingale_property asm unfolding min_def by fastforce thus "AE \ in M. min (X i \) (Y i \) \ cond_exp M (F i) (\\. min (X j \) (Y j \)) \" using sigma_finite_subalgebra.cond_exp_min[OF _ integrable Y.integrable, of "F i" j j] asm by (fast intro: order.trans) } show "\i. t\<^sub>0 \ i \ (\\. min (X i \) (Y i \)) \ borel_measurable (F i)" "\i. t\<^sub>0 \ i \ integrable M (\\. min (X i \) (Y i \))" by (force intro: Y.integrable integrable assms)+ qed lemma min_0: shows "supermartingale M F t\<^sub>0 (\i \. min 0 (X i \))" proof - interpret zero: martingale_linorder M F t\<^sub>0 "\_ _. 0" by (force intro: martingale_linorder.intro) show ?thesis by (intro zero.min supermartingale_linorder.intro supermartingale_axioms) qed end lemma (in sigma_finite_filtered_measure) supermartingale_of_cond_exp_diff_le_zero: assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and diff_le_zero: "\i j. t\<^sub>0 \ i \ i \ j \ AE x in M. cond_exp M (F i) (\\. X j \ - X i \) x \ 0" shows "supermartingale M F t\<^sub>0 X" proof interpret adapted_process M F t\<^sub>0 X by (rule adapted) { fix i j :: 'b assume asm: "t\<^sub>0 \ i" "i \ j" thus "AE \ in M. X i \ \ cond_exp M (F i) (X j) \" using diff_le_zero[OF asm] sigma_finite_subalgebra.cond_exp_diff[OF _ integrable(1,1), of "F i" j i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, of i] by fastforce } qed (auto intro: integrable adapted[THEN adapted_process.adapted]) lemma (in sigma_finite_filtered_measure) supermartingale_of_set_integral_ge: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F t\<^sub>0 X" and integrable: "\i. t\<^sub>0 \ i \ integrable M (X i)" and "\A i j. t\<^sub>0 \ i \ i \ j \ A \ F i \ set_lebesgue_integral M A (X j) \ set_lebesgue_integral M A (X i)" shows "supermartingale M F t\<^sub>0 X" proof - interpret adapted_process M F t\<^sub>0 X by (rule adapted) note * = set_integral_uminus[unfolded set_integrable_def, OF integrable_mult_indicator[OF _ integrable]] have "supermartingale M F t\<^sub>0 (-(- X))" using ord_eq_le_trans[OF * ord_le_eq_trans[OF le_imp_neg_le[OF assms(3)] *[symmetric]]] sets_F_subset[THEN subsetD] by (intro submartingale.uminus submartingale_of_set_integral_le[OF uminus_adapted]) (clarsimp simp add: fun_Compl_def integrable | fastforce)+ thus ?thesis unfolding fun_Compl_def by simp qed text \Many of the statements we have made concerning martingales can be simplified when the indexing set is the natural numbers. Given a point in time \<^term>\i \ \\, it suffices to consider the successor \<^term>\i + 1\, instead of all future times \<^term>\j \ i\.\ subsection \Discrete Time Martingales\ context nat_sigma_finite_filtered_measure begin text \A predictable martingale is necessarily constant.\ lemma predictable_const: assumes "martingale M F 0 X" and "predictable_process M F 0 X" shows "AE \ in M. X i \ = X j \" proof - interpret martingale M F 0 X by (rule assms) have *: "AE \ in M. X i \ = X 0 \" for i proof (induction i) case 0 then show ?case by (simp add: bot_nat_def) next case (Suc i) interpret S: adapted_process M F 0 "\i. X (Suc i)" by (intro predictable_implies_adapted_Suc assms) show ?case using Suc S.adapted[of i] martingale_property[OF _ le_SucI, of i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable, of "F i" "Suc i"] by fastforce qed show ?thesis using *[of i] *[of j] by force qed lemma martingale_of_set_integral_eq_Suc: assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\A i. A \ F i \ set_lebesgue_integral M A (X i) = set_lebesgue_integral M A (X (Suc i))" shows "martingale M F 0 X" proof (intro martingale_of_set_integral_eq adapted integrable) fix i j A assume asm: "i \ j" "A \ sets (F i)" show "set_lebesgue_integral M A (X i) = set_lebesgue_integral M A (X j)" using asm proof (induction "j - i" arbitrary: i j) case 0 then show ?case using asm by simp next case (Suc n) hence *: "n = j - Suc i" by linarith have "Suc i \ j" using Suc(2,3) by linarith thus ?case using sets_F_mono[OF _ le_SucI] Suc(4) Suc(1)[OF *] by (auto intro: assms(3)[THEN trans]) qed qed lemma martingale_nat: assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. X i \ = cond_exp M (F i) (X (Suc i)) \" shows "martingale M F 0 X" proof (unfold_locales) interpret adapted_process M F 0 X by (rule adapted) fix i j :: nat assume asm: "i \ j" show "AE \ in M. X i \ = cond_exp M (F i) (X j) \" using asm proof (induction "j - i" arbitrary: i j) case 0 hence "j = i" by simp thus ?case using sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable adapted, THEN AE_symmetric] by blast next case (Suc n) have j: "j = Suc (n + i)" using Suc by linarith have n: "n = n + i - i" using Suc by linarith have *: "AE \ in M. cond_exp M (F (n + i)) (X j) \ = X (n + i) \" unfolding j using assms(3)[THEN AE_symmetric] by blast have "AE \ in M. cond_exp M (F i) (X j) \ = cond_exp M (F i) (cond_exp M (F (n + i)) (X j)) \" by (intro cond_exp_nested_subalg integrable subalg, simp add: subalgebra_def sets_F_mono) hence "AE \ in M. cond_exp M (F i) (X j) \ = cond_exp M (F i) (X (n + i)) \" using cond_exp_cong_AE[OF integrable_cond_exp integrable *] by force thus ?case using Suc(1)[OF n] by fastforce qed qed (auto simp add: integrable adapted[THEN adapted_process.adapted]) lemma martingale_of_cond_exp_diff_Suc_eq_zero: assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. cond_exp M (F i) (\\. X (Suc i) \ - X i \) \ = 0" shows "martingale M F 0 X" proof (intro martingale_nat integrable adapted) interpret adapted_process M F 0 X by (rule adapted) fix i show "AE \ in M. X i \ = cond_exp M (F i) (X (Suc i)) \" using cond_exp_diff[OF integrable(1,1), of i "Suc i" i] cond_exp_F_meas[OF integrable adapted, of i] assms(3)[of i] by fastforce qed end subsection \Discrete Time Submartingales\ context nat_sigma_finite_filtered_measure begin lemma predictable_mono: assumes "submartingale M F 0 X" and "predictable_process M F 0 X" "i \ j" shows "AE \ in M. X i \ \ X j \" using assms(3) proof (induction "j - i" arbitrary: i j) case 0 then show ?case by simp next case (Suc n) hence *: "n = j - Suc i" by linarith interpret submartingale M F 0 X by (rule assms) interpret S: adapted_process M F 0 "\i. X (Suc i)" by (intro predictable_implies_adapted_Suc assms) have "Suc i \ j" using Suc(2,3) by linarith thus ?case using Suc(1)[OF *] S.adapted[of i] submartingale_property[OF _ le_SucI, of i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable, of "F i" "Suc i"] by fastforce qed lemma submartingale_of_set_integral_le_Suc: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\A i. A \ F i \ set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X (Suc i))" shows "submartingale M F 0 X" proof (intro submartingale_of_set_integral_le adapted integrable) fix i j A assume asm: "i \ j" "A \ sets (F i)" show "set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X j)" using asm proof (induction "j - i" arbitrary: i j) case 0 then show ?case using asm by simp next case (Suc n) hence *: "n = j - Suc i" by linarith have "Suc i \ j" using Suc(2,3) by linarith thus ?case using sets_F_mono[OF _ le_SucI] Suc(4) Suc(1)[OF *] by (auto intro: assms(3)[THEN order_trans]) qed qed lemma submartingale_nat: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. X i \ \ cond_exp M (F i) (X (Suc i)) \" shows "submartingale M F 0 X" proof - show ?thesis using subalg assms(3) integrable by (intro submartingale_of_set_integral_le_Suc adapted integrable ord_le_eq_trans[OF set_integral_mono_AE_banach cond_exp_set_integral[symmetric]]) (meson in_mono integrable_mult_indicator set_integrable_def subalgebra_def, meson integrable_cond_exp in_mono integrable_mult_indicator set_integrable_def subalgebra_def, fast+) qed lemma submartingale_of_cond_exp_diff_Suc_nonneg: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. cond_exp M (F i) (\\. X (Suc i) \ - X i \) \ \ 0" shows "submartingale M F 0 X" proof (intro submartingale_nat integrable adapted) interpret adapted_process M F 0 X by (rule assms) fix i show "AE \ in M. X i \ \ cond_exp M (F i) (X (Suc i)) \" using cond_exp_diff[OF integrable(1,1), of i "Suc i" i] cond_exp_F_meas[OF integrable adapted, of i] assms(3)[of i] by fastforce qed lemma submartingale_partial_sum_scaleR: assumes "submartingale_linorder M F 0 X" and "adapted_process M F 0 C" "\i. AE \ in M. 0 \ C i \" "\i. AE \ in M. C i \ \ R" shows "submartingale M F 0 (\n \. \i *\<^sub>R (X (Suc i) \ - X i \))" proof - interpret submartingale_linorder M F 0 X by (rule assms) interpret C: adapted_process M F 0 C by (rule assms) interpret C': adapted_process M F 0 "\i \. C (i - 1) \ *\<^sub>R (X i \ - X (i - 1) \)" by (intro adapted_process.scaleR_right_adapted adapted_process.diff_adapted, unfold_locales) (auto intro: adaptedD C.adaptedD)+ interpret S: adapted_process M F 0 "\n \. \i *\<^sub>R (X (Suc i) \ - X i \)" using C'.adapted_process_axioms[THEN partial_sum_Suc_adapted] diff_Suc_1 by simp have "integrable M (\x. C i x *\<^sub>R (X (Suc i) x - X i x))" for i using assms(3,4)[of i] by (intro Bochner_Integration.integrable_bound[OF integrable_scaleR_right, OF Bochner_Integration.integrable_diff, OF integrable(1,1), of R "Suc i" i]) (auto simp add: mult_mono) moreover have "AE \ in M. 0 \ cond_exp M (F i) (\\. (\i *\<^sub>R (X (Suc i) \ - X i \)) - (\i *\<^sub>R (X (Suc i) \ - X i \))) \" for i using sigma_finite_subalgebra.cond_exp_measurable_scaleR[OF _ calculation _ C.adapted, of i] cond_exp_diff_nonneg[OF _ le_SucI, OF _ order.refl, of i] assms(3,4)[of i] by (fastforce simp add: scaleR_nonneg_nonneg integrable) ultimately show ?thesis by (intro submartingale_of_cond_exp_diff_Suc_nonneg S.adapted_process_axioms Bochner_Integration.integrable_sum, blast+) qed lemma submartingale_partial_sum_scaleR': assumes "submartingale_linorder M F 0 X" and "predictable_process M F 0 C" "\i. AE \ in M. 0 \ C i \" "\i. AE \ in M. C i \ \ R" shows "submartingale M F 0 (\n \. \i *\<^sub>R (X (Suc i) \ - X i \))" proof - interpret Suc_C: adapted_process M F 0 "\i. C (Suc i)" using predictable_implies_adapted_Suc assms by blast show ?thesis by (intro submartingale_partial_sum_scaleR[OF assms(1), of _ R] assms) (intro_locales) qed end subsection \Discrete Time Supermartingales\ context nat_sigma_finite_filtered_measure begin lemma predictable_mono': assumes "supermartingale M F 0 X" and "predictable_process M F 0 X" "i \ j" shows "AE \ in M. X i \ \ X j \" using assms(3) proof (induction "j - i" arbitrary: i j) case 0 then show ?case by simp next case (Suc n) hence *: "n = j - Suc i" by linarith interpret supermartingale M F 0 X by (rule assms) interpret S: adapted_process M F 0 "\i. X (Suc i)" by (intro predictable_implies_adapted_Suc assms) have "Suc i \ j" using Suc(2,3) by linarith thus ?case using Suc(1)[OF *] S.adapted[of i] supermartingale_property[OF _ le_SucI, of i] sigma_finite_subalgebra.cond_exp_F_meas[OF _ integrable, of "F i" "Suc i"] by fastforce qed lemma supermartingale_of_set_integral_ge_Suc: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\A i. A \ F i \ set_lebesgue_integral M A (X i) \ set_lebesgue_integral M A (X (Suc i))" shows "supermartingale M F 0 X" proof - interpret adapted_process M F 0 X by (rule assms) interpret uminus_X: adapted_process M F 0 "-X" by (rule uminus_adapted) note * = set_integral_uminus[unfolded set_integrable_def, OF integrable_mult_indicator[OF _ integrable]] have "supermartingale M F 0 (-(- X))" using ord_eq_le_trans[OF * ord_le_eq_trans[OF le_imp_neg_le[OF assms(3)] *[symmetric]]] sets_F_subset[THEN subsetD] by (intro submartingale.uminus submartingale_of_set_integral_le_Suc[OF uminus_adapted]) (clarsimp simp add: fun_Compl_def integrable | fastforce)+ thus ?thesis unfolding fun_Compl_def by simp qed lemma supermartingale_nat: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. X i \ \ cond_exp M (F i) (X (Suc i)) \" shows "supermartingale M F 0 X" proof - interpret adapted_process M F 0 X by (rule assms) have "AE \ in M. - X i \ \ cond_exp M (F i) (\x. - X (Suc i) x) \" for i using assms(3) cond_exp_uminus[OF integrable, of i "Suc i"] by force hence "supermartingale M F 0 (-(- X))" by (intro submartingale.uminus submartingale_nat[OF uminus_adapted]) (auto simp add: fun_Compl_def integrable) thus ?thesis unfolding fun_Compl_def by simp qed lemma supermartingale_of_cond_exp_diff_Suc_le_zero: fixes X :: "_ \ _ \ _ :: {linorder_topology}" assumes adapted: "adapted_process M F 0 X" and integrable: "\i. integrable M (X i)" and "\i. AE \ in M. cond_exp M (F i) (\\. X (Suc i) \ - X i \) \ \ 0" shows "supermartingale M F 0 X" proof (intro supermartingale_nat integrable adapted) interpret adapted_process M F 0 X by (rule assms) fix i show "AE \ in M. X i \ \ cond_exp M (F i) (X (Suc i)) \" using cond_exp_diff[OF integrable(1,1), of i "Suc i" i] cond_exp_F_meas[OF integrable adapted, of i] assms(3)[of i] by fastforce qed end end \ No newline at end of file diff --git a/thys/Skip_Lists/Skip_List.thy b/thys/Skip_Lists/Skip_List.thy --- a/thys/Skip_Lists/Skip_List.thy +++ b/thys/Skip_Lists/Skip_List.thy @@ -1,1348 +1,1348 @@ (* File: Skip_List.thy Authors: Max W. Haslbeck, Manuel Eberl *) section \Randomized Skip Lists\ theory Skip_List imports Geometric_PMF Misc "Monad_Normalisation.Monad_Normalisation" begin text \Conflicting notation from \<^theory>\HOL-Analysis.Infinite_Sum\\ no_notation Infinite_Sum.abs_summable_on (infixr "abs'_summable'_on" 46) subsection \Preliminaries\ lemma bind_pmf_if': "(do {c \ C; ab \ (if c then A else B); D ab}::'a pmf) = do {c \ C; (if c then (A \ D) else (B \ D))}" by (metis (mono_tags, lifting)) abbreviation (input) Max\<^sub>0 where "Max\<^sub>0 \ (\A. Max (A \ {0}))" subsection \Definition of a Randomised Skip List\ text \ Given a set A we assign a geometric random variable (counting the number of failed Bernoulli trials before the first success) to every element in A. That means an arbitrary element of A is on level n with probability $(1-p)^{n}p$. We define he height of the skip list as the maximum assigned level. So a skip list with only one level has height 0 but the calculation of the expected height is cleaner this way. \ locale random_skip_list = fixes p::real begin definition q where "q = 1 - p" definition SL :: "('a::linorder) set \ ('a \ nat) pmf" where "SL A = Pi_pmf A 0 (\_. geometric_pmf p)" definition SL\<^sub>N :: "nat \ (nat \ nat) pmf" where "SL\<^sub>N n = SL {..Height of Skip List\ definition H where "H A = map_pmf (\f. Max\<^sub>0 (f ` A)) (SL A)" definition H\<^sub>N :: "nat \ nat pmf" where "H\<^sub>N n = H {.. The height of a skip list is independent of the values in a set A. For simplicity we can therefore work on the skip list over the set @{term "{..< card A}"} \ lemma assumes "finite A" shows "H A = H\<^sub>N (card A)" proof - define f' where "f' = (\x. if x \ A then the_inv_into {..0 ((f \ f') ` A) = Max\<^sub>0 (f ` {.. nat" using bij_betw_imp_surj_on bij_f' image_comp by metis have "H A = map_pmf (\f. Max\<^sub>0 (f ` A)) (map_pmf (\g. g \ f') (SL\<^sub>N (card A)))" using assms bij_f' unfolding H_def SL_def SL\<^sub>N_def by (subst Pi_pmf_bij_betw[of _ f' "{.. = H\<^sub>N (card A)" unfolding H\<^sub>N_def H_def SL\<^sub>N_def using * by (auto intro!: bind_pmf_cong simp add: map_pmf_def) finally show ?thesis by simp qed text \ The cumulative distribution function (CDF) of the height is the CDF of the geometric PMF to the power of n \ lemma prob_Max_IID_geometric_atMost: assumes "p \ {0..1}" shows "measure_pmf.prob (H\<^sub>N n) {..i} = (measure_pmf.prob (geometric_pmf p) {..i}) ^ n" (is "?lhs = ?rhs") proof - note SL_def[simp] SL\<^sub>N_def[simp] H_def[simp] H\<^sub>N_def[simp] have "{f. Max\<^sub>0 (f ` {.. i} = {.. {..i}" by auto then have "?lhs = measure_pmf.prob (SL\<^sub>N n) ({.. {..i})" by (simp add: vimage_def) also have "\ = measure_pmf.prob (SL\<^sub>N n) (PiE_dflt {.._. {..i}))" by (intro measure_prob_cong_0) (auto simp add: PiE_dflt_def pmf_Pi split: if_splits) also have "\ = measure_pmf.prob (geometric_pmf p) {..i} ^ n" using assms by (auto simp add: measure_Pi_pmf_PiE_dflt) finally show ?thesis by simp qed lemma prob_Max_IID_geometric_greaterThan: assumes "p \ {0<..1}" shows "measure_pmf.prob (H\<^sub>N n) {i<..} = 1 - (1 - q ^ (i + 1)) ^ n" proof - have "UNIV - {..i} = {i<..}" by auto then have "measure_pmf.prob (H\<^sub>N n) {i<..} = measure_pmf.prob (H\<^sub>N n) (space (measure_pmf (H\<^sub>N n)) - {..i})" by (auto) also have "\ = 1 - (measure_pmf.prob (geometric_pmf p) {..i}) ^ n" using assms by (subst measure_pmf.prob_compl) (auto simp add: prob_Max_IID_geometric_atMost) also have "\ = 1 - (1 - q ^ (i + 1)) ^ n" using assms unfolding q_def by (subst geometric_pmf_prob_atMost) auto finally show ?thesis by simp qed end (* context includes monad_normalisation *) end (* locale skip_list *) text \ An alternative definition of the expected value of a non-negative random variable \footnote{\url{https://en.wikipedia.org/w/index.php?title=Expected\_value&oldid=881384346\#Formula\_for\_non-negative\_random\_variables}} \ lemma expectation_prob_atLeast: assumes "(\i. measure_pmf.prob N {i..}) abs_summable_on {1..}" shows "measure_pmf.expectation N real = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" "integrable N real" proof - have "(\(x, y). pmf N y) abs_summable_on Sigma {Suc 0..} atLeast" using assms by (auto simp add: measure_pmf_conv_infsetsum abs_summable_on_Sigma_iff) then have summable: "(\(x, y). pmf N x) abs_summable_on Sigma {Suc 0..} (atLeastAtMost (Suc 0))" by (subst abs_summable_on_reindex_bij_betw[of "\(x,y). (y,x)", symmetric]) (auto intro!: bij_betw_imageI simp add: inj_on_def case_prod_beta) have "measure_pmf.expectation N real = (\\<^sub>ax. pmf N x *\<^sub>R real x)" by (auto simp add: infsetsum_def integral_density measure_pmf_eq_density) also have "\ = (\\<^sub>ax \ ({0} \ {Suc 0..}). pmf N x *\<^sub>R real x)" by (auto intro!: infsetsum_cong) also have "\ = (\\<^sub>ax\{Suc 0..}. pmf N x * real x)" proof - have "(\x. pmf N x *\<^sub>R real x) abs_summable_on {0} \ {Suc 0..}" using summable by (subst (asm) abs_summable_on_Sigma_iff) (auto simp add: mult.commute) then show ?thesis by (subst infsetsum_Un_Int) auto qed also have "\ = (\\<^sub>a(x, y)\Sigma {Suc 0..} (atLeastAtMost (Suc 0)). pmf N x)" using summable by (subst infsetsum_Sigma) (auto simp add: mult.commute) also have "\ = (\\<^sub>ax\Sigma {Suc 0..} atLeast. pmf N (snd x))" by (subst infsetsum_reindex_bij_betw[of "\(x,y). (y,x)", symmetric]) (auto intro!: bij_betw_imageI simp add: inj_on_def case_prod_beta) also have "\ = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" using assms by (subst infsetsum_Sigma) (auto simp add: measure_pmf_conv_infsetsum abs_summable_on_Sigma_iff infsetsum_Sigma') finally show "measure_pmf.expectation N real = infsetsum (\i. measure_pmf.prob N {i..}) {1..}" by simp have "(\x. pmf N x *\<^sub>R real x) abs_summable_on {0} \ {Suc 0..}" using summable by (subst (asm) abs_summable_on_Sigma_iff) (auto simp add: mult.commute) then have "(\x. pmf N x *\<^sub>R real x) abs_summable_on UNIV" by (simp add: atLeast_Suc) then have "integrable (count_space UNIV) (\x. pmf N x *\<^sub>R real x)" by (subst abs_summable_on_def[symmetric]) blast then show "integrable N real" by (subst measure_pmf_eq_density, subst integrable_density) auto qed text \ The expected height of a skip list has no closed-form expression but we can approximate it. We start by showing how we can calculate an infinite sum over the natural numbers with an integral over the positive reals and the floor function. \ lemma infsetsum_set_nn_integral_reals: assumes "f abs_summable_on UNIV" "\n. f n \ 0" shows "infsetsum f UNIV = set_nn_integral lborel {0::real..} (\x. f (nat (floor x)))" proof - have "x < 1 + (floor x)"for x::real by linarith then have "\n. real n \ x \ x < 1 + real n" if "x \ 0" for x using that of_nat_floor by (intro exI[of _ "nat (floor x)"]) auto then have "{0..} = (\n. {real n..\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel = + then have "(\\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel) = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f (nat \x\))\lborel)" by (auto simp add: disjoint_family_on_def nn_integral_disjoint_family) also have "\ = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f n)\lborel)" by(subst suminf_cong, rule nn_integral_cong_AE) (auto intro!: eventuallyI simp add: indicator_def floor_eq4) also have "\ = (\n. ennreal (f n))" by (auto intro!: suminf_cong simp add: nn_integral_cmult) also have "\ = infsetsum f {0..}" using assms suminf_ennreal2 abs_summable_on_nat_iff' summable_norm_cancel by (auto simp add: infsetsum_nat) finally show ?thesis by simp qed lemma nn_integral_nats_reals: - shows "(\\<^sup>+ i. ennreal (f i) \count_space UNIV) = \\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel" + shows "(\\<^sup>+ i. ennreal (f i) \count_space UNIV) = (\\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel)" proof - have "x < 1 + (floor x)"for x::real by linarith then have "\n. real n \ x \ x < 1 + real n" if "x \ 0" for x using that of_nat_floor by (intro exI[of _ "nat (floor x)"]) auto then have "{0..} = (\n. {real n..\<^sup>+x\{0::real..}. f (nat \x\)\lborel = + then have "(\\<^sup>+x\{0::real..}. f (nat \x\)\lborel) = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f (nat \x\))\lborel)" by (auto simp add: disjoint_family_on_def nn_integral_disjoint_family) also have "\ = (\n. \\<^sup>+x\{real n..<1 + real n}. ennreal (f n)\lborel)" by(subst suminf_cong,rule nn_integral_cong_AE) (auto intro!: eventuallyI simp add: indicator_def floor_eq4) also have "\ = (\n. ennreal (f n))" by (auto intro!: suminf_cong simp add: nn_integral_cmult) also have "\ = (\\<^sup>+ i. ennreal (f i) \count_space UNIV)" by (simp add: nn_integral_count_space_nat) finally show ?thesis by simp qed lemma nn_integral_floor_less_eq: assumes "\x y. x \ y \ f y \ f x" - shows "\\<^sup>+x\{0::real..}. ennreal (f x)\lborel \ \\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel" + shows "(\\<^sup>+x\{0::real..}. ennreal (f x)\lborel) \ (\\<^sup>+x\{0::real..}. ennreal (f (nat \x\))\lborel)" using assms by (auto simp add: indicator_def intro!: nn_integral_mono ennreal_leI) lemma nn_integral_finite_imp_abs_sumable_on: fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes "nn_integral (count_space A) (\x. norm (f x)) < \" shows "f abs_summable_on A" using assms unfolding abs_summable_on_def integrable_iff_bounded by auto lemma nn_integral_finite_imp_abs_sumable_on': assumes "nn_integral (count_space A) (\x. ennreal (f x)) < \" "\x. f x \ 0" shows "f abs_summable_on A" using assms unfolding abs_summable_on_def integrable_iff_bounded by auto text \ We now show that $\int_0^\infty 1 - (1 - q^x) ^ n\;dx = \frac{- H_n}{\ln q}$ if $0 < q < 1$. \ lemma harm_integral_x_raised_n: "set_integrable lborel {0::real..1} (\x. (\i\{..i\{..x. (\i\{..i\{..i\{..x. (\i\{..auto simp add: einterval_def\) moreover have "set_integrable lborel (einterval 0 1) (\x. (x ^ n))" proof - have "set_integrable lborel {0::real..1} (\x. (x ^ n))" by (rule borel_integrable_atLeastAtMost') (auto intro!: borel_integrable_atLeastAtMost' continuous_intros) then show ?thesis by (rule set_integrable_subset) (auto simp add: einterval_def) qed ultimately show ?thesis by (auto intro!: borel_integrable_atLeastAtMost' simp add: interval_lebesgue_integrable_def) qed also have "(LBINT x=0..1. x ^ n) = 1 / (1 + real n)" proof - - have "(LBINT x=0..1. x ^ n) = LBINT x. x ^ n * indicator {0..1} x " + have "(LBINT x=0..1. x ^ n) = (LBINT x. x ^ n * indicator {0..1} x)" proof - have "AE x in lborel. x ^ n * indicator {0..1} x = indicator (einterval 0 1) x * x ^ n" by(rule eventually_mono[OF eventually_conj[OF AE_lborel_singleton[of 1] AE_lborel_singleton[of 0]]]) (auto simp add: indicator_def einterval_def) then show ?thesis using integral_cong_AE unfolding interval_lebesgue_integral_def set_lebesgue_integral_def by (auto intro!: integral_cong_AE) qed then show ?thesis by (auto simp add: integral_power) qed finally show ?case using Suc by (auto simp add: harm_def inverse_eq_divide) qed (auto simp add: harm_def) qed lemma harm_integral_0_1_fraction: "set_integrable lborel {0::real..1} (\x. (1 - x ^ n) / (1 - x))" "(LBINT x = 0..1. ((1 - x ^ n) / (1 - x))) = harm n" proof - show "set_integrable lborel {0::real..1} (\x. (1 - x ^ n) / (1 - x))" proof - have "AE x\{0::real..1} in lborel. (1 - x ^ n) / (1 - x) = sum ((^) x) {..{0::real<..<1} in lborel. (1 - x ^ n) / (1 - x) = sum ((^) x) {.. {0<..<1}" shows "set_integrable lborel (einterval 0 \) (\x. (1 - (1 - q powr x) ^ n))" "(LBINT x=0..\. 1 - (1 - q powr x) ^ n) = - harm n / ln q" proof - have [simp]: "q powr (log q (1-x)) = 1 - x" if "x \ {0<..<1}" for x using that assms by (subst powr_log_cancel) auto have 1: "((ereal \ (\x. log q (1 - x)) \ real_of_ereal) \ 0) (at_right 0)" using assms unfolding zero_ereal_def ereal_tendsto_simps by (auto intro!: tendsto_eq_intros) have 2: "((ereal \ (\x. log q (1-x)) \ real_of_ereal) \ \) (at_left 1)" proof - have "filterlim ((-) 1) (at_right 0) (at_left (1::real))" by (intro filterlim_at_withinI eventually_at_leftI[of 0]) (auto intro!: tendsto_eq_intros) then have "LIM x at_left 1. - inverse (ln q) * - ln (1 - x) :> at_top" using assms by (intro filterlim_tendsto_pos_mult_at_top [OF tendsto_const]) (auto simp: filterlim_uminus_at_top intro!: filterlim_compose[OF ln_at_0]) then show ?thesis unfolding one_ereal_def ereal_tendsto_simps log_def by (simp add: field_simps) qed have 3: "set_integrable lborel (einterval 0 1) (\x. (1 - (1 - q powr (log q (1 - x))) ^ n) * (- 1 / (ln q * (1 - x))))" proof - have "set_integrable lborel (einterval 0 1) (\x. - (1 / ln q) * ((1 - x ^ n) / (1 - x)))" by(intro set_integrable_mult_right) (auto intro!: harm_integral_0_1_fraction intro: set_integrable_subset simp add: einterval_def) then show ?thesis by(subst set_integrable_cong_AE[where g="\x. - (1 / ln q) * ((1 - x ^ n) / (1 - x))"]) (auto intro!: eventuallyI simp add: einterval_def) qed have 4: "LBINT x=0..1. - ((1 - (1 - q powr log q (1 - x)) ^ n) / (ln q * (1 - x))) = - (harm n / ln q)" (is "?lhs = ?rhs") proof - have "?lhs = LBINT x=0..1. ((1 - x ^ n) / (1 - x)) * (- 1 / ln q)" using assms by (intro interval_integral_cong_AE) (auto intro!: eventuallyI simp add: max_def einterval_def field_simps) also have "\ = harm n * (-1 / ln q)" using harm_integral_0_1_fraction by (subst interval_lebesgue_integral_mult_left) auto finally show ?thesis by auto qed note sub = interval_integral_substitution_nonneg [where f = "(\x. (1 - (1 - q powr x) ^ n))" and g="(\x. log q (1-x))" and g'="(\x. - 1 / (ln q * (1 - x)))" and a = 0 and b = 1] show "set_integrable lborel (einterval 0 \) (\x. (1 - (1 - q powr x) ^ n))" using assms 1 2 3 4 by (intro sub) (auto intro!: derivative_eq_intros mult_nonneg_nonpos2 tendsto_intros power_le_one) show "(LBINT x=0..\. 1 - (1 - q powr x) ^ n) = - harm n / ln q" using assms 1 2 3 4 by (subst sub) (auto intro!: derivative_eq_intros mult_nonneg_nonpos2 tendsto_intros power_le_one) qed lemma one_minus_one_minus_q_x_n_nn_integral: fixes q::real assumes "q \ {0<..<1}" shows "set_nn_integral lborel {0..} (\x. (1 - (1 - q powr x) ^ n)) = LBINT x=0..\. 1 - (1 - q powr x) ^ n" proof - have "set_nn_integral lborel {0..} (\x. (1 - (1 - q powr x) ^ n)) = nn_integral lborel (\x. indicator (einterval 0 \) x * (1 - (1 - q powr x) ^ n))" using assms by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]]) (auto simp add: indicator_def einterval_def) also have "\ = ennreal (LBINT x. indicator (einterval 0 \) x * (1 - (1 - q powr x) ^ n))" using one_minus_one_minus_q_x_n_integral assms by(intro nn_integral_eq_integral) (auto simp add: indicator_def einterval_def set_integrable_def intro!: eventuallyI power_le_one powr_le1) finally show ?thesis by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def) qed text \ We can now derive bounds for the expected height. \ context random_skip_list begin definition EH\<^sub>N where "EH\<^sub>N n = measure_pmf.expectation (H\<^sub>N n) real" lemma EH\<^sub>N_bounds': fixes n::nat assumes "p \ {0<..<1}" "0 < n" shows "- harm n / ln q - 1 \ EH\<^sub>N n" "EH\<^sub>N n \ - harm n / ln q" "integrable (H\<^sub>N n) real" proof - define f where "f = (\x. 1 - (1 - q ^ x) ^ n)" define f' where "f' = (\x. 1 - (1 - q powr x) ^ n)" have q: "q \ {0<..<1}" unfolding q_def using assms by auto have f_descending: "f y \ f x" if "x \ y" for x y unfolding f_def using that q by (auto intro!: power_mono simp add: power_decreasing power_le_one_iff) have f'_descending: "f' y \ f' x" if "x \ y" "0 \ x" for x y unfolding f'_def using that q by (auto intro!: power_mono simp add: ln_powr powr_def mult_nonneg_nonpos) have [simp]: "harm n / ln q <= 0" using harm_nonneg ln_ge_zero_imp_ge_one q by (intro divide_nonneg_neg) auto have f_nn_integral_harm: "- harm n / ln q \ \\<^sup>+ x. (f x) \count_space UNIV" "(\\<^sup>+ i. f (i + 1) \count_space UNIV) \ - harm n / ln q" proof - have "(\\<^sup>+ i. f (i + 1) \count_space UNIV) = (\\<^sup>+x\{0::real..}. (f (nat \x\ + 1))\lborel)" using nn_integral_nats_reals by auto - also have "\ = \\<^sup>+x\{0::real..}. ennreal (f' (nat \x\ + 1))\lborel" + also have "\ = (\\<^sup>+x\{0::real..}. ennreal (f' (nat \x\ + 1))\lborel)" proof - have "0 \ x \ (1 - q * q ^ nat \x\) ^ n = (1 - q powr (1 + real_of_int \x\)) ^ n" for x::real using q by (subst powr_realpow [symmetric]) (auto simp: powr_add) then show ?thesis unfolding f_def f'_def using q by (auto intro!: nn_integral_cong ennreal_cong simp add: powr_real_of_int indicator_def) qed also have "\ \ set_nn_integral lborel {0..} f'" proof - have "x \ 1 + real_of_int \x\" for x by linarith then show ?thesis by (auto simp add: indicator_def intro!: f'_descending nn_integral_mono ennreal_leI) qed also have harm_integral_f': "\ = - harm n / ln q" unfolding f'_def using q by (auto intro!: ennreal_cong simp add: one_minus_one_minus_q_x_n_nn_integral one_minus_one_minus_q_x_n_integral) finally show "(\\<^sup>+ i. f (i + 1) \count_space UNIV) \ - harm n / ln q" by simp note harm_integral_f'[symmetric] - also have "set_nn_integral lborel {0..} f' \ \\<^sup>+x\{0::real..}. f' (nat \x\)\lborel" + also have "set_nn_integral lborel {0..} f' \ (\\<^sup>+x\{0::real..}. f' (nat \x\)\lborel)" using assms f'_descending by (auto simp add: indicator_def intro!: nn_integral_mono ennreal_leI) - also have "\ = \\<^sup>+x\{0::real..}. f (nat \x\)\lborel" + also have "\ = (\\<^sup>+x\{0::real..}. f (nat \x\)\lborel)" unfolding f_def f'_def using q by (auto intro!: nn_integral_cong ennreal_cong simp add: powr_real_of_int indicator_def) also have "\ = (\\<^sup>+ x. f x \count_space UNIV)" using nn_integral_nats_reals by auto finally show "- harm n / ln q \ \\<^sup>+ x. f x \count_space UNIV" by simp qed then have f1_abs_summable_on: "(\i. f (i + 1)) abs_summable_on UNIV" unfolding f_def using q by (intro nn_integral_finite_imp_abs_sumable_on') (auto simp add: f_def le_less_trans intro!: power_le_one mult_le_one) then have f_abs_summable_on: "f abs_summable_on {1..}" using Suc_le_lessD greaterThan_0 by (subst abs_summable_on_reindex_bij_betw[symmetric, where g="\x. x + 1" and A="UNIV"]) auto also have "(f abs_summable_on {1..}) = ((\x. measure_pmf.prob (H\<^sub>N n) {x..}) abs_summable_on {1..})" proof - have "((\x. measure_pmf.prob (H\<^sub>N n) {x..}) abs_summable_on {1..}) = ((\x. measure_pmf.prob (H\<^sub>N n) {x - 1<..}) abs_summable_on {1..})" by (auto intro!: measure_prob_cong_0 abs_summable_on_cong) also have "\ = (f abs_summable_on {1..})" using assms by (intro abs_summable_on_cong) (auto simp add: f_def prob_Max_IID_geometric_greaterThan) finally show ?thesis by simp qed finally have EH\<^sub>N_sum: "EH\<^sub>N n = (\\<^sub>ai\{1..}. measure_pmf.prob (H\<^sub>N n) {i..})" "integrable (measure_pmf (H\<^sub>N n)) real" unfolding EH\<^sub>N_def using expectation_prob_atLeast by auto then show "integrable (measure_pmf (H\<^sub>N n)) real" by simp have EH\<^sub>N_sum': "EH\<^sub>N n = infsetsum f {1..}" proof - have "EH\<^sub>N n = (\\<^sub>ak\{1..}. measure_pmf.prob (H\<^sub>N n) {k - 1<..})" unfolding EH\<^sub>N_sum by (auto intro!: measure_prob_cong_0 infsetsum_cong) also have "\ = infsetsum f {1..}" using assms by (intro infsetsum_cong) (auto simp add: f_def prob_Max_IID_geometric_greaterThan) finally show ?thesis by simp qed also have "\ = (\\<^sub>ak. f (k + 1))" using Suc_le_lessD greaterThan_0 by (subst infsetsum_reindex_bij_betw[symmetric, where g="\x. x + 1" and A="UNIV"]) auto also have "ennreal \ = (\\<^sup>+x\{0::real..}. f (nat \x\ + 1)\lborel)" using f1_abs_summable_on q by (intro infsetsum_set_nn_integral_reals) (auto simp add: f_def mult_le_one power_le_one) also have "\ = (\\<^sup>+ i. f (i + 1) \count_space UNIV)" using nn_integral_nats_reals by auto also have "\ \ - harm n / ln q" using f_nn_integral_harm by auto finally show "EH\<^sub>N n \ - harm n / ln q" by (subst (asm) ennreal_le_iff) (auto) have "EH\<^sub>N n + 1 = (\\<^sub>ax\{Suc 0..}. f x) + (\\<^sub>ax\{0}. f x)" using assms by (subst EH\<^sub>N_sum') (auto simp add: f_def) also have "\ = infsetsum f UNIV" using f_abs_summable_on by (subst infsetsum_Un_disjoint[symmetric]) (auto intro!: infsetsum_cong) also have "\ = (\\<^sup>+x\{0::real..}. f (nat \x\)\lborel)" proof - have "f abs_summable_on ({0} \ {1..})" using f_abs_summable_on by (intro abs_summable_on_union) (auto) also have "{0::nat} \ {1..} = UNIV" by auto finally show ?thesis using q by (intro infsetsum_set_nn_integral_reals) (auto simp add: f_def mult_le_one power_le_one) qed also have "\ = (\\<^sup>+ x. f x \count_space UNIV)" using nn_integral_nats_reals by auto also have "... \ - harm n / ln q" using f_nn_integral_harm by auto finally have "- harm n / ln q \ EH\<^sub>N n + 1" by (subst (asm) ennreal_le_iff) (auto simp add: EH\<^sub>N_def) then show "- harm n / ln q - 1 \ EH\<^sub>N n" by simp qed theorem EH\<^sub>N_bounds: fixes n::nat assumes "p \ {0<..<1}" shows "- harm n / ln q - 1 \ EH\<^sub>N n" "EH\<^sub>N n \ - harm n / ln q" "integrable (H\<^sub>N n) real" proof - show "- harm n / ln q - 1 \ EH\<^sub>N n" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def SL_def harm_expand) show "EH\<^sub>N n \ - harm n / ln q" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def SL_def harm_expand) show "integrable (H\<^sub>N n) real" using assms EH\<^sub>N_bounds' by (cases "n = 0") (auto simp add: H\<^sub>N_def H_def SL_def intro!: integrable_measure_pmf_finite) qed end (* context random_skip_list *) subsection \Expected Length of Search Path\ text \ Let @{term "A::'a::linorder set"} and @{term "f::'a \ nat"} where f is an abstract description of a skip list (assign each value its maximum level). steps A f s u l starts on the rightmost element on level s in the skip lists. If possible it moves up, if not it moves to the left. For every step up it adds cost u and for every step to the left it adds cost l. steps A f 0 1 1 therefore walks from the bottom right corner of a skip list to the top left corner of a skip list and counts all steps. \ \ \NOTE: You could also define steps with lsteps and then prove that the following recursive definition holds\ function steps :: "'a :: linorder set \ ('a \ nat) \ nat \ nat \ nat \ nat" where "steps A f l up left = (if A = {} \ infinite A then 0 else (let m = Max A in (if f m < l then steps (A - {m}) f l up left else (if f m > l then up + steps A f (l + 1) up left else left + steps (A - {m}) f l up left))))" by pat_completeness auto termination proof (relation "(\(A,f,l,a,b). card A) <*mlex*> (\(A,f,l,a,b). Max (f ` A) - l) <*mlex*> {}", goal_cases) case 1 then show ?case by(intro wf_mlex wf_empty) next case 2 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) next case (3 A f l a b x) then have "Max (f ` A) - Suc l < Max (f ` A) - l" by (meson Max_gr_iff Max_in diff_less_mono2 finite_imageI imageI image_is_empty lessI) with 3 have "((A, f, l + 1, a, b), A, f, l, a, b) \ (\(A, f, l, a, b). Max (f ` A) - l) <*mlex*> {}" by (intro mlex_less) (auto) with 3 show ?case apply - apply(rule mlex_leq) by auto next case 4 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) qed declare steps.simps[simp del] text \ lsteps is similar to steps but is using lists instead of sets. This makes the proofs where we use induction easier. \ function lsteps :: "'a list \ ('a \ nat) \ nat \ nat \ nat \ nat" where "lsteps [] f l up left = 0" | "lsteps (x#xs) f l up left = (if f x < l then lsteps xs f l up left else (if f x > l then up + lsteps (x#xs) f (l + 1) up left else left + lsteps xs f l up left))" by pat_completeness auto termination proof (relation "(\(xs,f,l,a,b). length xs) <*mlex*> (\(xs,f,l,a,b). Max (f ` set xs) - l) <*mlex*> {}", goal_cases) case 1 then show ?case by(intro wf_mlex wf_empty) next case 2 then show ?case by (auto intro: mlex_less simp: card_gt_0_iff) next case (3 n f l a b) show ?case by (rule mlex_leq) (use 3 in \auto intro: mlex_less mlex_leq intro!: diff_less_mono2 simp add: Max_gr_iff\) next case 4 then show ?case by (intro mlex_less) (auto simp: card_gt_0_iff) qed declare lsteps.simps(2)[simp del] lemma steps_empty [simp]: "steps {} f l up left = 0" by (simp add: steps.simps) lemma steps_lsteps: "steps A f l u v = lsteps (rev (sorted_list_of_set A)) f l u v" proof (cases "finite A \ A \ {}") case True then show ?thesis proof(induction "(rev (sorted_list_of_set A))" f l u v arbitrary: A rule: lsteps.induct) case (2 y ys f l u v A) then have y_ys: "y = Max A" "ys = rev (sorted_list_of_set (A - {y}))" by (auto simp add: sorted_list_of_set_Max_snoc) consider (a) "l < f y" | (b) "f y < l" | (c) "f y = l" by fastforce then have "steps A f l u v = lsteps (y#ys) f l u v" proof cases case a then show ?thesis by (subst steps.simps, subst lsteps.simps) (use y_ys 2 in auto) next case b then show ?thesis using y_ys 2(1) by (cases "ys = []") (auto simp add: steps.simps lsteps.simps) next case c then have "steps (A - {Max A}) f l u v = lsteps (rev (sorted_list_of_set (A - {Max A}))) f l u v" by (cases "A = {Max A}") (use y_ys 2 in \auto intro!: 2(3) simp add: steps.simps\) then show ?thesis by (subst steps.simps, subst lsteps.simps) (use y_ys 2 in auto) qed then show ?case using 2 by simp qed (auto simp add: steps.simps) qed (auto simp add: steps.simps) lemma lsteps_comp_map: "lsteps zs (f \ g) l u v = lsteps (map g zs) f l u v" by (induction zs "f \ g" l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_image: assumes "finite A" "mono_on A g" "inj_on g A" shows "steps A (f \ g) l u v = steps (g ` A) f l u v" proof - have "(sorted_list_of_set (g ` A)) = map g (sorted_list_of_set A)" using sorted_list_of_set_image assms by auto also have "rev \ = map g (rev (sorted_list_of_set A))" using rev_map by auto finally show ?thesis by (simp add: steps_lsteps lsteps_comp_map) qed lemma lsteps_cong: assumes "ys = xs" "\x. x \ set xs \ f x = g x" "l = l'" shows "lsteps xs f l u v = lsteps ys g l' u v" using assms proof (induction xs f l u v arbitrary: ys l' rule: lsteps.induct) case (2 x xs f l up left) then show ?case by (subst \ys = x # xs\, subst lsteps.simps, subst (2) lsteps.simps) auto qed (auto) lemma steps_cong: assumes "A = B" "\x. x \ A \ f x = g x" "l = l'" shows "steps A f l u v = steps B g l' u v" using assms by (cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps intro!: lsteps_cong) lemma lsteps_f_add': shows "lsteps xs f l u v = lsteps xs (\x. f x + m) (l + m) u v" by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_f_add': shows "steps A f l u v = steps A (\x. f x + m) (l + m) u v" by (cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps intro!: lsteps_f_add') lemma lsteps_smaller_set: assumes "m \ l" shows "lsteps xs f l u v = lsteps [x \ xs. m \ f x] f l u v" using assms by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma steps_smaller_set: assumes "finite A" "m \ l" shows "steps A f l u v = steps {x\A. f x \ m} f l u v" using assms by(cases "A = {} \ infinite A") (auto simp add: steps_lsteps steps.simps rev_filter sorted_list_of_set_filter intro!: lsteps_smaller_set) lemma lsteps_level_greater_fun_image: assumes "\x. x \ set xs \ f x < l" shows "lsteps xs f l u v = 0" using assms by (induction xs f l u v rule: lsteps.induct) (auto simp add: lsteps.simps) lemma lsteps_smaller_card_Max_fun': assumes "\x \ set xs. l \ f x" shows "lsteps xs f l u v + l * u \ v * length xs + u * Max ((f ` (set xs)) \ {0})" using assms proof (induction xs f l u v rule: lsteps.induct) case (1 f l up left) then show ?case by (simp) next case (2 x xs f l up left) consider "l = f x" "\y\set xs. l \ f y" | "f x = l" "\ (\y\set xs. l \ f y)" | "f x < l" | "l < f x" by fastforce then show ?case proof cases assume a: "l = f x" "\y\set xs. l \ f y" have "lsteps (x # xs) f l up left + l * up = lsteps xs f l up left + f x * up + left" using a by (auto simp add: lsteps.simps) also have "lsteps xs f l up left + f x * up \ left * length xs + up * Max (f ` set xs \ {0})" using a 2 by blast also have "up * Max (f ` set xs \ {0}) \ up * Max (insert (f x) (f ` set xs))" by simp finally show ?case by auto next assume a: "f x = l" "\ (\y\set xs. l \ f y)" have "lsteps (x # xs) f l up left + l * up = lsteps xs f l up left + f x * up + left" using a by (auto simp add: lsteps.simps) also have "lsteps xs f l up left = 0" using a by (subst lsteps_level_greater_fun_image) auto also have "f x * up \ up * Max (insert (f x) (f ` set xs))" by simp finally show ?case by simp next assume a: "f x < l" then have "lsteps (x # xs) f l up left = lsteps xs f l up left" by (auto simp add: lsteps.simps) also have "\ + l * up \ left * length (x # xs) + up * Max (insert 0 (f ` set xs))" using a 2 by auto also have "Max (insert 0 (f ` set xs)) \ Max (f ` set (x # xs) \ {0})" by simp finally show ?case by simp next assume "f x > l" then show ?case using 2 by (subst lsteps.simps) auto qed qed lemma steps_smaller_card_Max_fun': assumes "finite A" "\x\A. l \ f x" shows "steps A f l up left + l * up \ left * card A + up * Max\<^sub>0 (f ` A)" proof - let ?xs = "rev (sorted_list_of_set A)" have "steps A f l up left = lsteps (rev (sorted_list_of_set A)) f l up left" using steps_lsteps by blast also have "\ + l * up \ left * length ?xs + up * Max (f ` set ?xs \ {0})" using assms by (intro lsteps_smaller_card_Max_fun') auto also have "left * length ?xs = left * card A" using assms sorted_list_of_set_length by (auto) also have "set ?xs = A" using assms by (auto) finally show ?thesis by simp qed lemma lsteps_height: assumes "\x \ set xs. l \ f x" shows "lsteps xs f l up 0 + up * l = up * Max\<^sub>0 (f ` (set xs))" using assms proof (induction xs f l up "0::nat" rule: lsteps.induct) case (2 x xs f l up) consider "l = f x" "\y\set xs. l \ f y" | "f x = l" "\ (\y\set xs. l \ f y)" | "f x < l" | "l < f x" by fastforce then show ?case proof cases assume 0: "l = f x" "\y\set xs. l \ f y" then have 1: "set xs \ {}" using 2 by auto then have "\xa\set xs. f x \ f xa" using 0 2 by force then have "f x \ Max (f ` set xs)" using 0 2 by (subst Max_ge_iff) auto then have "max (f x) (Max (f ` set xs)) = (Max (f ` set xs))" using 0 2 by (auto intro!: simp add: max_def) then show ?case using 0 1 2 by (subst lsteps.simps) (auto) next assume 0: "f x = l" "\ (\y\set xs. l \ f y)" then have "Max (insert l (f ` set xs)) = l" by (intro Max_eqI) (auto) moreover have "lsteps xs f l up 0 = 0" using 0 by (subst lsteps_level_greater_fun_image) auto ultimately show ?case using 0 by (subst lsteps.simps) auto next assume 0: "f x < l" then have 1: "set xs \ {}" using 2 by auto then have "\xa\set xs. f x \ f xa" using 0 2 by force then have " f x \ Max (f ` set xs)" using 0 2 by (subst Max_ge_iff) auto then have "max (f x) (Max (f ` set xs)) = Max (f ` set xs)" using 0 2 by (auto intro!: simp add: max_def) then show ?case using 0 1 2 by (subst lsteps.simps) (auto) next assume "f x > l" then show ?case using 2 by (subst lsteps.simps) auto qed qed (simp) lemma steps_height: assumes "finite A" shows "steps A f 0 up 0 = up * Max\<^sub>0 (f ` A)" proof - have "steps A f 0 up 0 = lsteps (rev (sorted_list_of_set A)) f 0 up 0 + up * 0" by (subst steps_lsteps) simp also have "\ = up * Max (f ` A \ {0})" if "A \ {}" using assms that by (subst lsteps_height) auto finally show ?thesis using assms by (cases "A = {}") (auto) qed context random_skip_list begin text \ We can now define the pmf describing the length of the search path in a skip list. Like the height it only depends on the number of elements in the skip list's underlying set. \ definition R where "R A u l = map_pmf (\f. steps A f 0 u l) (SL A)" definition R\<^sub>N :: "nat \ nat \ nat \ nat pmf" where "R\<^sub>N n u l = R {..N_alt_def: "R\<^sub>N n u l = map_pmf (\f. steps {..N n)" unfolding SL\<^sub>N_def R\<^sub>N_def R_def by simp context includes monad_normalisation begin lemma R_R\<^sub>N: assumes "finite A" "p \ {0..1}" shows "R A u l = R\<^sub>N (card A) u l" proof - let ?steps = "\A f. steps A f 0 u l" let ?f' = "bij_mono_map_set_to_nat A" have "R A u l = SL A \ (\f. return_pmf (?steps A f))" unfolding R_def map_pmf_def by simp also have "\ = SL\<^sub>N (card A) \ (\f. return_pmf (?steps A (f \ ?f')))" proof - have "?f' x \ {.. A" for x using that unfolding bij_mono_map_set_to_nat_def by (auto) then show ?thesis using assms bij_mono_map_set_to_nat unfolding SL_def SL\<^sub>N_def by (subst Pi_pmf_bij_betw[of _ ?f' "{.. = SL\<^sub>N (card A) \ (\f. return_pmf (?steps {..N_def R_def SL\<^sub>N_def SL_def by (simp add: map_pmf_def) qed text \ @{const R\<^sub>N} fulfills a recurrence relation. If we move up or to the left the ``remaining'' length of the search path is again a slightly different probability distribution over the length. \ lemma R\<^sub>N_recurrence: assumes "0 < n" "p \ {0<..1}" shows "R\<^sub>N n u l = do { b \ bernoulli_pmf p; if b then \ \leftwards\ map_pmf (\n. n + l) (R\<^sub>N (n - 1) u l) else do { \ \upwards\ m \ binomial_pmf (n - 1) (1 - p); map_pmf (\n. n + u) (R\<^sub>N (m + 1) u l) } }" proof - define B where "B = (\b. insert (n-1) {x \ {.. b x})" have "R\<^sub>N n u l = map_pmf (\f. steps {..N n)" by (auto simp add: R\<^sub>N_def R_def SL\<^sub>N_def) also have "\ = map_pmf (\f. steps {..(y, f). f(n-1 := y)) (pair_pmf (geometric_pmf p) (SL\<^sub>N (n - 1))))" proof - have "{.._. geometric_pmf p)) = map_pmf (\(y, f). f(n - 1 := y)) (pair_pmf (geometric_pmf p) (Pi_pmf {.._. geometric_pmf p)))" using assms by (subst Pi_pmf_insert[of "{.._. geometric_pmf p", symmetric]) (auto) then show ?thesis by (simp add: SL\<^sub>N_def SL_def) qed also have "\ = do { g \ geometric_pmf p; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. = do { b \ bernoulli_pmf p; g \ if b then return_pmf 0 else map_pmf Suc (geometric_pmf p); f \ SL\<^sub>N (n - 1); return_pmf (steps {.. = do { b \ bernoulli_pmf p; if b then do { g \ return_pmf 0; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. map_pmf Suc (geometric_pmf p); f \ SL\<^sub>N (n - 1); return_pmf (steps {.. return_pmf 0; f \ SL\<^sub>N (n - 1); return_pmf (steps {.. SL\<^sub>N (n - 1); return_pmf (steps {.. = map_pmf (\n. n + l) (map_pmf (\f. steps {..N (n - 1)))" proof - have I: "{.. = map_pmf (\n. n + l) (R\<^sub>N (n - 1) u l)" unfolding R\<^sub>N_def R_def SL\<^sub>N_def by simp also have "map_pmf Suc (geometric_pmf p) \ (\g. SL\<^sub>N (n - 1) \ (\f. return_pmf (steps {.._. bernoulli_pmf p) \ (\b. map_pmf Suc (geometric_pmf p) \ (\g. Pi_pmf {x \ {.. b x} 0 (\_. map_pmf Suc (geometric_pmf p)) \ (\f. return_pmf (steps {..N_def SL_def by (subst Pi_pmf_geometric_filter) (auto) also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ Pi_pmf (insert (n-1) {x \ {.. b x}) 0 (\_. map_pmf Suc (geometric_pmf p)); return_pmf (steps {.. = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ Pi_pmf (B b) 1 (\_. map_pmf Suc (geometric_pmf p)); return_pmf (steps {..x. if x \ (B b) then f x else 0) 0 u l)}" by (subst Pi_pmf_default_swap[symmetric, of _ _ _ 1]) (auto simp add: map_pmf_def B_def) also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); f \ SL (B b); return_pmf (steps {..x. if x \ (B b) then Suc (f x) else 0) 0 u l)}" proof - have *: "(Suc \ f) x = Suc (f x)" for x and f::"nat \ nat" by simp have "(\f. return_pmf (steps {..x. if x \ B b then (Suc \ f) x else 0) 0 u l)) = (\f. return_pmf (steps {..x. if x \ B b then Suc (f x) else 0) 0 u l))" for b by (subst *) (simp) then show ?thesis by (subst Pi_pmf_map[of _ _ 0]) (auto simp add: map_pmf_def B_def SL_def) qed also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf p); r \ R (B b) u l; return_pmf (u + r)}" proof - have "steps {..x. if x \ B b then Suc (f x) else 0) 0 u l = u + steps (B b) f 0 u l" for f b proof - have "Max {..x. if x \ B b then Suc (f x) else 0) 0 u l = u + (steps {..x. if x \ (B b) then Suc (f x) else 0) 1 u l)" unfolding B_def using assms by (subst steps.simps) (auto simp add: Let_def) also have "steps {..x. if x \ (B b) then Suc (f x) else 0) 1 u l = steps (B b) (\x. if x \ (B b) then Suc (f x) else 0) 1 u l" proof - have "{x \ {.. (if x \ B b then Suc (f x) else 0)} = B b" using assms unfolding B_def by force then show ?thesis by (subst steps_smaller_set[of _ 1]) auto qed also have "\ = steps (B b) (\x. f x + 1) 1 u l" by (rule steps_cong) (auto) also have "\ = steps (B b) f 0 u l" by (subst (2) steps_f_add'[of _ _ _ _ _ 1]) simp finally show ?thesis by auto qed then show ?thesis by (simp add: R_def map_pmf_def) qed also have "\ = do { b \ Pi_pmf {.._. bernoulli_pmf (1 - p)); let m = 1 + card {x. x < n - 1 \ b x}; r \ R {.. b x}) = (Suc (card {x. x < n - 1 \ b x}))" for b using assms by (auto simp add: card_insert_if) have "Pi_pmf {.._. bernoulli_pmf p) = Pi_pmf {.._. map_pmf Not (bernoulli_pmf (1 - p)))" using assms by (subst bernoulli_pmf_Not) auto also have "\ = map_pmf ((\) Not) (Pi_pmf {.._. bernoulli_pmf (1 - p)))" using assms by (subst Pi_pmf_map[of _ _ False]) auto finally show ?thesis unfolding B_def using assms * by (subst R_R\<^sub>N) (auto simp add: R_R\<^sub>N map_pmf_def) qed also have "\ = binomial_pmf (n - 1) (1 - p) \ (\m. map_pmf (\n. n + u) (R\<^sub>N (m + 1) u l))" using assms by (subst binomial_pmf_altdef'[where A = "{..N_def R_def SL_def map_pmf_def ac_simps) finally show ?thesis by simp qed end (* context includes monad_normalisation *) text \ The expected height and length of search path defined as non-negative integral. It's easier to prove the recurrence relation of the expected length of the search path using non-negative integrals. \ definition NH\<^sub>N where "NH\<^sub>N n = nn_integral (H\<^sub>N n) real" definition NR\<^sub>N where "NR\<^sub>N n u l = nn_integral (R\<^sub>N n u l) real" lemma NH\<^sub>N_EH\<^sub>N: assumes "p \ {0<..<1}" shows "NH\<^sub>N n = EH\<^sub>N n" using assms EH\<^sub>N_bounds unfolding EH\<^sub>N_def NH\<^sub>N_def by (subst nn_integral_eq_integral) (auto) lemma R\<^sub>N_0 [simp]: "R\<^sub>N 0 u l = return_pmf 0" unfolding R\<^sub>N_def R_def SL_def by (auto simp add: steps.simps) lemma NR\<^sub>N_bounds: fixes u l::nat shows "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" proof - have "NR\<^sub>N n u l = \\<^sup>+ x. x \measure_pmf (R\<^sub>N n u l)" unfolding NR\<^sub>N_def R\<^sub>N_alt_def by (simp add: ennreal_of_nat_eq_real_of_nat) also have "\ \ \\<^sup>+ x. x \(measure_pmf (map_pmf (\f. l * n + u * Max\<^sub>0 (f ` {..N n)))" using of_nat_mono[OF steps_smaller_card_Max_fun'[of "{..N_alt_def by (cases "n = 0") (auto intro!: nn_integral_mono) also have "\ = l * n + u * NH\<^sub>N n" unfolding NH\<^sub>N_def H\<^sub>N_def H_def SL\<^sub>N_def by (auto simp add: nn_integral_add nn_integral_cmult ennreal_of_nat_eq_real_of_nat ennreal_mult) finally show "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" by simp qed lemma NR\<^sub>N_recurrence: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u l = (p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (pmf (binomial_pmf (n - 1) q) k)))) / (1 - (q ^ n))" proof - define B where "B = (\n k. pmf (binomial_pmf n q) k)" have q: "q \ {0<..<1}" using assms unfolding q_def by auto then have "q ^ n < 1" using assms power_Suc_less_one by (induction n) (auto) then have qn: "q ^ n \ {0<..<1}" using assms q by (auto) have "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * (u + \\<^sup>+ k. NR\<^sub>N (k + 1) u l \measure_pmf (binomial_pmf (n - 1) q))" using assms unfolding NR\<^sub>N_def by(subst R\<^sub>N_recurrence) (auto simp add: field_simps nn_integral_add q_def ennreal_of_nat_eq_real_of_nat) also have "(\\<^sup>+ m. NR\<^sub>N (m + 1) u l \measure_pmf (binomial_pmf (n - 1) q)) = (\k\n - 1. NR\<^sub>N (k + 1) u l * B (n - 1) k)" using assms unfolding B_def q_def by (auto simp add: nn_integral_measure_pmf_finite) also have "\ = (\k\{.. {n - 1}. NR\<^sub>N (k + 1) u l * B (n - 1) k)" by (rule sum.cong) (auto) also have "\ = (\kN (k + 1) u l * B (n - 1) k) + NR\<^sub>N n u l * q ^ (n - 1)" unfolding B_def q_def using assms by (subst sum.union_disjoint) (auto) finally have "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * ((\kN (k + 1) u l * B (n - 1) k) + u) + NR\<^sub>N n u l * (q ^ (n - 1)) * q" using assms by (auto simp add: field_simps numerals) also have "NR\<^sub>N n u l * (q ^ (n - 1)) * q = (q ^ n) * NR\<^sub>N n u l" using q power_minus_mult[of _ q] assms by (subst mult_ac, subst ennreal_mult[symmetric], auto simp add: mult_ac) finally have 1: "NR\<^sub>N n u l = p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (B (n - 1) k))) + (q ^ n) * NR\<^sub>N n u l " by (simp add: add_ac) have "x - z = y" if "x = y + z" "z \ \" for x y z::ennreal using that by (subst that) (auto) have "NR\<^sub>N n u l \ l * n + u * NH\<^sub>N n" using NR\<^sub>N_bounds by (auto simp add: ennreal_of_nat_eq_real_of_nat) also have "NH\<^sub>N n = EH\<^sub>N n" using assms NH\<^sub>N_EH\<^sub>N by auto also have "(l * n) + u * ennreal (EH\<^sub>N n) < \" by (simp add: ennreal_mult_less_top of_nat_less_top) finally have 3: "NR\<^sub>N n u l \ \" by simp have 2: "x = y / (1 - a)" if "x = y + a * x" and t: "x \ \" "a \ {0<..<1}" for x y::ennreal and a::real proof - have "y = x - a * x" using t by (subst that) (auto simp add: ennreal_mult_eq_top_iff) also have "\ = x * (ennreal 1 - ennreal a)" using that by (auto simp add: mult_ac ennreal_right_diff_distrib) also have "ennreal 1 - ennreal a = ennreal (1 - a)" using that by (subst ennreal_minus) (auto) also have "x * (1 - a) / (1 - a) = x" using that ennreal_minus_eq_0 not_less by (subst mult_divide_eq_ennreal) auto finally show ?thesis by simp qed have "NR\<^sub>N n u l = (p * (l + NR\<^sub>N (n - 1) u l) + q * (u + (\kN (k + 1) u l * (B (n - 1) k)))) / (1 - (q ^ n))" using 1 3 assms qn by (intro 2) auto then show ?thesis unfolding B_def by simp qed lemma NR\<^sub>n_NH\<^sub>N: "NR\<^sub>N n u 0 = u * NH\<^sub>N n" proof - have "NR\<^sub>N n u 0 = \\<^sup>+ f. steps {..measure_pmf (SL\<^sub>N n)" unfolding NR\<^sub>N_def R\<^sub>N_alt_def by (auto simp add: ennreal_of_nat_eq_real_of_nat) also have "\ = \\<^sup>+ f. of_nat u * of_nat (Max\<^sub>0 (f ` {..measure_pmf (SL\<^sub>N n)" by (intro nn_integral_cong) (auto simp add: steps_height) also have "\ = u * NH\<^sub>N n" by (auto simp add: NH\<^sub>N_def H\<^sub>N_def H_def SL\<^sub>N_def ennreal_of_nat_eq_real_of_nat nn_integral_cmult) finally show ?thesis by simp qed lemma NR\<^sub>N_recurrence': assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u l = (p * l + p * NR\<^sub>N (n - 1) u l + q * u + q * (\kN (k + 1) u l * (pmf (binomial_pmf (n - 1) q) k))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (auto simp add: field_simps ennreal_of_nat_eq_real_of_nat ennreal_mult' ennreal_mult'') lemma NR\<^sub>N_l_0: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n u 0 = (p * NR\<^sub>N (n - 1) u 0 + q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (n - 1) q) k)))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (simp) lemma NR\<^sub>N_u_0: assumes "0 < n" "p \ {0<..<1}" shows "NR\<^sub>N n 0 l = (p * (l + NR\<^sub>N (n - 1) 0 l) + q * (\kN (k + 1) 0 l * (pmf (binomial_pmf (n - 1) q) k))) / (1 - (q ^ n))" unfolding NR\<^sub>N_recurrence[OF assms] by (simp) lemma NR\<^sub>N_0[simp]: "NR\<^sub>N 0 u l = 0" unfolding NR\<^sub>N_def R\<^sub>N_def R_def by (auto) lemma NR\<^sub>N_1: assumes "p \ {0<..<1}" shows "NR\<^sub>N 1 u l = (u * q + l * p) / p" proof - have "NR\<^sub>N 1 u l = (ennreal p * of_nat l + ennreal q * of_nat u) / ennreal (1 - q)" using assms by (subst NR\<^sub>N_recurrence) auto also have "(ennreal p * of_nat l + ennreal q * of_nat u) = (u * q + l * p)" using assms q_def by (subst ennreal_plus) (auto simp add: field_simps ennreal_mult' ennreal_of_nat_eq_real_of_nat) also have "\ / ennreal (1 - q) = ennreal ((u * q + l * p) / (1 - q))" using q_def assms by (intro divide_ennreal) auto finally show ?thesis unfolding q_def by simp qed lemma NR\<^sub>N_NR\<^sub>N_l_0: assumes n: "0 < n" and p: "p \ {0<..<1}" and "u \ 1" shows "NR\<^sub>N n u 0 = (u * q / (u * q + l * p)) * NR\<^sub>N n u l" using n proof (induction n rule: less_induct) case (less i) have 1: "0 < u * q" unfolding q_def using assms by simp moreover have "0 \ l * p" using assms by auto ultimately have 2: "0 < u * q + l * p" by arith define c where "c = ennreal (u * q / (u * q + l * p))" have [simp]: "c / c = 1" proof - have "u * q / (u * q + l * p) \ 0" using assms q_def 2 by auto then show ?thesis unfolding c_def using p q_def by (auto intro!: ennreal_divide_self) qed show ?case proof (cases "i = 1") case True have "c * NR\<^sub>N i u l = c * ((u * q + l * p) / p)" unfolding c_def True by (subst NR\<^sub>N_1[OF p]) auto also have "\ = ennreal ((u * q / (u * q + l * p)) * ((u * q + l * p) / p))" unfolding c_def using assms q_def by (subst ennreal_mult'') auto also have "(u * q / (u * q + l * p)) * ((u * q + l * p) / p) = u * q / p" proof - have I: "(a / b) * (b / c) = a / c" if "0 < b" for a b c::"real" using that by (auto) show ?thesis using 2 q_def by (intro I) auto qed also have "\ = NR\<^sub>N i u 0" unfolding True c_def by (subst NR\<^sub>N_1[OF p]) (auto) finally show ?thesis unfolding c_def using True by simp next case False then have i: "i > 1" using less by auto define c where "c = ennreal (u * q / (u * q + l * p))" define B where "B = (\kN (k + 1) u l * ennreal (pmf (binomial_pmf (i - 1) q) k))" have "NR\<^sub>N i u 0 = (p * NR\<^sub>N (i - 1) u 0 + q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k)))) / (1 - (q ^ i))" using less assms by (subst NR\<^sub>N_l_0) auto also have "q * (u + (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k))) = q * u + q * (\kN (k + 1) u 0 * (pmf (binomial_pmf (i - 1) q) k))" using assms q_def by (auto simp add: field_simps ennreal_of_nat_eq_real_of_nat ennreal_mult) also have "NR\<^sub>N (i - 1) u 0 = c * NR\<^sub>N (i - 1) u l" unfolding c_def using less i by (intro less) (auto) also have "(\kN (k + 1) u 0 * ennreal (pmf (binomial_pmf (i - 1) q) k)) = (\kN (k + 1) u l * ennreal (pmf (binomial_pmf (i - 1) q) k))" by (auto intro!: sum.cong simp add: less c_def) also have "\ = c * B" unfolding B_def by (subst sum_distrib_left) (auto intro!: sum.cong mult_ac) also have "q * (c * B) = c * (q * B)" by (simp add: mult_ac) also have "ennreal (q * real u) = q * u * ((u * q + l * p) / (u * q + l * p))" using assms 2 by (auto simp add: field_simps q_def) also have "\ = c * (real u * q + real l * p)" unfolding c_def using 2 by (subst ennreal_mult''[symmetric]) (auto simp add: mult_ac) also have "c * ennreal (real u * q + real l * p) + c * (ennreal q * B) = c * (ennreal (real u * q + real l * p) + (ennreal q * B))" by (auto simp add: field_simps) also have "ennreal p * (c * NR\<^sub>N (i - 1) u l) = c * (ennreal p * NR\<^sub>N (i - 1) u l)" by (simp add: mult_ac) also have "(c * (ennreal p * NR\<^sub>N (i - 1) u l) + c * (ennreal (u * q + l * p) + ennreal q * B)) = c * ((ennreal p * NR\<^sub>N (i - 1) u l) + (ennreal (u * q + l * p) + ennreal q * B))" by (auto simp add: field_simps) also have " c * (ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (u * q + l * p) + ennreal q * B)) / ennreal (1 - q ^ i) = c * ((ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (u * q + l * p) + ennreal q * B)) / ennreal (1 - q ^ i))" by (auto simp add: ennreal_times_divide) also have "(ennreal p * NR\<^sub>N (i - 1) u l + (ennreal (real u * q + real l * p) + ennreal q * B)) / ennreal (1 - q ^ i) = NR\<^sub>N i u l" apply(subst (2) NR\<^sub>N_recurrence') using i assms q_def by (auto simp add: field_simps B_def ennreal_of_nat_eq_real_of_nat ennreal_mult' ennreal_mult'') finally show ?thesis unfolding c_def by simp qed qed text \ Assigning 1 as the cost for going up and/or left, we can now show the relation between the expected length of the reverse search path and the expected height. \ definition EL\<^sub>N where "EL\<^sub>N n = measure_pmf.expectation (R\<^sub>N n 1 1) real" theorem EH\<^sub>N_EL\<^sub>s\<^sub>p: assumes "p \ {0<..<1}" shows "1 / q * EH\<^sub>N n = EL\<^sub>N n" proof - have 1: "ennreal (1 / y * x) = r" if "ennreal x = y * r" "x \ 0" "y > 0" for x y::real and r::ennreal proof - have "ennreal ((1 / y) * x) = ennreal (1 / y) * ennreal x" using that apply(subst ennreal_mult'') by auto also note that(1) also have "ennreal (1 / y) * (ennreal y * r) = ennreal ((1 / y) * y) * r" using that by (subst ennreal_mult'') (auto simp add: mult_ac) also have "(1 / y) * y = 1" using that by (auto) finally show ?thesis by auto qed have "EH\<^sub>N n = NH\<^sub>N n" using NH\<^sub>N_EH\<^sub>N assms by auto also have "NH\<^sub>N n = NR\<^sub>N n 1 0" using NR\<^sub>n_NH\<^sub>N by auto also have "NR\<^sub>N n 1 0 = q * NR\<^sub>N n 1 1" if "n > 0" using NR\<^sub>N_NR\<^sub>N_l_0[of _ 1 1] that assms q_def by force finally have "ennreal (EH\<^sub>N n) = q * NR\<^sub>N n 1 1" if "n > 0" using that by blast then have "1 / q * EH\<^sub>N n = NR\<^sub>N n 1 1" if "n > 0" using that assms q_def by (intro 1) (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def) moreover have "1 / q * EH\<^sub>N n = NR\<^sub>N n 1 1" if "n = 0" unfolding that by (auto simp add: EH\<^sub>N_def H\<^sub>N_def H_def) ultimately have 2: "ennreal (1 / q * EH\<^sub>N n) = NR\<^sub>N n 1 1" by blast also have "NR\<^sub>N n 1 1 = EL\<^sub>N n" using 2 assms EH\<^sub>N_bounds unfolding EL\<^sub>N_def NR\<^sub>N_def by(subst nn_integral_eq_integral) (auto intro!: integrableI_nn_integral_finite[where x="EH\<^sub>N n / q"]) finally show ?thesis using assms q_def ennreal_inj unfolding EL\<^sub>N_def EH\<^sub>N_def H\<^sub>N_def H_def SL_def by (auto) qed end (* context random_skip_list *) thm random_skip_list.EH\<^sub>N_EL\<^sub>s\<^sub>p[unfolded random_skip_list.q_def] random_skip_list.EH\<^sub>N_bounds'[unfolded random_skip_list.q_def] end diff --git a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy --- a/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy +++ b/thys/Zeta_3_Irrational/Zeta_3_Irrational.thy @@ -1,2092 +1,2092 @@ (* File: Zeta_3_Irrational.thy Author: Manuel Eberl, TU München *) section \The Irrationality of $\zeta(3)$\ theory Zeta_3_Irrational imports "E_Transcendental.E_Transcendental" "Prime_Number_Theorem.Prime_Number_Theorem" "Prime_Distribution_Elementary.PNT_Consequences" begin hide_const (open) UnivPoly.coeff UnivPoly.up_ring.monom hide_const (open) Module.smult Coset.order text \ Ap\'{e}ry's original proof of the irrationality of $\zeta(3)$ contained several tricky identities of sums involving binomial coefficients that are difficult to prove. I instead follow Beukers's proof, which consists of several fairly straightforward manipulations of integrals with absolutely no caveats. Both Ap\'{e}ry and Beukers make use of an asymptotic upper bound on $\text{lcm}\{1\ldots n\}$ -- namely $\text{lcm}\{1\ldots n\} \in o(c^n)$ for any $c > e$, which is a consequence of the Prime Number Theorem (which, fortunately, is available in the \emph{Archive of Formal Proofs}~\<^cite>\"afp_primes1" and "afp_primes2"\). I follow the concise presentation of Beukers's proof in Filaseta's lecture notes~\<^cite>\"filaseta"\, which turned out to be very amenable to formalisation. There is another earlier formalisation of the irrationality of $\zeta(3)$ by Mahboubi\ \emph{et al.}~\<^cite>\"mahboubi"\, who followed Ap\'{e}ry's original proof, but were ultimately forced to find a more elementary way to prove the asymptotics of $\text{lcm}\{1\ldots n\}$ than the Prime Number Theorem. First, we will require some auxiliary material before we get started with the actual proof. \ (* TODO: A lot of this (if not all of it) should really be in the library *) subsection \Auxiliary facts about polynomials\ lemma degree_higher_pderiv: "degree ((pderiv ^^ n) p) = degree p - n" for p :: "'a::{comm_semiring_1,semiring_no_zero_divisors,semiring_char_0} poly" by (induction n arbitrary: p) (auto simp: degree_pderiv) lemma pcompose_power_left: "pcompose (p ^ n) q = pcompose p q ^ n" by (induction n) (auto simp: pcompose_mult pcompose_1) lemma pderiv_sum: "pderiv (\x\A. f x) = (\x\A. pderiv (f x))" by (induction A rule: infinite_finite_induct) (auto simp: pderiv_add) lemma higher_pderiv_minus: "(pderiv ^^ n) (-p :: 'a :: idom poly) = -(pderiv ^^ n) p" by (induction n) (auto simp: pderiv_minus) lemma pderiv_power: "pderiv (p ^ n) = smult (of_nat n) (p ^ (n - 1)) * pderiv p" by (cases n) (simp_all add: pderiv_power_Suc del: power_Suc) lemma pderiv_monom: "pderiv (monom c n) = monom (of_nat n * c) (n - 1)" by (simp add: monom_altdef pderiv_smult pderiv_power pderiv_pCons mult_ac) lemma higher_pderiv_monom: "k \ n \ (pderiv ^^ k) (monom c n) = monom (of_nat (pochhammer (n - k + 1) k) * c) (n - k)" by (induction k) (auto simp: pderiv_monom pochhammer_rec Suc_diff_le Suc_diff_Suc mult_ac) lemma higher_pderiv_mult: "(pderiv ^^ n) (p * q) = (\k\n. Polynomial.smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (n - k)) q))" proof (induction n) case (Suc n) have eq: "(Suc n choose k) = (n choose k) + (n choose (k-1))" if "k > 0" for k using that by (cases k) auto have "(pderiv ^^ Suc n) (p * q) = (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q))" by (simp add: Suc pderiv_sum pderiv_smult pderiv_mult sum.distrib smult_add_right algebra_simps Suc_diff_le) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) = (\k\insert 0 {1..n}. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q" by (subst sum.insert) (auto simp: add_ac) also have "(\k\n. smult (of_nat (n choose k)) ((pderiv ^^ Suc k) p * (pderiv ^^ (n - k)) q)) = (\k=1..n+1. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto also have "\ = (\k\insert (n+1) {1..n}. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (intro sum.cong) auto also have "\ = (\k=1..n. smult (of_nat (n choose (k-1))) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + (pderiv ^^ Suc n) p * q" by (subst sum.insert) (auto simp: add_ac) also have "(\k=1..n. smult (of_nat (n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + \ = (\k=1..n. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q)) + p * (pderiv ^^ Suc n) q + (pderiv ^^ Suc n) p * q" by (simp add: sum.distrib algebra_simps smult_add_right eq smult_add_left) also have "\ = (\k\{1..n}\{0,Suc n}. smult (of_nat (Suc n choose k)) ((pderiv ^^ k) p * (pderiv ^^ (Suc n - k)) q))" by (subst sum.union_disjoint) (auto simp: algebra_simps) also have "{1..n}\{0,Suc n} = {..Suc n}" by auto finally show ?case . qed auto subsection \Auxiliary facts about integrals\ theorem (in pair_sigma_finite) Fubini_set_integrable: fixes f :: "_ \ _::{banach, second_countable_topology}" assumes f[measurable]: "set_borel_measurable (M1 \\<^sub>M M2) (A \ B) f" and integ1: "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2)" and integ2: "AE x\A in M1. set_integrable M2 B (\y. f (x, y))" shows "set_integrable (M1 \\<^sub>M M2) (A \ B) f" unfolding set_integrable_def proof (rule Fubini_integrable) note integ1 also have "set_integrable M1 A (\x. \y\B. norm (f (x, y)) \M2) \ integrable M1 (\x. LINT y|M2. norm (indicat_real (A \ B) (x, y) *\<^sub>R f (x, y)))" unfolding set_integrable_def by (intro Bochner_Integration.integrable_cong) (auto simp: indicator_def set_lebesgue_integral_def) finally show \ . next from integ2 show "AE x in M1. integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof eventually_elim case (elim x) show "integrable M2 (\y. indicat_real (A \ B) (x, y) *\<^sub>R f (x, y))" proof (cases "x \ A") case True with elim have "set_integrable M2 B (\y. f (x, y))" by simp also have "?this \ ?thesis" unfolding set_integrable_def using True by (intro Bochner_Integration.integrable_cong) (auto simp: indicator_def) finally show ?thesis . qed auto qed qed (insert assms, auto simp: set_borel_measurable_def) lemma (in pair_sigma_finite) set_integral_fst': fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\x\A. (\y\B. f (x, y) \M2) \M1)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\x. \y. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M2 \M1)" using assms by (subst integral_fst' [symmetric]) (auto simp: set_integrable_def) also have "\ = (\x\A. (\y\B. f (x,y) \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) set_integral_snd: fixes f :: "_ \ 'c :: {second_countable_topology, banach}" assumes "set_integrable (M1 \\<^sub>M M2) (A \ B) f" shows "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\y\B. (\x\A. f (x, y) \M1) \M2)" proof - have "set_lebesgue_integral (M1 \\<^sub>M M2) (A \ B) f = (\z. indicator (A \ B) z *\<^sub>R f z \(M1 \\<^sub>M M2))" by (simp add: set_lebesgue_integral_def) also have "\ = (\y. \x. indicator (A \ B) (x,y) *\<^sub>R f (x,y) \M1 \M2)" using assms by (subst integral_snd) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\y\B. (\x\A. f (x,y) \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong refl) (auto simp: indicator_def) finally show ?thesis . qed proposition (in pair_sigma_finite) Fubini_set_integral: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes f: "set_integrable (M1 \\<^sub>M M2) (A \ B) (case_prod f)" shows "(\y\B. (\x\A. f x y \M1) \M2) = (\x\A. (\y\B. f x y \M2) \M1)" proof - have "(\y\B. (\x\A. f x y \M1) \M2) = (\y. (\x. indicator (A \ B) (x, y) *\<^sub>R f x y \M1) \M2)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) also have "\ = (\x. (\y. indicator (A \ B) (x, y) *\<^sub>R f x y \M2) \M1)" using assms by (intro Fubini_integral) (auto simp: set_integrable_def case_prod_unfold) also have "\ = (\x\A. (\y\B. f x y \M2) \M1)" unfolding set_lebesgue_integral_def by (intro Bochner_Integration.integral_cong) (auto simp: indicator_def) finally show ?thesis . qed lemma (in pair_sigma_finite) nn_integral_swap: assumes [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+x. f x \(M1 \\<^sub>M M2)) = (\\<^sup>+(y,x). f (x,y) \(M2 \\<^sub>M M1))" by (subst distr_pair_swap, subst nn_integral_distr) (auto simp: case_prod_unfold) lemma set_integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" shows "set_integrable M A f \ set_borel_measurable M A g \ (AE x in M. x \ A \ norm (g x) \ norm (f x)) \ set_integrable M A g" unfolding set_integrable_def apply (erule Bochner_Integration.integrable_bound) apply (simp add: set_borel_measurable_def) apply (erule eventually_mono) apply (auto simp: indicator_def) done lemma set_integrableI_nonneg: fixes f :: "'a \ real" assumes "set_borel_measurable M A f" assumes "AE x in M. x \ A \ 0 \ f x" "(\\<^sup>+x\A. f x \M) < \" shows "set_integrable M A f" unfolding set_integrable_def proof (rule integrableI_nonneg) from assms show "(\x. indicator A x *\<^sub>R f x) \ borel_measurable M" by (simp add: set_borel_measurable_def) from assms(2) show "AE x in M. 0 \ indicat_real A x *\<^sub>R f x" by eventually_elim (auto simp: indicator_def) have "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) = (\\<^sup>+x\A. f x \M)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ < \" by fact finally show "(\\<^sup>+x. ennreal (indicator A x *\<^sub>R f x) \M) < \" . qed lemma set_integral_eq_nn_integral: assumes "set_borel_measurable M A f" assumes "set_nn_integral M A f = ennreal x" "x \ 0" assumes "AE x in M. x \ A \ f x \ 0" shows "set_integrable M A f" and "set_lebesgue_integral M A f = x" proof - have eq: "(\x. (if x \ A then 1 else 0) * f x) = (\x. if x \ A then f x else 0)" "(\x. if x \ A then ennreal (f x) else 0) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" "(\x. ennreal (f x * (if x \ A then 1 else 0))) = (\x. ennreal (f x) * (if x \ A then 1 else 0))" by auto from assms show *: "set_integrable M A f" by (intro set_integrableI_nonneg) auto have "set_lebesgue_integral M A f = enn2real (set_nn_integral M A f)" unfolding set_lebesgue_integral_def using assms(1,4) * eq by (subst integral_eq_nn_integral) (auto intro!: nn_integral_cong simp: indicator_def of_bool_def set_integrable_def mult_ac) also have "\ = x" using assms by simp finally show "set_lebesgue_integral M A f = x" . qed lemma set_integral_0 [simp, intro]: "set_integrable M A (\y. 0)" by (simp add: set_integrable_def) lemma set_integrable_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_integrable M A (\y. \x\B. f x y)" using assms by (induction rule: finite_induct) (auto intro!: set_integral_add) lemma set_integral_sum: fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" assumes "finite B" assumes "\x. x \ B \ set_integrable M A (f x)" shows "set_lebesgue_integral M A (\y. \x\B. f x y) = (\x\B. set_lebesgue_integral M A (f x))" using assms apply (induction rule: finite_induct) apply simp apply simp apply (subst set_integral_add) apply (auto intro!: set_integrable_sum) done lemma set_nn_integral_cong: assumes "M = M'" "A = B" "\x. x \ space M \ A \ f x = g x" shows "set_nn_integral M A f = set_nn_integral M' B g" using assms unfolding assms(1,2) by (intro nn_integral_cong) (auto simp: indicator_def) lemma set_nn_integral_mono: assumes "\x. x \ space M \ A \ f x \ g x" shows "set_nn_integral M A f \ set_nn_integral M A g" using assms by (intro nn_integral_mono) (auto simp: indicator_def) lemma fixes f :: "real \ real" assumes "a \ b" assumes deriv: "\x. a \ x \ x \ b \ (F has_field_derivative f x) (at x within {a..b})" assumes cont: "continuous_on {a..b} f" shows has_bochner_integral_FTC_Icc_real: "has_bochner_integral lborel (\x. f x * indicator {a .. b} x) (F b - F a)" (is ?has) and integral_FTC_Icc_real: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) proof - have 1: "\x. a \ x \ x \ b \ (F has_vector_derivative f x) (at x within {a .. b})" unfolding has_real_derivative_iff_has_vector_derivative[symmetric] using deriv by auto show ?has ?eq using has_bochner_integral_FTC_Icc[OF \a \ b\ 1 cont] integral_FTC_Icc[OF \a \ b\ 1 cont] by (auto simp: mult.commute) qed lemma integral_by_parts_integrable: fixes f g F G::"real \ real" assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "integrable lborel (\x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" proof - have "integrable lborel (\x. indicator {a..b} x *\<^sub>R ((F x) * (g x) + (f x) * (G x)))" by (intro borel_integrable_compact continuous_intros assms) (auto intro!: DERIV_continuous_on assms) thus ?thesis by (simp add: mult_ac) qed lemma integral_by_parts: fixes f g F G::"real \ real" assumes [arith]: "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(\x. (F x * g x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a - \x. (f x * G x) * indicator {a .. b} x \lborel" proof- have 0: "(\x. (F x * g x + f x * G x) * indicator {a .. b} x \lborel) = F b * G b - F a * G a" by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros) (auto intro!: assms DERIV_continuous_on) have [continuous_intros]: "continuous_on {a..b} F" by (rule DERIV_continuous_on assms)+ have [continuous_intros]: "continuous_on {a..b} G" by (rule DERIV_continuous_on assms)+ have "(\x. indicator {a..b} x *\<^sub>R (F x * g x + f x * G x) \lborel) = (\x. indicator {a..b} x *\<^sub>R(F x * g x) \lborel) + \x. indicator {a..b} x *\<^sub>R (f x * G x) \lborel" apply (subst Bochner_Integration.integral_add[symmetric]) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (rule borel_integrable_compact; force intro!: continuous_intros assms) apply (simp add: algebra_simps) done thus ?thesis using 0 by (simp add: algebra_simps) qed lemma interval_lebesgue_integral_by_parts: assumes "a \ b" assumes cont_f[intro]: "continuous_on {a..b} f" assumes cont_g[intro]: "continuous_on {a..b} g" assumes [intro]: "\x. x \ {a..b} \ (F has_field_derivative f x) (at x within {a..b})" assumes [intro]: "\x. x \ {a..b} \ (G has_field_derivative g x) (at x within {a..b})" shows "(LBINT x=a..b. F x * g x) = F b * G b - F a * G a - (LBINT x=a..b. f x * G x)" using integral_by_parts[of a b f g F G] assms by (simp add: interval_integral_Icc set_lebesgue_integral_def mult_ac) lemma interval_lebesgue_integral_by_parts_01: assumes cont_f[intro]: "continuous_on {0..1} f" assumes cont_g[intro]: "continuous_on {0..1} g" assumes [intro]: "\x. x \ {0..1} \ (F has_field_derivative f x) (at x within {0..1})" assumes [intro]: "\x. x \ {0..1} \ (G has_field_derivative g x) (at x within {0..1})" shows "(LBINT x=0..1. F x * g x) = F 1 * G 1 - F 0 * G 0 - (LBINT x=0..1. f x * G x)" using interval_lebesgue_integral_by_parts[of 0 1 f g F G] assms by (simp add: zero_ereal_def one_ereal_def) lemma continuous_on_imp_set_integrable_cbox: fixes h :: "'a :: euclidean_space \ real" assumes "continuous_on (cbox a b) h" shows "set_integrable lborel (cbox a b) h" proof - from assms have "h absolutely_integrable_on cbox a b" by (rule absolutely_integrable_continuous) moreover have "(\x. indicat_real (cbox a b) x *\<^sub>R h x) \ borel_measurable borel" by (rule borel_measurable_continuous_on_indicator) (use assms in auto) ultimately show ?thesis unfolding set_integrable_def using assms by (subst (asm) integrable_completion) auto qed subsection \Shifted Legendre polynomials\ text \ The first ingredient we need to show Apéry's theorem is the \<^emph>\shifted Legendre polynomials\ \[P_n(X) = \frac{1}{n!} \frac{\partial^n}{\partial X^n} (X^n(1-X)^n)\] and the auxiliary polynomials \[Q_{n,k}(X) = \frac{\partial^k}{\partial X^k} (X^n(1-X)^n)\ .\] Note that $P_n$ is in fact an \emph{integer} polynomial. Only some very basic properties of these will be proven, since that is all we will need. \ context fixes n :: nat begin definition gen_shleg_poly :: "nat \ int poly" where "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" definition shleg_poly where "shleg_poly = gen_shleg_poly n div [:fact n:]" text \ We can easily prove the following more explicit formula for $Q_{n,k}$: \[Q_{n,k}(X) = \sum_{i=0}^k (-1)^{k-1} {k \choose i} n^{\underline{i}}\, n^{\underline{k-i}}\, X^{n-i}\, (1-X)^{n-k+i}\] \ lemma gen_shleg_poly_altdef: assumes "k \ n" shows "gen_shleg_poly k = (\i\k. smult ((-1)^(k-i) * of_nat (k choose i) * pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n-k+i)))" proof - have *: "(pderiv ^^ i) (x \\<^sub>p [:1, -1:]) = smult ((-1) ^ i) ((pderiv ^^ i) x \\<^sub>p [:1, -1:])" for i and x :: "int poly" by (induction i arbitrary: x) (auto simp: pderiv_smult pderiv_pcompose funpow_Suc_right pderiv_pCons higher_pderiv_minus simp del: funpow.simps(2)) have "gen_shleg_poly k = (pderiv ^^ k) ([:0, 1, -1:] ^ n)" by (simp add: gen_shleg_poly_def) also have "[:0, 1, -1::int:] = [:0, 1:] * [:1, -1:]" by simp also have "\ ^ n = [:0, 1:] ^ n * [:1, -1:] ^ n" by (simp flip: power_mult_distrib) also have "(pderiv ^^ k) \ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) ([:0, 1:] ^ n) * (pderiv ^^ (k - i)) ([:1, -1:] ^ n)))" by (simp add: higher_pderiv_mult) also have "\ = (\i\k. smult (of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * (pderiv ^^ (k - i)) (monom 1 n \\<^sub>p [:1, -1:])))" by (simp add: monom_altdef pcompose_pCons pcompose_power_left) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) ((pderiv ^^ i) (monom 1 n) * ((pderiv ^^ (k - i)) (monom 1 n) \\<^sub>p [:1, -1:])))" by (simp add: * mult_ac) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i)) (monom (pochhammer (n - i + 1) i) (n - i) * monom (pochhammer (n - k + i + 1) (k - i)) (n - k + i) \\<^sub>p [:1, -1:]))" using assms by (simp add: higher_pderiv_monom) also have "\ = (\i\k. smult ((-1) ^ (k - i) * of_nat (k choose i) * pochhammer (n - i + 1) i * pochhammer (n - k + i + 1) (k - i)) ([:0, 1:] ^ (n - i) * [:1, -1:] ^ (n - k + i)))" by (simp add: monom_altdef algebra_simps pcompose_smult pcompose_power_left pcompose_pCons) finally show ?thesis . qed lemma degree_gen_shleg_poly [simp]: "degree (gen_shleg_poly k) = 2 * n - k" by (simp add: gen_shleg_poly_def degree_higher_pderiv degree_power_eq) lemma gen_shleg_poly_n: "gen_shleg_poly n = smult (fact n) shleg_poly" proof - obtain r where r: "gen_shleg_poly n = [:fact n:] * r" unfolding gen_shleg_poly_def using fact_dvd_higher_pderiv[of n "[:0,1,-1:]^n"] by blast have "smult (fact n) shleg_poly = smult (fact n) (gen_shleg_poly n div [:fact n:])" by (simp add: shleg_poly_def) also note r also have "[:fact n:] * r div [:fact n:] = r" by (rule nonzero_mult_div_cancel_left) auto finally show ?thesis by (simp add: r) qed lemma degree_shleg_poly [simp]: "degree shleg_poly = n" using degree_gen_shleg_poly[of n] by (simp add: gen_shleg_poly_n) lemma pderiv_gen_shleg_poly [simp]: "pderiv (gen_shleg_poly k) = gen_shleg_poly (Suc k)" by (simp add: gen_shleg_poly_def) text \ The following functions are the interpretation of the shifted Legendre polynomials and the auxiliary polynomials as a function from reals to reals. \ definition Gen_Shleg :: "nat \ real \ real" where "Gen_Shleg k x = poly (of_int_poly (gen_shleg_poly k)) x" definition Shleg :: "real \ real" where "Shleg = poly (of_int_poly shleg_poly)" lemma Gen_Shleg_altdef: assumes "k \ n" shows "Gen_Shleg k x = (\i\k. (-1)^(k-i) * of_nat (k choose i) * of_int (pochhammer (n-i+1) i * pochhammer (n-k+i+1) (k-i)) * x ^ (n - i) * (1 - x) ^ (n-k+i))" using assms by (simp add: Gen_Shleg_def gen_shleg_poly_altdef poly_sum mult_ac) lemma Gen_Shleg_0 [simp]: "k < n \ Gen_Shleg k 0 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_1 [simp]: "k < n \ Gen_Shleg k 1 = 0" by (simp add: Gen_Shleg_altdef zero_power) lemma Gen_Shleg_n_0 [simp]: "Gen_Shleg n 0 = fact n" proof - have "Gen_Shleg n 0 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{n}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ (n - i))" by (intro sum.mono_neutral_right) auto also have "\ = fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Gen_Shleg_n_1 [simp]: "Gen_Shleg n 1 = (-1) ^ n * fact n" proof - have "Gen_Shleg n 1 = (\i\n. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (simp add: Gen_Shleg_altdef) also have "\ = (\i\{0}. (-1) ^ (n-i) * real (n choose i) * (real (pochhammer (Suc (n-i)) i) * real (pochhammer (Suc i) (n-i))) * 0 ^ i)" by (intro sum.mono_neutral_right) auto also have "\ = (-1) ^ n * fact n" by (simp add: pochhammer_fact flip: pochhammer_of_nat) finally show ?thesis . qed lemma Shleg_altdef: "Shleg x = Gen_Shleg n x / fact n" by (simp add: Shleg_def Gen_Shleg_def gen_shleg_poly_n) lemma Shleg_0 [simp]: "Shleg 0 = 1" and Shleg_1 [simp]: "Shleg 1 = (-1) ^ n" by (simp_all add: Shleg_altdef) lemma Gen_Shleg_0_left: "Gen_Shleg 0 x = x ^ n * (1 - x) ^ n" by (simp add: Gen_Shleg_def gen_shleg_poly_def power_mult_distrib) lemma has_field_derivative_Gen_Shleg: "(Gen_Shleg k has_field_derivative Gen_Shleg (Suc k) x) (at x)" proof - note [derivative_intros] = poly_DERIV show ?thesis unfolding Gen_Shleg_def by (rule derivative_eq_intros) auto qed lemma continuous_on_Gen_Shleg: "continuous_on A (Gen_Shleg k)" by (auto simp: Gen_Shleg_def intro!: continuous_intros) lemma continuous_on_Gen_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Gen_Shleg k (f x))" by (rule continuous_on_compose2[OF continuous_on_Gen_Shleg[of UNIV]]) auto lemma continuous_on_Shleg: "continuous_on A Shleg" by (auto simp: Shleg_def intro!: continuous_intros) lemma continuous_on_Shleg' [continuous_intros]: "continuous_on A f \ continuous_on A (\x. Shleg (f x))" by (rule continuous_on_compose2[OF continuous_on_Shleg[of UNIV]]) auto lemma measurable_Gen_Shleg [measurable]: "Gen_Shleg n \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Gen_Shleg) lemma measurable_Shleg [measurable]: "Shleg \ borel_measurable borel" by (intro borel_measurable_continuous_onI continuous_on_Shleg) end subsection \Auxiliary facts about the \\\ function\ lemma Re_zeta_ge_1: assumes "x > 1" shows "Re (zeta (of_real x)) \ 1" proof - have *: "(\n. real (Suc n) powr -x) sums Re (zeta (complex_of_real x))" using sums_Re[OF sums_zeta[of "of_real x"]] assms by (simp add: powr_Reals_eq) show "Re (zeta (of_real x)) \ 1" proof (rule sums_le[OF _ _ *]) show "(\n. if n = 0 then 1 else 0) sums 1" by (rule sums_single) qed auto qed lemma sums_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. 1 / (k + 1) ^ n) sums zeta (of_nat n)" using sums_zeta[of "of_nat n"] n by (simp add: powr_minus field_simps flip: of_nat_Suc) from sums_split_initial_segment[OF this, of r] have "(\k. 1 / (r + k + 1) ^ n) sums (zeta (of_nat n) - (\kkk=1..r. 1 / k ^ n)" by (intro sum.reindex_bij_witness[of _ "\k. k - 1" Suc]) auto finally show ?thesis . qed lemma sums_Re_zeta_of_nat_offset: fixes r :: nat assumes n: "n > 1" shows "(\k. 1 / (r + k + 1) ^ n) sums (Re (zeta (of_nat n)) - (\k=1..r. 1 / k ^ n))" proof - have "(\k. Re (1 / (r + k + 1) ^ n)) sums (Re (zeta (of_nat n) - (\k=1..r. 1 / k ^ n)))" by (intro sums_Re sums_zeta_of_nat_offset assms) thus ?thesis by simp qed subsection \Divisor of a sum of rationals\ text \ A finite sum of rationals of the form $\frac{a_1}{b_1} + \ldots + \frac{a_n}{b_n}$ can be brought into the form $\frac{c}{d}$, where $d$ is the LCM of the $b_i$ (or some integer multiple thereof). \ lemma sum_rationals_common_divisor: fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" using assms proof (induction rule: finite_induct) case empty thus ?case by auto next case (insert x A) define d where "d = (LCM x\A. g x)" from insert have [simp]: "d \ 0" by (auto simp: d_def Lcm_0_iff) from insert have [simp]: "g x \ 0" by auto from insert obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d" by (auto simp: d_def) define e1 where "e1 = lcm d (g x) div d" define e2 where "e2 = lcm d (g x) div g x" have "(\y\insert x A. f y / g y) = c / d + f x / g x" using insert c by simp also have "c / d = (c * e1) / lcm d (g x)" by (simp add: e1_def real_of_int_div) also have "f x / g x = (f x * e2) / lcm d (g x)" by (simp add: e2_def real_of_int_div) also have "(c * e1) / lcm d (g x) + \ = (c * e1 + f x * e2) / (LCM x\insert x A. g x)" using insert by (simp add: add_divide_distrib lcm.commute d_def) finally show ?case .. qed lemma sum_rationals_common_divisor': fixes f g :: "'a \ int" assumes "finite A" assumes "\x. x \ A \ g x \ 0" and "(\x. x \ A \ g x dvd d)" and "d \ 0" shows "\c. (\x\A. f x / g x) = real_of_int c / real_of_int d" proof - define d' where "d' = (LCM x\A. g x)" have "d' dvd d" unfolding d'_def using assms(3) by (auto simp: Lcm_dvd_iff) then obtain e where e: "d = d' * e" by blast have "\c. (\x\A. f x / g x) = real_of_int c / (LCM x\A. g x)" by (rule sum_rationals_common_divisor) fact+ then obtain c where c: "(\x\A. f x / g x) = real_of_int c / real_of_int d'" unfolding d'_def .. also have "\ = real_of_int (c * e) / real_of_int d" using \d \ 0\ by (simp add: e) finally show ?thesis .. qed subsection \The first double integral\ text \ We shall now investigate the double integral \[I_1 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy}\,x^r y^s\, \text{d}x\,\text{d}y\ .\] Since everything is non-negative for now, we can work over the extended non-negative real numbers and the issues of integrability or summability do not arise at all. \ definition beukers_nn_integral1 :: "nat \ nat \ ennreal" where "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" definition beukers_integral1 :: "nat \ nat \ real" where "beukers_integral1 r s = (\(x,y)\{0<..<1}\{0<..<1}. (-ln (x*y) / (1 - x*y) * x^r * y^s) \lborel)" lemma fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0..1}" shows beukers_denom_ineq: "(1 - x * y) * z < 1" and beukers_denom_neq: "(1 - x * y) * z \ 1" proof - from xyz have *: "x * y < 1 * 1" by (intro mult_strict_mono) auto from * have "(1 - x * y) * z \ (1 - x * y) * 1" using xyz by (intro mult_left_mono) auto also have "\ < 1 * 1" using xyz by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * z < 1" "(1 - x * y) * z \ 1" by simp_all qed text \ We first evaluate the improper integral \[\int_0^1 -\ln x \cdot x^e\,\text{d}x = \frac{1}{(e+1)^2}\ .\] for any $e>-1$. \ lemma integral_0_1_ln_times_powr: assumes "e > -1" shows "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" and "interval_lebesgue_integrable lborel 0 1 (\x. -ln x * x powr e)" proof - define f where "f = (\x. -ln x * x powr e)" define F where "F = (\x. x powr (e + 1) * (1 - (e + 1) * ln x) / (e + 1) ^ 2)" have 0: "isCont f x" if "x \ {0<..<1}" for x using that by (auto intro!: continuous_intros simp: f_def) have 1: "(F has_real_derivative f x) (at x)" if "x \ {0<..<1}" for x proof - show "(F has_real_derivative f x) (at x)" unfolding F_def f_def using that assms apply (insert that assms) apply (rule derivative_eq_intros refl | simp)+ apply (simp add: divide_simps) apply (simp add: power2_eq_square algebra_simps powr_add power_numeral_reduce) done qed have 2: "AE x in lborel. ereal 0 < ereal x \ ereal x < ereal 1 \ 0 \ f x" by (intro AE_I2) (auto simp: f_def mult_nonpos_nonneg) have 3: "((F \ real_of_ereal) \ 0) (at_right (ereal 0))" unfolding ereal_tendsto_simps F_def using assms by real_asymp have 4: "((F \ real_of_ereal) \ F 1) (at_left (ereal 1))" unfolding ereal_tendsto_simps F_def using assms by real_asymp (simp add: field_simps) have "(LBINT x=ereal 0..ereal 1. f x) = F 1 - 0" by (rule interval_integral_FTC_nonneg[where F = F]) (use 0 1 2 3 4 in auto) thus "(LBINT x=0..1. -ln x * x powr e) = 1 / (e + 1)\<^sup>2" by (simp add: F_def zero_ereal_def one_ereal_def f_def) have "set_integrable lborel (einterval (ereal 0) (ereal 1)) f" by (rule interval_integral_FTC_nonneg) (use 0 1 2 3 4 in auto) thus "interval_lebesgue_integrable lborel 0 1 f" by (simp add: interval_lebesgue_integrable_def einterval_def) qed lemma interval_lebesgue_integral_lborel_01_cong: assumes "\x. x \ {0<..<1} \ f x = g x" shows "interval_lebesgue_integral lborel 0 1 f = interval_lebesgue_integral lborel 0 1 g" using assms by (subst (1 2) interval_integral_Ioo) (auto intro!: set_lebesgue_integral_cong assms) lemma nn_integral_0_1_ln_times_powr: assumes "e > -1" shows "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = ennreal (1 / (e + 1)\<^sup>2)" proof - have *: "(LBINT x=0..1. - ln x * x powr e = 1 / (e + 1)\<^sup>2)" "interval_lebesgue_integrable lborel 0 1 (\x. - ln x * x powr e)" using integral_0_1_ln_times_powr[OF assms] by simp_all have eq: "(\y. (if 0 < y \ y < 1 then 1 else 0) * ln y * y powr e) = (\y. if 0 < y \ y < 1 then ln y * y powr e else 0)" by auto have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr e) \lborel) = (\\<^sup>+y. ennreal (-ln y * y powr e * indicator {0<..<1} y) \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def) also have "\ = ennreal (1 / (e + 1)\<^sup>2)" using * eq by (subst nn_integral_eq_integral) (auto intro!: AE_I2 simp: indicator_def interval_lebesgue_integrable_def set_integrable_def one_ereal_def zero_ereal_def interval_integral_Ioo mult_ac mult_nonpos_nonneg set_lebesgue_integral_def) finally show ?thesis . qed lemma nn_integral_0_1_ln_times_power: "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = ennreal (1 / (n + 1)\<^sup>2)" proof - have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y ^ n) \lborel) = (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y powr real n) \lborel)" by (intro set_nn_integral_cong) (auto simp: powr_realpow) also have "\ = ennreal (1 / (n + 1)^2)" by (subst nn_integral_0_1_ln_times_powr) auto finally show ?thesis by simp qed text \ Next, we also evaluate the more trivial integral \[\int_0^1 x^n\ \text{d}x\ .\] \ lemma nn_integral_0_1_power: "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = ennreal (1 / (n + 1))" proof - have *: "((\a. a ^ (n + 1) / real (n + 1)) has_real_derivative x ^ n) (at x)" for x by (rule derivative_eq_intros refl | simp)+ have "(\\<^sup>+y\{0<..<1}. ennreal (y ^ n) \lborel) = (\\<^sup>+y\{0..1}. ennreal (y ^ n) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: indicator_def emeasure_lborel_countable) also have "\ = ennreal (1 ^ (n + 1) / (n + 1) - 0 ^ (n + 1) / (n + 1))" using * by (intro nn_integral_FTC_Icc) auto also have "\ = ennreal (1 / (n + 1))" by simp finally show ?thesis by simp qed text \ $I_1$ can alternatively be written as the triple integral \[\int_0^1\int_0^1\int_0^1 \frac{x^r y^s}{1-(1-xy)w}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ lemma beukers_nn_integral1_altdef: "beukers_nn_integral1 r s = (\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel)" proof - have "(\\<^sup>+(w,x,y)\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\\<^sup>+w\{0<..<1}. ennreal (1 / (1-(1-x*y)*w) * x^r * y^s) \lborel) \lborel)" by (subst lborel_prod [symmetric], subst lborel_pair.nn_integral_snd [symmetric]) (auto simp: case_prod_unfold indicator_def simp flip: lborel_prod intro!: nn_integral_cong) also have "\ = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y)/(1-x*y) * x^r * y^s) \lborel)" proof (intro nn_integral_cong, clarify) fix x y :: real have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (-ln (x*y)*x^r*y^s/(1-x*y))" if xy: "(x, y) \ {0<..<1} \ {0<..<1}" proof - from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have deriv: "((\w. -ln (1-(1-x*y)*w) / (1-x*y)) has_real_derivative 1/(1-(1-x*y)*w)) (at w)" if w: "w \ {0..1}" for w by (insert xy w \x*y<1\ beukers_denom_ineq[of x y w]) (rule derivative_eq_intros refl | simp add: divide_simps)+ have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) = ennreal (x^r*y^s) * (\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" using xy by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = (\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel)" by (intro nn_integral_cong_AE AE_I[of _ _ "{0,1}"]) (auto simp: emeasure_lborel_countable indicator_def) also have "(\\<^sup>+w\{0..1}. ennreal (1/(1-(1-x*y)*w)) \lborel) = ennreal (-ln (1-(1-x*y)*1)/(1-x*y) - (-ln (1-(1-x*y)*0)/(1-x*y)))" using xy deriv less_imp_le[OF beukers_denom_ineq[of x y]] by (intro nn_integral_FTC_Icc) auto finally show ?thesis using xy by (simp flip: ennreal_mult' ennreal_mult'' add: mult_ac) qed thus "(\\<^sup>+w\{0<..<1}. ennreal (1/(1-(1-x*y)*w)*x^r*y^s) \lborel) * indicator ({0<..<1}\{0<..<1}) (x, y) = ennreal (-ln (x*y)/(1-x*y)*x^r*y^s) * indicator ({0<..<1}\{0<..<1}) (x, y)" by (auto simp: indicator_def) qed also have "\ = beukers_nn_integral1 r s" by (simp add: beukers_nn_integral1_def) finally show ?thesis .. qed context fixes r s :: nat and I1 I2' :: real and I2 :: ennreal and D :: "(real \ real \ real) set" assumes rs: "s \ r" defines "D \ {0<..<1}\{0<..<1}\{0<..<1}" begin text \ By unfolding the geometric series, pulling the summation out and evaluating the integrals, we find that \[I_1 = \sum_{k=0}^\infty \frac{1}{(k+r+1)^2(k+s+1)} + \frac{1}{(k+r+1)(k+s+1)^2}\ .\] \ lemma beukers_nn_integral1_series: "beukers_nn_integral1 r s = (\k. ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2)))" proof - have "beukers_nn_integral1 r s = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel)" unfolding beukers_nn_integral1_def proof (intro set_nn_integral_cong refl, clarify) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by simp have "(\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) = ennreal (-ln (x*y) * x^r * y^s) * (\k. ennreal ((x*y)^k))" using xy by (subst ennreal_suminf_cmult [symmetric], subst ennreal_mult'' [symmetric]) (auto simp: power_add mult_ac power_mult_distrib) also have "(\k. ennreal ((x*y)^k)) = ennreal (1 / (1 - x*y))" using geometric_sums[of "x*y"] \x * y < 1\ xy by (intro suminf_ennreal_eq) auto also have "ennreal (-ln (x*y) * x^r * y^s) * \ = ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s)" using \x * y < 1\ by (subst ennreal_mult'' [symmetric]) auto finally show "ennreal (-ln (x*y) / (1 - x*y) * x^r * y^s) = (\k. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)))" .. qed also have "\ = (\k. (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. (ennreal (-ln (x*y) * x^(k+r) * y^(k+s))) \lborel))" unfolding case_prod_unfold by (subst nn_integral_suminf [symmetric]) (auto simp flip: borel_prod) also have "\ = (\k. ennreal (1 / ((k+r+1)^2*(k+s+1)) + 1 / ((k+r+1)*(k+s+1)^2)))" proof (rule suminf_cong) fix k :: nat define F where "F = (\x y::real. x + y)" have "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+x\{0<..<1}. (\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) \lborel)" unfolding case_prod_unfold lborel_prod [symmetric] by (subst lborel.nn_integral_fst [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2) \lborel)" proof (intro set_nn_integral_cong refl, clarify) fix x :: real assume x: "x \ {0<..<1}" have "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = (\\<^sup>+y\{0<..<1}. (ennreal (-ln x * x^(k+r) * y^(k+s)) + ennreal (-ln y * x^(k+r) * y^(k+s))) \lborel)" by (intro set_nn_integral_cong) (use x in \auto simp: ln_mult ring_distribs mult_nonpos_nonneg simp flip: ennreal_plus\) also have "\ = (\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) + (\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln x * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult'') also have "(\\<^sup>+y\{0<..<1}. ennreal (y^(k+s)) \lborel) = ennreal (1/(k+s+1))" by (subst nn_integral_0_1_power) simp also have "ennreal (-ln x * x^(k+r)) * \ = ennreal (-ln x * x^(k+r) / (k+s+1))" by (subst ennreal_mult'' [symmetric]) auto also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * x^(k+r) * y^(k+s)) \lborel) = ennreal (x^(k+r)) * (\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel)" by (subst nn_integral_cmult [symmetric]) (use x in \auto intro!: nn_integral_cong simp: indicator_def mult_ac simp flip: ennreal_mult'\) also have "(\\<^sup>+y\{0<..<1}. ennreal (-ln y * y^(k+s)) \lborel) = ennreal (1 / (k + s + 1)\<^sup>2)" by (subst nn_integral_0_1_ln_times_power) simp also have "ennreal (x ^ (k + r)) * \ = ennreal (x ^ (k + r) / (k + s + 1) ^ 2)" by (subst ennreal_mult'' [symmetric]) auto also have "ennreal (- ln x * x ^ (k + r) / (k + s + 1)) + \ = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" using x by (subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) finally show "(\\<^sup>+y\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = ennreal (-ln x * x^(k+r) / (k+s+1) + x^(k+r)/(k+s+1)^2)" . qed also have "\ = (\\<^sup>+x\{0<..<1}. (ennreal (-ln x * x^(k+r) / (k+s+1)) + ennreal (x^(k+r)/(k+s+1)^2)) \lborel)" by (intro set_nn_integral_cong refl, subst ennreal_plus) (auto simp: mult_nonpos_nonneg divide_nonpos_nonneg) also have "\ = (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) + (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel)" by (subst nn_integral_add [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def) also have "(\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r) / (k+s+1)) \lborel) = ennreal (1 / (k+s+1)) * (\\<^sup>+x\{0<..<1}. ennreal (-ln x * x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1 / ((k+s+1) * (k+r+1)^2))" by (subst nn_integral_0_1_ln_times_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "(\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)/(k+s+1)^2) \lborel) = ennreal (1/(k+s+1)^2) * (\\<^sup>+x\{0<..<1}. ennreal (x^(k+r)) \lborel)" by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "\ = ennreal (1/((k+r+1)*(k+s+1)^2))" by (subst nn_integral_0_1_power, subst ennreal_mult [symmetric]) (auto simp: algebra_simps) also have "ennreal (1 / ((k+s+1) * (k+r+1)^2)) + \ = ennreal (1/((k+r+1)^2*(k+s+1)) + 1/((k+r+1)*(k+s+1)^2))" by (subst ennreal_plus [symmetric]) (auto simp: algebra_simps) finally show "(\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (-ln (x*y) * x^(k+r) * y^(k+s)) \lborel) = \" . qed finally show ?thesis . qed text \ Remembering that $\zeta(3) = \sum k^{-3}$, it is easy to see that if $r = s$, this sum is simply \[2\left(\zeta(3) - \sum_{k=1}^r \frac{1}{k^3}\right)\ .\] \ lemma beukers_nn_integral1_same: assumes "r = s" shows "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" and "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" proof - from assms have [simp]: "s = r" by simp have *: "Suc 2 = 3" by simp have "beukers_nn_integral1 r s = (\k. ennreal (2 / (r + k + 1) ^ 3))" unfolding beukers_nn_integral1_series by (simp only: assms power_Suc [symmetric] mult.commute[of "x ^ 2" for x] * times_divide_eq_right mult_1_right add_ac flip: mult_2) also have **: "(\k. 2 / (r + k + 1) ^ 3) sums (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" using sums_mult[OF sums_Re_zeta_of_nat_offset[of 3], of 2] by simp hence "(\k. ennreal (2 / (r + k + 1) ^ 3)) = ennreal \" by (intro suminf_ennreal_eq) auto finally show "beukers_nn_integral1 r s = ennreal (2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)))" . show "2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3)) \ 0" by (rule sums_le[OF _ sums_zero **]) auto qed lemma beukers_integral1_same: assumes "r = s" shows "beukers_integral1 r s = 2 * (Re (zeta 3) - (\k=1..r. 1 / k ^ 3))" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_same[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed text \ In contrast, for \r > s\, we find that \[I_1 = \frac{1}{r-s} \sum_{k=s+1}^r \frac{1}{k^2}\ .\] \ lemma beukers_nn_integral1_different: assumes "r > s" shows "beukers_nn_integral1 r s = ennreal ((\k\{s<..r}. 1 / k ^ 2) / (r - s))" proof - have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) sums (1 / (r - s) * ((Re (zeta (of_nat 2)) - (\k=1..s. 1/k^2)) - (Re (zeta (of_nat 2)) - (\k=1..r. 1/k^2))))" (is "_ sums ?S") by (intro sums_mult sums_diff sums_Re_zeta_of_nat_offset) auto also have "?S = ((\k=1..r. 1/k^2) - (\k=1..s. 1/k^2)) / (r - s)" by (simp add: algebra_simps diff_divide_distrib) also have "(\k=1..r. 1/k^2) - (\k=1..s. 1/k^2) = (\k\{1..r}-{1..s}. 1/k^2)" using assms by (subst Groups_Big.sum_diff) auto also have "{1..r} - {1..s} = {s<..r}" by auto also have "(\k. 1 / (r - s) * (1 / (s + k + 1) ^ 2 - 1 / (r + k + 1) ^ 2)) = (\k. 1 / ((k+r+1) * (k+s+1)^2) + 1 / ((k+r+1)^2 * (k+s+1)))" proof (intro ext, goal_cases) case (1 k) define x where "x = real (k + r + 1)" define y where "y = real (k + s + 1)" have [simp]: "x \ 0" "y \ 0" by (auto simp: x_def y_def) have "(x\<^sup>2 * y + x * y\<^sup>2) * (real r - real s) = x * y * (x\<^sup>2 - y\<^sup>2)" by (simp add: algebra_simps power2_eq_square x_def y_def) hence "1 / (x*y^2) + 1 / (x^2*y) = 1 / (r - s) * (1 / y^2 - 1 / x^2)" using assms by (simp add: divide_simps of_nat_diff) thus ?case by (simp add: x_def y_def algebra_simps) qed finally show ?thesis unfolding beukers_nn_integral1_series by (intro suminf_ennreal_eq) (auto simp: add_ac) qed lemma beukers_integral1_different: assumes "r > s" shows "beukers_integral1 r s = (\k\{s<..r}. 1 / k ^ 2) / (r - s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using that mult_strict_mono[of a 1 b 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) auto thus ?thesis using beukers_nn_integral1_different[OF assms] unfolding beukers_nn_integral1_def beukers_integral1_def by (intro set_integral_eq_nn_integral AE_I2) (auto simp flip: lborel_prod simp: case_prod_unfold set_borel_measurable_def intro: divide_nonpos_nonneg mult_nonpos_nonneg intro!: sum_nonneg divide_nonneg_nonneg) qed end text \ It is also easy to see that if we exchange \r\ and \s\, nothing changes. \ lemma beukers_nn_integral1_swap: "beukers_nn_integral1 r s = beukers_nn_integral1 s r" unfolding beukers_nn_integral1_def lborel_prod [symmetric] by (subst lborel_pair.nn_integral_swap, simp) (intro nn_integral_cong, auto simp: indicator_def algebra_simps split: if_splits) lemma beukers_nn_integral1_finite: "beukers_nn_integral1 r s < \" using beukers_nn_integral1_different[of r s] beukers_nn_integral1_different[of s r] by (cases r s rule: linorder_cases) (simp_all add: beukers_nn_integral1_same beukers_nn_integral1_swap) lemma beukers_integral1_integrable: "set_integrable lborel ({0<..<1}\{0<..<1}) (\(x,y). (-ln (x*y) / (1 - x*y) * x^r * y^s :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y :: real assume xy: "x \ {0<..<1}" "y \ {0<..<1}" have "0 \ ln (x * y) / (1 - x * y) * x ^ r * y ^ s" using mult_strict_mono[of x 1 y 1] by (intro mult_nonpos_nonneg divide_nonpos_nonneg) (use xy in auto) thus "0 \ - ln (x * y) / (1 - x * y) * x ^ r * y ^ s" by simp next - show "\\<^sup>+x\{0<..<1}\{0<..<1}. ennreal (case x of (x, y) \ - - ln (x * y) / (1 - x * y) * x ^ r * y ^ s) \lborel < \" + show "(\\<^sup>+x\{0<..<1}\{0<..<1}. ennreal (case x of (x, y) \ + - ln (x * y) / (1 - x * y) * x ^ r * y ^ s) \lborel) < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_def case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_integrable': "set_integrable lborel ({0<..<1}\{0<..<1}\{0<..<1}) (\(z,x,y). (x^r * y^s / (1 - (1 - x*y) * z) :: real))" proof (intro set_integrableI_nonneg AE_I2; clarify?) fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" show "0 \ x^r * y^s / (1 - (1 - x*y) * z)" using mult_strict_mono[of x 1 y 1] xyz beukers_denom_ineq[of x y z] by (intro mult_nonneg_nonneg divide_nonneg_nonneg) auto next - show "\\<^sup>+x\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (case x of (z,x,y) \ - x ^ r * y ^ s / (1-(1-x*y)*z)) \lborel < \" + show "(\\<^sup>+x\{0<..<1}\{0<..<1}\{0<..<1}. ennreal (case x of (z,x,y) \ + x ^ r * y ^ s / (1-(1-x*y)*z)) \lborel) < \" using beukers_nn_integral1_finite by (simp add: beukers_nn_integral1_altdef case_prod_unfold) qed (simp_all flip: lborel_prod add: set_borel_measurable_def) lemma beukers_integral1_conv_nn_integral: "beukers_integral1 r s = enn2real (beukers_nn_integral1 r s)" proof - have "ln (a * b) * a ^ r * b ^ s / (1 - a * b) \ 0" if "a \ {0<..<1}" "b \ {0<..<1}" for a b :: real using mult_strict_mono[of a 1 b 1] that by (intro divide_nonpos_nonneg mult_nonpos_nonneg) auto thus ?thesis unfolding beukers_integral1_def using beukers_nn_integral1_finite[of r s] by (intro set_integral_eq_nn_integral) (auto simp: case_prod_unfold beukers_nn_integral1_def set_borel_measurable_def simp flip: borel_prod intro!: AE_I2 intro: divide_nonpos_nonneg mult_nonpos_nonneg) qed lemma beukers_integral1_swap: "beukers_integral1 r s = beukers_integral1 s r" by (simp add: beukers_integral1_conv_nn_integral beukers_nn_integral1_swap) subsection \The second double integral\ context fixes n :: nat fixes D :: "(real \ real) set" and D' :: "(real \ real \ real) set" fixes P :: "real \ real" and Q :: "nat \ real \ real" defines "D \ {0<..<1} \ {0<..<1}" and "D' \ {0<..<1} \ {0<..<1} \ {0<..<1}" defines "Q \ Gen_Shleg n" and "P \ Shleg n" begin text \ The next integral to consider is the following variant of $I_1$: \[I_2 := \int_0^1 \int_0^1 -\frac{\ln(xy)}{1-xy} P_n(x) P_n(y)\,\text{d}x\,\text{d}y\ .\] \ definition beukers_integral2 :: real where "beukers_integral2 = (\(x,y)\D. (-ln (x*y) / (1-x*y) * P x * P y) \lborel)" text \ $I_2$ is simply a sum of integrals of type $I_1$, so using our results for $I_1$, we can write $I_2$ in the form $A \zeta(3) + \frac{B}{\text{lcm}\{1\ldots n\}^3}$ where $A$ and $B$ are integers and $A > 0$: \ lemma beukers_integral2_conv_int_combination: obtains A B :: int where "A > 0" and "beukers_integral2 = of_int A * Re (zeta 3) + of_int B / of_nat (Lcm {1..n} ^ 3)" proof - let ?I1 = "(\i. (i, i)) ` {..n}" let ?I2 = "Set.filter (\(i,j). i \ j) ({..n}\{..n})" let ?I3 = "Set.filter (\(i,j). i < j) ({..n}\{..n})" let ?I4 = "Set.filter (\(i,j). i > j) ({..n}\{..n})" define p where "p = shleg_poly n" define I where "I = (SIGMA i:{..n}. {1..i})" define J where "J = (SIGMA (i,j):?I4. {j<..i})" define h where "h = beukers_integral1" define A :: int where "A = (\i\n. 2 * poly.coeff p i ^ 2)" define B1 where "B1 = (\(i,k)\I. real_of_int (-2 * coeff p i ^ 2) / real_of_int (k ^ 3))" define B2 where "B2 = (\((i,j),k)\J. real_of_int (2 * coeff p i * coeff p j) / real_of_int (k^2*(i-j)))" define d where "d = Lcm {1..n} ^ 3" have [simp]: "h i j = h j i" for i j by (simp add: h_def beukers_integral1_swap) have "beukers_integral2 = (\(x,y)\D. (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * -ln (x*y) / (1-x*y) * x ^ i * y ^ j) \lborel)" unfolding beukers_integral2_def by (subst sum.cartesian_product [symmetric]) (simp add: poly_altdef P_def Shleg_def mult_ac case_prod_unfold p_def sum_distrib_left sum_distrib_right sum_negf sum_divide_distrib) also have "\ = (\(i,j)\{..n}\{..n}. coeff p i * coeff p j * h i j)" unfolding case_prod_unfold proof (subst set_integral_sum) fix ij :: "nat \ nat" have "set_integrable lborel D (\(x,y). real_of_int (coeff p (fst ij) * coeff p (snd ij)) * (-ln (x*y) / (1-x*y) * x ^ fst ij * y ^ snd ij))" unfolding case_prod_unfold using beukers_integral1_integrable[of "fst ij" "snd ij"] by (intro set_integrable_mult_right) (auto simp: D_def case_prod_unfold) thus "set_integrable lborel D (\pa. real_of_int (coeff p (fst ij) * coeff p (snd ij)) * -ln (fst pa * snd pa) / (1 - fst pa * snd pa) * fst pa ^ fst ij * snd pa ^ snd ij)" by (simp add: mult_ac case_prod_unfold) qed (auto simp: beukers_integral1_def h_def case_prod_unfold mult.assoc D_def simp flip: set_integral_mult_right) also have "\ = (\(i,j)\?I1\?I2. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I1. coeff p i * coeff p j * h i j) + (\(i,j)\?I2. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I1. coeff p i * coeff p j * h i j) = (\i\n. coeff p i ^ 2 * h i i)" by (subst sum.reindex) (auto intro: inj_onI simp: case_prod_unfold power2_eq_square) also have "\ = (\i\n. coeff p i ^ 2 * 2 * (Re (zeta 3) - (\k=1..i. 1 / k ^ 3)))" unfolding h_def D_def by (intro sum.cong refl, subst beukers_integral1_same) auto also have "\ = of_int A * Re (zeta 3) - (\i\n. 2 * coeff p i ^ 2 * (\k=1..i. 1 / k ^ 3))" by (simp add: sum_subtractf sum_distrib_left sum_distrib_right algebra_simps A_def) also have "\ = of_int A * Re (zeta 3) + B1" unfolding I_def B1_def by (subst sum.Sigma [symmetric]) (auto simp: sum_distrib_left sum_negf) also have "(\(i,j)\?I2. coeff p i * coeff p j * h i j) = (\(i,j)\?I3\?I4. coeff p i * coeff p j * h i j)" by (intro sum.cong) auto also have "\ = (\(i,j)\?I3. coeff p i * coeff p j * h i j) + (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.union_disjoint) auto also have "(\(i,j)\?I3. coeff p i * coeff p j * h i j) = (\(i,j)\?I4. coeff p i * coeff p j * h i j)" by (intro sum.reindex_bij_witness[of _ "\(i,j). (j,i)" "\(i,j). (j,i)"]) auto also have "\ + \ = 2 * \" by simp also have "\ = (\(i,j)\?I4. \k\{j<..i}. 2 * coeff p i * coeff p j / k ^ 2 / (i - j))" unfolding sum_distrib_left by (intro sum.cong refl) (auto simp: h_def beukers_integral1_different sum_divide_distrib sum_distrib_left mult_ac) also have "\ = B2" unfolding J_def B2_def by (subst sum.Sigma [symmetric]) (auto simp: case_prod_unfold) also have "\B1'. B1 = real_of_int B1' / real_of_int d" unfolding B1_def case_prod_unfold by (rule sum_rationals_common_divisor') (auto simp: d_def I_def) then obtain B1' where "B1 = real_of_int B1' / real_of_int d" .. also have "\B2'. B2 = real_of_int B2' / real_of_int d" unfolding B2_def case_prod_unfold J_def proof (rule sum_rationals_common_divisor'; clarsimp?) fix i j k :: nat assume ijk: "i \ n" "j < k" "k \ i" have "int (k ^ 2 * (i - j)) dvd int (Lcm {1..n} ^ 2 * Lcm {1..n})" unfolding int_dvd_int_iff using ijk by (intro mult_dvd_mono dvd_power_same dvd_Lcm) auto also have "\ = d" by (simp add: d_def power_numeral_reduce) finally show "int k ^ 2 * int (i - j) dvd int d" by simp qed(auto simp: d_def J_def intro!: Nat.gr0I) then obtain B2' where "B2 = real_of_int B2' / real_of_int d" .. finally have "beukers_integral2 = of_int A * Re (zeta 3) + of_int (B1' + B2') / of_nat (Lcm {1..n} ^ 3)" by (simp add: add_divide_distrib d_def) moreover have "coeff p 0 = P 0" unfolding P_def p_def Shleg_def by (simp add: poly_0_coeff_0) hence "coeff p 0 = 1" by (simp add: P_def) hence "A > 0" unfolding A_def by (intro sum_pos2[of _ 0]) auto ultimately show ?thesis by (intro that[of A "B1' + B2'"]) auto qed lemma beukers_integral2_integrable: "set_integrable lborel D (\(x,y). -ln (x*y) / (1 - x*y) * P x * P y)" proof - have "bounded (P ` {0..1})" unfolding P_def Shleg_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast have [measurable]: "P \ borel_measurable borel" by (simp add: P_def) from C[of 0] have "C \ 0" by simp show ?thesis proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C ^ 2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y :: real assume xy: "(x, y) \ D" from xy have "x * y < 1" using mult_strict_mono[of x 1 y 1] by (simp add: D_def) have "norm (-ln (x*y) / (1 - x*y) * P x * P y) = (-ln (x*y)) / (1 - x*y) * norm (P x) * norm (P y)" using xy \x * y < 1\ by (simp add: abs_mult abs_divide D_def) also have "\ \ (-ln (x*y)) / (1-x*y) * C * C" using xy C[of x] C[of y] \x * y < 1\ \C \ 0\ by (intro mult_mono divide_left_mono) (auto simp: D_def divide_nonpos_nonneg mult_nonpos_nonneg) also have "\ = norm ((-ln (x*y)) / (1-x*y) * C * C)" using xy \x * y < 1\ \C \ 0\ by (simp add: abs_divide abs_mult D_def) finally show "norm (-ln (x*y) / (1 - x*y) * P x * P y) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (auto simp: algebra_simps power2_eq_square abs_mult abs_divide) qed (auto simp: D_def set_borel_measurable_def case_prod_unfold simp flip: lborel_prod) qed subsection \The triple integral\ text \ Lastly, we turn to the triple integral \[I_3 := \int_0^1 \int_0^1 \int_0^1 \frac{(x(1-x)y(1-y)w(1-w))^n}{(1-(1-xy)w)^{n+1}}\ \text{d}x\,\text{d}y\,\text{d}w\ .\] \ definition beukers_nn_integral3 :: ennreal where "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" definition beukers_integral3 :: real where "beukers_integral3 = (\(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" text \ We first prove the following bound (which is a consequence of the arithmetic--geometric mean inequality) that will help us to bound the triple integral. \ lemma beukers_integral3_integrand_bound: fixes x y z :: real assumes xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" shows "(x*(1-x)*y*(1-y)*z*(1-z)) / (1-(1-x*y)*z) \ 1 / 27" (is "?lhs \ _") proof - have ineq1: "x * (1 - x) \ 1 / 4" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) - 1 / 4 = -((x - 1 / 2) ^ 2)" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by simp finally show ?thesis by simp qed have ineq2: "x * (1 - x) ^ 2 \ 4 / 27" if x: "x \ {0..1}" for x :: real proof - have "x * (1 - x) ^ 2 - 4 / 27 = (x - 4 / 3) * (x - 1 / 3) ^ 2" by (simp add: algebra_simps power2_eq_square) also have "\ \ 0" by (rule mult_nonpos_nonneg) (use x in auto) finally show ?thesis by simp qed have "1 - (1-x*y)*z = (1 - z) + x * y * z" by (simp add: algebra_simps) also have "\ \ 2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z" using arith_geo_mean_sqrt[of "1 - z" "x * y * z"] xyz by (auto simp: real_sqrt_mult) finally have *: "?lhs \ (x*(1-x)*y*(1-y)*z*(1-z)) / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z)" using xyz beukers_denom_ineq[of x y z] by (intro divide_left_mono mult_nonneg_nonneg mult_pos_pos) auto have "(x*(1-x)*y*(1-y)*z*(1-z)) = (sqrt x * sqrt x * (1-x) * sqrt y * sqrt y * (1-y) * sqrt z * sqrt z * sqrt (1-z) * sqrt (1-z))" using xyz by simp also have "\ / (2 * sqrt (1 - z) * sqrt x * sqrt y * sqrt z) = sqrt (x * (1 - x) ^ 2) * sqrt (y * (1 - y) ^ 2) * sqrt (z * (1 - z)) / 2" using xyz by (simp add: divide_simps real_sqrt_mult del: real_sqrt_mult_self) also have "\ \ sqrt (4 / 27) * sqrt (4 / 27) * sqrt (1 / 4) / 2" using xyz by (intro divide_right_mono mult_mono real_sqrt_le_mono ineq1 ineq2) auto also have "\ = 1 / 27" by (simp add: real_sqrt_divide) finally show ?thesis using * by argo qed text \ Connecting the above bound with our results of $I_1$, it is easy to see that $I_3 \leq 2 \cdot 27^{-n} \cdot \zeta(3)$: \ lemma beukers_nn_integral3_le: "beukers_nn_integral3 \ ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have "beukers_nn_integral3 = (\\<^sup>+(w,x,y)\D'. ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) \lborel)" by (simp add: beukers_nn_integral3_def) also have "\ \ (\\<^sup>+(w,x,y)\D'. ((1 / 27) ^ n / (1-(1-x*y)*w)) \lborel)" proof (intro set_nn_integral_mono ennreal_leI, clarify, goal_cases) case (1 w x y) have "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) = ((x*(1-x)*y*(1-y)*w*(1-w)) / (1-(1-x*y)*w)) ^ n / (1-(1-x*y)*w)" by (simp add: divide_simps) also have "\ \ (1 / 27) ^ n / (1 - (1 - x * y) * w)" using beukers_denom_ineq[of x y w] 1 by (intro divide_right_mono power_mono beukers_integral3_integrand_bound) (auto simp: D'_def) finally show ?case . qed also have "\ = ennreal ((1 / 27) ^ n) * (\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel)" unfolding lborel_prod [symmetric] case_prod_unfold by (subst nn_integral_cmult [symmetric]) (auto intro!: nn_integral_cong simp: indicator_def simp flip: ennreal_mult') also have "(\\<^sup>+(w,x,y)\D'. (1 / (1-(1-x*y)*w)) \lborel) = (\\<^sup>+(x,y)\{0<..<1}\{0<..<1}. ennreal (- (ln (x * y) / (1 - x * y))) \lborel)" using beukers_nn_integral1_altdef[of 0 0] by (simp add: beukers_nn_integral1_def D'_def case_prod_unfold) also have "\ = ennreal (2 * Re (zeta 3))" using beukers_nn_integral1_same[of 0 0] by (simp add: D_def beukers_nn_integral1_def) also have "ennreal ((1 / 27) ^ n) * \ = ennreal (2 * (1 / 27) ^ n * Re (zeta 3))" by (subst ennreal_mult' [symmetric]) (simp_all add: mult_ac) finally show ?thesis . qed lemma beukers_nn_integral3_finite: "beukers_nn_integral3 < \" by (rule le_less_trans, rule beukers_nn_integral3_le) simp_all lemma beukers_integral3_integrable: "set_integrable lborel D' (\(w,x,y). (x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1))" unfolding case_prod_unfold using less_imp_le[OF beukers_denom_ineq] beukers_nn_integral3_finite by (intro set_integrableI_nonneg AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod intro!: divide_nonneg_nonneg mult_nonneg_nonneg) lemma beukers_integral3_conv_nn_integral: "beukers_integral3 = enn2real beukers_nn_integral3" unfolding beukers_integral3_def using beukers_nn_integral3_finite less_imp_le[OF beukers_denom_ineq] by (intro set_integral_eq_nn_integral AE_I2 impI) (auto simp: D'_def set_borel_measurable_def beukers_nn_integral3_def case_prod_unfold simp flip: lborel_prod) lemma beukers_integral3_le: "beukers_integral3 \ 2 * (1 / 27) ^ n * Re (zeta 3)" proof - have "beukers_integral3 = enn2real beukers_nn_integral3" by (rule beukers_integral3_conv_nn_integral) also have "\ \ enn2real (ennreal (2 * (1 / 27) ^ n * Re (zeta 3)))" by (intro enn2real_mono beukers_nn_integral3_le) auto also have "\ = 2 * (1 / 27) ^ n * Re (zeta 3)" using Re_zeta_ge_1[of 3] by (intro enn2real_ennreal mult_nonneg_nonneg) auto finally show ?thesis . qed text \ It is also easy to see that $I_3 > 0$. \ lemma beukers_nn_integral3_pos: "beukers_nn_integral3 > 0" proof - have D' [measurable]: "D' \ sets (borel \\<^sub>M borel \\<^sub>M borel)" unfolding D'_def by (simp flip: borel_prod) have *: "\(AE (w,x,y) in lborel. ennreal ((x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1)) * indicator D' (w,x,y) \ 0)" (is "\(AE z in lborel. ?P z)") proof - { fix w x y :: real assume xyw: "(w,x,y) \ D'" hence "(x*(1-x)*y*(1-y)*w*(1-w))^n / (1-(1-x*y)*w)^(n+1) > 0" using beukers_denom_ineq[of x y w] by (intro divide_pos_pos mult_pos_pos zero_less_power) (auto simp: D'_def) with xyw have "\?P (w,x,y)" by (auto simp: indicator_def D'_def) } hence *: "\?P z" if "z \ D'" for z using that by blast hence "{z\space lborel. \?P z} = D'" by auto moreover have "emeasure lborel D' = 1" proof - have "D' = box (0,0,0) (1,1,1)" by (auto simp: D'_def box_def Basis_prod_def) also have "emeasure lborel \ = 1" by (subst emeasure_lborel_box) (auto simp: Basis_prod_def) finally show ?thesis by simp qed ultimately show ?thesis by (subst AE_iff_measurable[of D']) (simp_all flip: borel_prod) qed hence "nn_integral lborel (\_::real\real\real. 0) < beukers_nn_integral3" unfolding beukers_nn_integral3_def by (intro nn_integral_less) (simp_all add: case_prod_unfold flip: lborel_prod) thus ?thesis by simp qed lemma beukers_integral3_pos: "beukers_integral3 > 0" proof - have "0 < enn2real beukers_nn_integral3" using beukers_nn_integral3_pos beukers_nn_integral3_finite by (subst enn2real_positive_iff) auto also have "\ = beukers_integral3" by (rule beukers_integral3_conv_nn_integral [symmetric]) finally show ?thesis . qed subsection \Connecting the double and triple integral\ text \ In this section, we will prove the most technically involved part, namely that $I_2 = I_3$. I will not go into detail about how this works -- the reader is advised to simply look at Filaseta's presentation of the proof. The basic idea is to integrate by parts \n\ times with respect to \y\ to eliminate the factor $P(y)$, then change variables $z = \frac{1-w}{1-(1-xy)w}$, and then apply the same integration by parts \n\ times to \x\ to eliminate $P(x)$. \ text \ The first expand \[-\frac{\ln (xy)}{1-xy} = \int_0^1 \frac{1}{1-(1-xy)z}\,\text{d}z\ .\] \ lemma beukers_aux_ln_conv_integral: fixes x y :: real assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "-ln (x*y) / (1-x*y) = (LBINT z=0..1. 1 / (1-(1-x*y)*z))" proof - have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - x * y) * u < 1" if u: "u \ {0..1}" for u proof - from u \x * y < 1\ have "(1 - x * y) * u \ (1 - x * y) * 1" by (intro mult_left_mono) auto also have "\ < 1 * 1" using xy by (intro mult_strict_right_mono) auto finally show "(1 - x * y) * u < 1" by simp qed have neq: "(1 - x * y) * u \ 1" if "u \ {0..1}" for u using less[of u] that by simp let ?F = "\z. ln (1-(1-x*y)*z)/(x*y-1)" have "(LBINT z=ereal 0..ereal 1. 1 / (1-(1-x*y)*z)) = ?F 1 - ?F 0" proof (rule interval_integral_FTC_finite, goal_cases cont deriv) case cont show ?case using neq by (intro continuous_intros) auto next case (deriv z) show ?case unfolding has_real_derivative_iff_has_vector_derivative [symmetric] by (insert less[of z] xy \x * y < 1\ deriv) (rule derivative_eq_intros refl | simp)+ qed also have "\ = -ln (x*y) / (1-x*y)" using \x * y < 1\ by (simp add: field_simps) finally show ?thesis by (simp add: zero_ereal_def one_ereal_def) qed text \ The first part we shall show is the integration by parts. \ lemma beukers_aux_by_parts_aux: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" and "k \ n" shows "(LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z))) = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" using assms(3) proof (induction k) case (Suc k) note [derivative_intros] = DERIV_chain2[OF has_field_derivative_Gen_Shleg] define G where "G = (\y. -fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1))" define g where "g = (\y. fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))" have less: "(1 - x * y) * z < 1" and neq: "(1 - x * y) * z \ 1" if y: "y \ {0..1}" for y proof - from y xz have "x * y \ x * 1" by (intro mult_left_mono) auto also have "\ < 1" using xz by simp finally have "(1 - x * y) * z \ 1 * z" using xz y by (intro mult_right_mono) auto also have "\ < 1" using xz by simp finally show "(1 - x * y) * z < 1" by simp thus "(1 - x * y) * z \ 1" by simp qed have cont: "continuous_on {0..1} g" using neq by (auto simp: g_def intro!: continuous_intros) have deriv: "(G has_real_derivative g y) (at y within {0..1})" if y: "y \ {0..1}" for y unfolding G_def by (insert neq xz y, (rule derivative_eq_intros refl power_not_zero)+) (auto simp: divide_simps g_def) have deriv2: "(Q (n - Suc k) has_real_derivative Q (n - k) y) (at y within {0..1})" for y using Suc.prems by (auto intro!: derivative_eq_intros simp: Suc_diff_Suc Q_def) have "(LBINT y=0..1. Q (n-Suc k) y * (fact (Suc k) * (x*z)^Suc k / (1-(1-x*y)*z) ^ (k+2))) = (LBINT y=0..1. Q (n-Suc k) y * g y)" by (simp add: g_def) also have "(LBINT y=0..1. Q (n-Suc k) y * g y) = -(LBINT y=0..1. Q (n-k) y * G y)" using Suc.prems deriv deriv2 cont by (subst interval_lebesgue_integral_by_parts_01[where f = "Q (n-k)" and G = G]) (auto intro!: continuous_intros simp: Q_def) also have "\ = (LBINT y=0..1. Q (n-k) y * (fact k * (x*z)^k / (1-(1-x*y)*z) ^ (k+1)))" by (simp add: G_def flip: interval_lebesgue_integral_uminus) finally show ?case using Suc by simp qed auto lemma beukers_aux_by_parts: assumes xz: "x \ {0<..<1}" "z \ {0<..<1}" shows "(LBINT y=0..1. P y / (1-(1-x*y)*z)) = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have "(LBINT y=0..1. P y * (1/(1-(1-x*y)*z))) = 1 / fact n * (LBINT y=0..1. Q n y * (1/(1-(1-x*y)*z)))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: P_def Q_def Shleg_altdef) also have "\ = (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" by (subst beukers_aux_by_parts_aux[OF assms, of n], simp, subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: Q_def mult_ac Gen_Shleg_0_left power_mult_distrib) finally show ?thesis by simp qed text \ We then get the following by applying the integration by parts to \y\: \ lemma beukers_aux_integral_transform1: fixes z :: real assumes z: "z \ {0<..<1}" shows "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "1 - (1 - fst p * snd p) * z \ 0" by simp qed hence integrable: "set_integrable lborel (box (0,0) (1,1)) (\(x, y). P x * P y / (1 - (1 - x * y) * z))" by (rule set_integrable_subset) (auto simp: box simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0, 0) (1, 1)" have "(1 - fst p * snd p) * z \ 1 * z" using mult_mono[of "fst p" 1 "snd p" 1] p z cbox by (intro mult_right_mono) auto also have "\ < 1" using z by simp finally show "(1 - (1 - fst p * snd p) * z) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x, y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (rule set_integrable_subset) (auto simp: box D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = (LBINT x=0..1. (LBINT y=0..1. P x * P y / (1-(1-x*y)*z)))" unfolding D_def lborel_prod [symmetric] using box integrable by (subst lborel_pair.set_integral_fst') (simp_all add: interval_integral_Ioo lborel_prod) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT x=0..1. P x * (LBINT y=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using z by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) auto also have "\ = (LBINT x=0..1. (LBINT y=0..1. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac) also have "\ = (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding D_def lborel_prod [symmetric] using box integrable' by (subst lborel_pair.set_integral_fst') (simp_all add: D_def interval_integral_Ioo lborel_prod) finally show "(LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)) = \" . qed text \ We then change variables for \z\: \ lemma beukers_aux_integral_transform2: assumes xy: "x \ {0<..<1}" "y \ {0<..<1}" shows "(LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - define g where "g = (\z. (1 - z) / (1-(1-x*y)*z))" define g' where "g' = (\z. -x*y / (1-(1-x*y)*z)^2)" have "x * y < 1" using mult_strict_mono[of x 1 y 1] xy by simp have less: "(1 - (x * y)) * w < 1" and neq: "(1 - (x * y)) * w \ 1" if w: "w \ {0..1}" for w proof - have "(1 - (x * y)) * w \ (1 - (x * y)) * 1" using w \x * y < 1\ by (intro mult_left_mono) auto also have "\ < 1" using xy by simp finally show "(1 - (x * y)) * w < 1" . thus "(1 - (x * y)) * w \ 1" by simp qed have deriv: "(g has_real_derivative g' w) (at w within {0..1})" if "w \ {0..1}" for w unfolding g_def g'_def apply (insert that xy neq) apply (rule derivative_eq_intros refl)+ apply (simp_all add: divide_simps power2_eq_square) apply (auto simp: algebra_simps) done have "continuous_on {0..1} (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" using neq by (auto intro!: continuous_intros) moreover have "g ` {0..1} \ {0..1}" proof clarify fix w :: real assume w: "w \ {0..1}" have "(1 - x * y) * w \ 1 * w" using \x * y < 1\ xy w by (intro mult_right_mono) auto thus "g w \ {0..1}" unfolding g_def using less[of w] w by (auto simp: divide_simps) qed ultimately have cont: "continuous_on (g ` {0..1}) (\xa. (1 - xa) ^ n / (1 - (1 - x * y) * xa))" by (rule continuous_on_subset) have cont': "continuous_on {0..1} g'" using neq by (auto simp: g'_def intro!: continuous_intros) have "(LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = (1-y)^n * (LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) also have "(LBINT w=0..1. (1 - w)^n / (1-(1-x*y)*w)) = -(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w))" by (subst interval_integral_endpoints_reverse)(simp add: g_def zero_ereal_def one_ereal_def) also have "(LBINT w=g 0..g 1. (1 - w)^n / (1-(1-x*y)*w)) = (LBINT w=0..1. g' w * ((1 - g w) ^ n / (1 - (1-x*y) * g w)))" using deriv cont cont' by (subst interval_integral_substitution_finite [symmetric, where g = g and g' = g']) (simp_all add: zero_ereal_def one_ereal_def) also have "-\ = (LBINT z=0..1. ((x*y)^n * z^n / (1-(1-x*y)*z)^(n+1)))" unfolding interval_lebesgue_integral_uminus [symmetric] using xy apply (intro interval_lebesgue_integral_lborel_01_cong) apply (simp add: divide_simps g_def g'_def) apply (auto simp: algebra_simps power_mult_distrib power2_eq_square) done also have "(1-y)^n * \ = (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" unfolding interval_lebesgue_integral_mult_right [symmetric] by (simp add: algebra_simps power_mult_distrib) finally show "\ = (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" .. qed text \ Lastly, we apply the same integration by parts to \x\: \ lemma beukers_aux_integral_transform3: assumes w: "w \ {0<..<1}" shows "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" proof - have cbox: "cbox (0, 0) (1, 1) = ({0..1} \ {0..1} :: (real \ real) set)" by (auto simp: cbox_def Basis_prod_def inner_prod_def) have box: "box (0, 0) (1, 1) = ({0<..<1} \ {0<..<1} :: (real \ real) set)" by (auto simp: box_def Basis_prod_def inner_prod_def) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "1 - (1 - fst p * snd p) * w \ 0" by simp qed hence integrable: "set_integrable lborel D (\(x,y). P x * (1-y)^n/(1-(1-x*y)*w))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "set_integrable lborel (cbox (0,0) (1,1)) (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" unfolding lborel_prod case_prod_unfold P_def proof (intro continuous_on_imp_set_integrable_cbox continuous_intros ballI) fix p :: "real \ real" assume p: "p \ cbox (0,0) (1,1)" have "(1 - fst p * snd p) * w \ 1 * w" using p cbox w by (intro mult_right_mono) auto also have "\ < 1" using w by simp finally have "(1 - fst p * snd p) * w < 1" by simp thus "(1 - (1 - fst p * snd p) * w) ^ (n + 1) \ 0" by simp qed hence integrable': "set_integrable lborel D (\(x,y). (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" by (rule set_integrable_subset) (auto simp: D_def simp flip: borel_prod) have "(LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)) = (LBINT y=0..1. (LBINT x=0..1. P x * (1-y)^n / (1-(1-x*y)*w)))" using integrable unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. P x / (1-(1-y*x)*w)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (1-y)^n * (LBINT x=0..1. (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" using w by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_by_parts) (auto simp: mult_ac) also have "\ = (LBINT y=0..1. (LBINT x=0..1. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (auto simp: mult_ac) also have "\ = (LBINT (x,y):D. (1-y)^n * (x*y*w)^n * (1-x)^n / (1-(1-x*y)*w)^(n+1))" using integrable' unfolding case_prod_unfold D_def lborel_prod [symmetric] by (subst lborel_pair.set_integral_snd) (auto simp: interval_integral_Ioo) finally show ?thesis . qed text \ We need to show the existence of some of these triple integrals. \ lemma beukers_aux_integrable1: "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x,y),z). P x * P y / (1-(1-x*y)*z))" proof - have D [measurable]: "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp flip: borel_prod) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def case_prod_unfold proof (subst lborel_prod [symmetric], intro lborel_pair.Fubini_set_integrable AE_I2 impI; clarsimp?) fix x y :: real assume xy: "x > 0" "x < 1" "y > 0" "y < 1" have "x * y < 1" using xy mult_strict_mono[of x 1 y 1] by simp show "set_integrable lborel {0<..<1} (\z. P x * P y / (1-(1-x*y)*z))" by (rule set_integrable_subset[of _ "{0..1}"], rule borel_integrable_atLeastAtMost') (use \x*y<1\ beukers_denom_neq[of x y] xy in \auto intro!: continuous_intros simp: P_def\) next have "set_integrable lborel D (\(x,y). (\z\{0<..<1}. norm (P x * P y / (1-(1-x*y)*z)) \lborel))" proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel D (\(x,y). C\<^sup>2 * (-ln (x*y) / (1 - x*y)))" using beukers_integral1_integrable[of 0 0] unfolding case_prod_unfold by (intro set_integrable_mult_right) (auto simp: D_def) next fix x y assume xy: "(x, y) \ D" have "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) = norm (LBINT z:{0<..<1}. \P x\ * \P y\ * (1 / (1-(1-x*y)*z)))" proof (intro arg_cong[where f = norm] set_lebesgue_integral_cong allI impI, goal_cases) case (2 z) with beukers_denom_ineq[of x y z] xy show ?case by (auto simp: abs_mult D_def) qed (auto simp: abs_mult D_def) also have "\ = norm (\P x\ * \P y\ * (LBINT z=0..1. (1 / (1-(1-x*y)*z))))" by (subst set_integral_mult_right) (auto simp: interval_integral_Ioo) also have "\ = norm (norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)))" using beukers_aux_ln_conv_integral[of x y] xy by (simp add: D_def) also have "\ = norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y))" using xy mult_strict_mono[of x 1 y 1] by (auto simp: D_def divide_nonpos_nonneg abs_mult) also have "norm (P x) * norm (P y) * (- ln (x * y) / (1 - x * y)) \ norm (C * C * (- ln (x * y) / (1 - x * y)))" using xy C[of x] C[of y] mult_strict_mono[of x 1 y 1] unfolding norm_mult norm_divide by (intro mult_mono C) (auto simp: D_def divide_nonpos_nonneg) finally show "norm (LBINT z:{0<..<1}. norm (P x * P y / (1-(1-x*y)*z))) \ norm (case (x, y) of (x, y) \ C\<^sup>2 * (- ln (x * y) / (1 - x * y)))" by (simp add: power2_eq_square mult_ac) next show "set_borel_measurable lborel D (\(x, y). LBINT z:{0<..<1}. norm (P x * P y / (1 - (1 - x * y) * z)))" unfolding lborel_prod [symmetric] set_borel_measurable_def case_prod_unfold set_lebesgue_integral_def P_def by measurable qed thus "set_integrable lborel ({0<..<1} \ {0<..<1}) (\x. LBINT y:{0<..<1}. \P (fst x) * P (snd x)\ / \1 - (1 - fst x * snd x) * y\)" by (simp add: case_prod_unfold D_def) qed (auto simp: case_prod_unfold lborel_prod [symmetric] set_borel_measurable_def P_def) qed lemma beukers_aux_integrable2: "set_integrable lborel D' (\(z,x,y). P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z) ^ (n+1))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y z :: real assume xyz: "x \ {0<..<1}" "y \ {0<..<1}" "z \ {0<..<1}" have "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = norm (P x) * (1-y)^n * ((x*y*z) / (1-(1-x*y)*z))^n / (1-(1-x*y)*z)" using xyz beukers_denom_ineq[of x y z] by (simp add: abs_mult power_divide mult_ac) also have "(x*y*z) / (1-(1-x*y)*z) = 1/((1-z)/(z*x*y)+1)" using xyz by (simp add: field_simps) also have "norm (P x) * (1-y)^n * \^n / (1-(1-x*y)*z) \ C * 1^n * 1^n / (1-(1-x*y)*z)" using xyz C[of x] beukers_denom_ineq[of x y z] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*z)\" by linarith finally show "norm (P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) \ norm (case (z,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed lemma beukers_aux_integrable3: "set_integrable lborel D' (\(w,x,y). P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" proof - have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "bounded (P ` {0..1})" unfolding P_def by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then obtain C where C: "\x. x \ {0..1} \ norm (P x) \ C" unfolding bounded_iff by fast show ?thesis unfolding D'_def proof (rule set_integrable_bound[OF _ _ AE_I2]; clarify?) show "set_integrable lborel ({0<..<1} \ {0<..<1} \ {0<..<1}) (\(z,x,y). C * (1 / (1-(1-x*y)*z)))" unfolding case_prod_unfold using beukers_integral1_integrable'[of 0 0] by (intro set_integrable_mult_right) (auto simp: lborel_prod case_prod_unfold) next fix x y w :: real assume xyw: "x \ {0<..<1}" "y \ {0<..<1}" "w \ {0<..<1}" have "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) = norm (P x) * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)" using xyw beukers_denom_ineq[of x y w] by (simp add: abs_mult power_divide mult_ac) also have "\ \ C * 1^n * 1^n / (1-(1-x*y)*w)" using xyw C[of x] beukers_denom_ineq[of x y w] by (intro mult_mono divide_right_mono power_mono zero_le_power mult_nonneg_nonneg divide_nonneg_nonneg) (auto simp: field_simps) also have "\ \ \C * 1^n * 1^n / (1-(1-x*y)*w)\" by linarith finally show "norm (P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)) \ norm (case (w,x,y) of (z,x,y) \ C * (1 / (1-(1-x*y)*z)))" by (simp add: case_prod_unfold) qed (simp_all add: case_prod_unfold set_borel_measurable_def flip: borel_prod) qed text \ Now we only need to put all of these results together: \ lemma beukers_integral2_conv_3: "beukers_integral2 = beukers_integral3" proof - have cont_P: "continuous_on {0..1} P" by (auto simp: P_def intro: continuous_intros) have D [measurable]: "D \ sets borel" "D \ sets (borel \\<^sub>M borel)" unfolding D_def by (simp_all flip: borel_prod) have [measurable]: "P \ borel_measurable borel" unfolding P_def by (intro borel_measurable_continuous_onI continuous_intros) have "beukers_integral2 = (LBINT (x,y):D. P x * P y * (LBINT z=0..1. 1 / (1-(1-x*y)*z)))" unfolding beukers_integral2_def case_prod_unfold by (intro set_lebesgue_integral_cong allI impI, measurable) (subst beukers_aux_ln_conv_integral, auto simp: D_def) also have "\ = (LBINT (x,y):D. (LBINT z=0..1. P x * P y / (1-(1-x*y)*z)))" by (subst interval_lebesgue_integral_mult_right [symmetric]) auto also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * P y / (1-(1-x*y)*z)))" by (simp add: interval_integral_Ioo) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * P y / (1-(1-x*y)*z)))" proof (subst lborel_pair.Fubini_set_integral [symmetric]) have "set_integrable lborel (({0<..<1} \ {0<..<1}) \ {0<..<1}) (\((x, y), z). P x * P y / (1 - (1 - x * y) * z))" using beukers_aux_integrable1 by simp also have "?this \ set_integrable (lborel \\<^sub>M lborel) ({0<..<1} \ D) (\(z,x,y). P x * P y / (1 - (1 - x * y) * z))" unfolding set_integrable_def by (subst lborel_pair.integrable_product_swap_iff [symmetric], intro Bochner_Integration.integrable_cong) (auto simp: indicator_def case_prod_unfold lborel_prod D_def) finally show \ . qed (auto simp: case_prod_unfold) also have "\ = (LBINT z:{0<..<1}. (LBINT (x,y):D. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" by (rule set_lebesgue_integral_cong) (use beukers_aux_integral_transform1 in simp_all) also have "\ = (LBINT (x,y):D. (LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)))" using beukers_aux_integrable2 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT (x,y):D. (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" proof (intro set_lebesgue_integral_cong allI impI; clarify?) fix x y :: real assume xy: "(x, y) \ D" have "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = P x * (LBINT z=0..1. (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = P x * (LBINT w=0..1. (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" using xy by (subst beukers_aux_integral_transform2) (auto simp: D_def) also have "\ = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" by (subst interval_lebesgue_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) finally show "(LBINT z:{0<..<1}. P x * (x*y*z)^n * (1-y)^n / (1-(1-x*y)*z)^(n+1)) = (LBINT w:{0<..<1}. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w))" . qed (auto simp: D_def simp flip: borel_prod) also have "\ = (LBINT w:{0<..<1}. (LBINT (x,y):D. P x * (1-w)^n * (1-y)^n / (1-(1-x*y)*w)))" using beukers_aux_integrable3 by (subst lborel_pair.Fubini_set_integral [symmetric]) (auto simp: case_prod_unfold lborel_prod D_def D'_def) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. P x * (1-y)^n / (1-(1-x*y)*w)))" unfolding case_prod_unfold by (subst set_integral_mult_right [symmetric]) (simp add: mult_ac interval_integral_Ioo) also have "\ = (LBINT w=0..1. (1-w)^n * (LBINT (x,y):D. (x*y*w*(1-x)*(1-y))^n / (1-(1-x*y)*w)^(n+1)))" by (intro interval_lebesgue_integral_lborel_01_cong, subst beukers_aux_integral_transform3) (auto simp: mult_ac power_mult_distrib) also have "\ = (LBINT w=0..1. (LBINT (x,y):D. (x*y*w*(1-x)*(1-y)*(1-w))^n / (1-(1-x*y)*w)^(n+1)))" by (subst set_integral_mult_right [symmetric]) (auto simp: case_prod_unfold mult_ac power_mult_distrib) also have "\ = beukers_integral3" using beukers_integral3_integrable unfolding D'_def D_def beukers_integral3_def by (subst (2) lborel_prod [symmetric], subst lborel_pair.set_integral_fst') (auto simp: case_prod_unfold interval_integral_Ioo lborel_prod algebra_simps) finally show ?thesis . qed subsection \The main result\ text \ Combining all of the results so far, we can derive the key inequalities \[0 < A\zeta(3) + B < 2 \zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\] for integers $A$, $B$ with $A > 0$. \ lemma zeta_3_linear_combination_bounds: obtains A B :: int where "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" proof - define I where "I = beukers_integral2" define d where "d = Lcm {1..n} ^ 3" have "d > 0" by (auto simp: d_def intro!: Nat.gr0I) from beukers_integral2_conv_int_combination obtain A' B :: int where *: "A' > 0" "I = A' * Re (zeta 3) + B / d" unfolding I_def d_def . define A where "A = A' * d" from * have A: "A > 0" "I = (A * Re (zeta 3) + B) / d" using \d > 0\ by (simp_all add: A_def field_simps) have "0 < I" using beukers_integral3_pos by (simp add: I_def beukers_integral2_conv_3) with \d > 0\ have "A * Re (zeta 3) + B > 0" by (simp add: field_simps A) moreover have "I \ 2 * (1 / 27) ^ n * Re (zeta 3)" using beukers_integral2_conv_3 beukers_integral3_le by (simp add: I_def) hence "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * d / 27 ^ n" using \d > 0\ by (simp add: A field_simps) ultimately show ?thesis using A by (intro that[of A B]) (auto simp: d_def) qed text \ If $\zeta(3) = \frac{a}{b}$ for some integers $a$ and $b$, we can thus derive the inequality $2b\zeta(3) \cdot 27^{-n} \cdot \text{lcm}\{1\ldots n\}^3\geq 1$ for any natural number $n$. \ lemma beukers_key_inequality: fixes a :: int and b :: nat assumes "b > 0" and ab: "Re (zeta 3) = a / b" shows "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" proof - from zeta_3_linear_combination_bounds obtain A B :: int where AB: "A > 0" "A * Re (zeta 3) + B \ {0 <.. 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n}" . from AB have "0 < (A * Re (zeta 3) + B) * b" using \b > 0\ by (intro mult_pos_pos) auto also have "\ = A * (Re (zeta 3) * b) + B * b" using \b > 0\ by (simp add: algebra_simps) also have "Re (zeta 3) * b = a" using \b > 0\ by (simp add: ab) also have "of_int A * of_int a + of_int (B * b) = of_int (A * a + B * b)" by simp finally have "1 \ A * a + B * b" by linarith hence "1 \ real_of_int (A * a + B * b)" by linarith also have "\ = (A * (a / b) + B) * b" using \b > 0\ by (simp add: ring_distribs) also have "a / b = Re (zeta 3)" by (simp add: ab) also have "A * Re (zeta 3) + B \ 2 * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n" using AB by simp finally show "2 * b * Re (zeta 3) * Lcm {1..n} ^ 3 / 27 ^ n \ 1" using \b > 0\ by (simp add: mult_ac) qed end (* TODO: Move *) lemma smallo_power: "n > 0 \ f \ o[F](g) \ (\x. f x ^ n) \ o[F](\x. g x ^ n)" by (induction n rule: nat_induct_non_zero) (auto intro: landau_o.small.mult) text \ This is now a contradiction, since $\text{lcm}\{1\ldots n\} \in o(3^n)$ by the Prime Number Theorem -- hence the main result. \ theorem zeta_3_irrational: "zeta 3 \ \" proof assume "zeta 3 \ \" obtain a :: int and b :: nat where "b > 0" and ab': "zeta 3 = a / b" proof - from \zeta 3 \ \\ obtain r where r: "zeta 3 = of_rat r" by (elim Rats_cases) also have "r = rat_of_int (fst (quotient_of r)) / rat_of_int (snd (quotient_of r))" by (intro quotient_of_div) auto also have "of_rat \ = (of_int (fst (quotient_of r)) / of_int (snd (quotient_of r)) :: complex)" by (simp add: of_rat_divide) also have "of_int (snd (quotient_of r)) = of_nat (nat (snd (quotient_of r)))" using quotient_of_denom_pos'[of r] by auto finally have "zeta 3 = of_int (fst (quotient_of r)) / of_nat (nat (snd (quotient_of r)))" . thus ?thesis using quotient_of_denom_pos'[of r] by (intro that[of "nat (snd (quotient_of r))" "fst (quotient_of r)"]) auto qed hence ab: "Re (zeta 3) = a / b" by simp interpret prime_number_theorem by standard (rule prime_number_theorem) have Lcm_upto_smallo: "(\n. real (Lcm {1..n})) \ o(\n. c ^ n)" if c: "c > exp 1" for c proof - have "0 < exp (1::real)" by simp also note c finally have "c > 0" . have "(\n. real (Lcm {1..n})) = (\n. real (Lcm {1..nat \real n\}))" by simp also have "\ \ o(\n. c powr real n)" using Lcm_upto.smallo' by (rule landau_o.small.compose) (simp_all add: c filterlim_real_sequentially) also have "(\n. c powr real n) = (\n. c ^ n)" using c \c > 0\ by (subst powr_realpow) auto finally show ?thesis . qed have "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ O(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n)" using \b > 0\ Re_zeta_ge_1[of 3] by simp also have "exp 1 < (3 :: real)" using e_approx_32 by (simp add: abs_if split: if_splits) hence "(\n. real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\n. (3 ^ n) ^ 3 / 27 ^ n)" unfolding of_nat_power by (intro landau_o.small.divide_right smallo_power Lcm_upto_smallo) auto also have "(\n. (3 ^ n) ^ 3 / 27 ^ n :: real) = (\_. 1)" by (simp add: power_mult [of 3, symmetric] mult.commute[of _ 3] power_mult[of _ 3]) finally have *: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ o(\_. 1)" . have lim: "(\n. 2 * b * Re (zeta 3) * real (Lcm {1..n}) ^ 3 / 27 ^ n) \ 0" using smalloD_tendsto[OF *] by simp moreover have "1 \ real (2 * b) * Re (zeta 3) * real (Lcm {1..n} ^ 3) / 27 ^ n" for n using beukers_key_inequality[of b a] ab \b > 0\ by simp ultimately have "1 \ (0 :: real)" by (intro tendsto_lowerbound[OF lim] always_eventually allI) auto thus False by simp qed end \ No newline at end of file