diff -r cb27a55d868a -r 670063003ad3 src/HOL/Probability/Weak_Convergence.thy --- a/src/HOL/Probability/Weak_Convergence.thy Fri Feb 19 12:25:57 2016 +0100 +++ b/src/HOL/Probability/Weak_Convergence.thy Tue Feb 09 09:21:10 2016 +0100 @@ -27,11 +27,11 @@ section \Skorohod's theorem\ -locale right_continuous_mono = +locale right_continuous_mono = fixes f :: "real \ real" and a b :: real assumes cont: "\x. continuous (at_right x) f" assumes mono: "mono f" - assumes bot: "(f \ a) at_bot" + assumes bot: "(f \ a) at_bot" assumes top: "(f \ b) at_top" begin @@ -47,7 +47,7 @@ by (auto intro!: bdd_belowI[of _ y] elim: mono_invE[OF _ less_le_trans]) have ne: "?F \ {}" - using order_tendstoD(1)[OF top \\ < b\] + using order_tendstoD(1)[OF top \\ < b\] by (auto dest!: eventually_happens'[OF trivial_limit_at_top_linorder] intro: less_imp_le) show "\ \ f x \ I \ \ x" @@ -115,14 +115,14 @@ context fixes \ :: "nat \ real measure" and M :: "real measure" - assumes \: "\n. real_distribution (\ n)" + assumes \: "\n. real_distribution (\ n)" assumes M: "real_distribution M" assumes \_to_M: "weak_conv_m \ M" -begin +begin (* state using obtains? *) theorem Skorohod: - "\ (\ :: real measure) (Y_seq :: nat \ real \ real) (Y :: real \ real). + "\ (\ :: real measure) (Y_seq :: nat \ real \ real) (Y :: real \ real). prob_space \ \ (\n. Y_seq n \ measurable \ borel) \ (\n. distr \ borel (Y_seq n) = \ n) \ @@ -147,7 +147,7 @@ have Y_distr: "distr ?\ borel M.I = M" by (rule M.distr_I_eq_M) - have Y_cts_cnv: "(\n. \.I n \) \ M.I \" + have Y_cts_cnv: "(\n. \.I n \) \ M.I \" if \: "\ \ {0<..<1}" "isCont M.I \" for \ :: real proof (intro limsup_le_liminf_real) show "liminf (\n. \.I n \) \ M.I \" @@ -196,7 +196,7 @@ using \y < B\ by (intro Limsup_bounded[rotated]) (auto intro: le_less_trans elim: eventually_mono) qed simp - + have **: "(M.I \ ereal (M.I \)) (at_right \)" using \(2) by (auto intro: tendsto_within_subset simp: continuous_at) show "limsup (\n. \.I n \) \ M.I \" @@ -247,13 +247,13 @@ \ theorem weak_conv_imp_bdd_ae_continuous_conv: - fixes + fixes f :: "real \ 'a::{banach, second_countable_topology}" - assumes + assumes discont_null: "M ({x. \ isCont f x}) = 0" and f_bdd: "\x. norm (f x) \ B" and [measurable]: "f \ borel_measurable borel" - shows + shows "(\ n. integral\<^sup>L (\ n) f) \ integral\<^sup>L M f" proof - have "0 \ B" @@ -278,10 +278,10 @@ theorem weak_conv_imp_integral_bdd_continuous_conv: fixes f :: "real \ 'a::{banach, second_countable_topology}" - assumes + assumes "\x. isCont f x" and "\x. norm (f x) \ B" - shows + shows "(\ n. integral\<^sup>L (\ n) f) \ integral\<^sup>L M f" using assms by (intro weak_conv_imp_bdd_ae_continuous_conv) @@ -294,7 +294,7 @@ proof - interpret M: real_distribution M by fact interpret \: real_distribution "\ n" for n by fact - + have "(\n. (\x. indicator A x \\ n) :: real) \ (\x. indicator A x \M)" by (intro weak_conv_imp_bdd_ae_continuous_conv[where B=1]) (auto intro: assms simp: isCont_indicator) @@ -347,9 +347,9 @@ context fixes M_seq :: "nat \ real measure" and M :: "real measure" - assumes distr_M_seq [simp]: "\n. real_distribution (M_seq n)" + assumes distr_M_seq [simp]: "\n. real_distribution (M_seq n)" assumes distr_M [simp]: "real_distribution M" -begin +begin theorem continuity_set_conv_imp_weak_conv: fixes f :: "real \ real" @@ -364,9 +364,9 @@ theorem integral_cts_step_conv_imp_weak_conv: assumes integral_conv: "\x y. x < y \ (\n. integral\<^sup>L (M_seq n) (cts_step x y)) \ integral\<^sup>L M (cts_step x y)" shows "weak_conv_m M_seq M" - unfolding weak_conv_m_def weak_conv_def + unfolding weak_conv_m_def weak_conv_def proof (clarsimp) - interpret real_distribution M by (rule distr_M) + interpret real_distribution M by (rule distr_M) fix x assume "isCont (cdf M) x" hence left_cont: "continuous (at_left x) (cdf M)" unfolding continuous_at_split .. @@ -400,13 +400,13 @@ by (subst eventually_at_left[of "x - 1"]) (auto simp: ** intro: exI [of _ "x-1"]) qed (insert left_cont, auto simp: continuous_within) ultimately show "(\n. cdf (M_seq n) x) \ cdf M x" - by (elim limsup_le_liminf_real) + by (elim limsup_le_liminf_real) qed theorem integral_bdd_continuous_conv_imp_weak_conv: - assumes + assumes "\f. (\x. isCont f x) \ (\x. abs (f x) \ 1) \ (\n. integral\<^sup>L (M_seq n) f::real) \ integral\<^sup>L M f" - shows + shows "weak_conv_m M_seq M" apply (rule integral_cts_step_conv_imp_weak_conv [OF assms]) apply (rule continuous_on_interior)