author | paulson <lp15@cam.ac.uk> |
Fri, 03 Nov 2023 16:20:06 +0000 | |
changeset 78890 | d8045bc0544e |
parent 78517 | 28c1f4f5335f |
child 79599 | 2c18ac57e92e |
permissions | -rw-r--r-- |
63329 | 1 |
(* Title: HOL/Probability/Levy.thy |
2 |
Authors: Jeremy Avigad (CMU) |
|
62083 | 3 |
*) |
4 |
||
5 |
section \<open>The Levy inversion theorem, and the Levy continuity theorem.\<close> |
|
6 |
||
7 |
theory Levy |
|
8 |
imports Characteristic_Functions Helly_Selection Sinc_Integral |
|
9 |
begin |
|
10 |
||
11 |
subsection \<open>The Levy inversion theorem\<close> |
|
12 |
||
13 |
(* Actually, this is not needed for us -- but it is useful for other purposes. (See Billingsley.) *) |
|
14 |
lemma Levy_Inversion_aux1: |
|
15 |
fixes a b :: real |
|
16 |
assumes "a \<le> b" |
|
63589 | 17 |
shows "((\<lambda>t. (iexp (-(t * a)) - iexp (-(t * b))) / (\<i> * t)) \<longlongrightarrow> b - a) (at 0)" |
62083 | 18 |
(is "(?F \<longlongrightarrow> _) (at _)") |
19 |
proof - |
|
20 |
have 1: "cmod (?F t - (b - a)) \<le> a^2 / 2 * abs t + b^2 / 2 * abs t" if "t \<noteq> 0" for t |
|
21 |
proof - |
|
22 |
have "cmod (?F t - (b - a)) = cmod ( |
|
63589 | 23 |
(iexp (-(t * a)) - (1 + \<i> * -(t * a))) / (\<i> * t) - |
24 |
(iexp (-(t * b)) - (1 + \<i> * -(t * b))) / (\<i> * t))" |
|
25 |
(is "_ = cmod (?one / (\<i> * t) - ?two / (\<i> * t))") |
|
63167 | 26 |
using \<open>t \<noteq> 0\<close> by (intro arg_cong[where f=norm]) (simp add: field_simps) |
63589 | 27 |
also have "\<dots> \<le> cmod (?one / (\<i> * t)) + cmod (?two / (\<i> * t))" |
62083 | 28 |
by (rule norm_triangle_ineq4) |
63589 | 29 |
also have "cmod (?one / (\<i> * t)) = cmod ?one / abs t" |
62083 | 30 |
by (simp add: norm_divide norm_mult) |
63589 | 31 |
also have "cmod (?two / (\<i> * t)) = cmod ?two / abs t" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
32 |
by (simp add: norm_divide norm_mult) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
33 |
also have "cmod ?one / abs t + cmod ?two / abs t \<le> |
62083 | 34 |
((- (a * t))^2 / 2) / abs t + ((- (b * t))^2 / 2) / abs t" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
35 |
using iexp_approx1 [of "-(t * _)" 1] |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
36 |
by (intro add_mono divide_right_mono abs_ge_zero) (auto simp: field_simps eval_nat_numeral) |
62083 | 37 |
also have "\<dots> = a^2 / 2 * abs t + b^2 / 2 * abs t" |
63167 | 38 |
using \<open>t \<noteq> 0\<close> apply (case_tac "t \<ge> 0", simp add: field_simps power2_eq_square) |
39 |
using \<open>t \<noteq> 0\<close> by (subst (1 2) abs_of_neg, auto simp add: field_simps power2_eq_square) |
|
62083 | 40 |
finally show "cmod (?F t - (b - a)) \<le> a^2 / 2 * abs t + b^2 / 2 * abs t" . |
41 |
qed |
|
42 |
show ?thesis |
|
43 |
apply (rule LIM_zero_cancel) |
|
44 |
apply (rule tendsto_norm_zero_cancel) |
|
45 |
apply (rule real_LIM_sandwich_zero [OF _ _ 1]) |
|
46 |
apply (auto intro!: tendsto_eq_intros) |
|
47 |
done |
|
48 |
qed |
|
49 |
||
50 |
lemma Levy_Inversion_aux2: |
|
51 |
fixes a b t :: real |
|
52 |
assumes "a \<le> b" and "t \<noteq> 0" |
|
63589 | 53 |
shows "cmod ((iexp (t * b) - iexp (t * a)) / (\<i> * t)) \<le> b - a" (is "?F \<le> _") |
62083 | 54 |
proof - |
63589 | 55 |
have "?F = cmod (iexp (t * a) * (iexp (t * (b - a)) - 1) / (\<i> * t))" |
63167 | 56 |
using \<open>t \<noteq> 0\<close> by (intro arg_cong[where f=norm]) (simp add: field_simps exp_diff exp_minus) |
62083 | 57 |
also have "\<dots> = cmod (iexp (t * (b - a)) - 1) / abs t" |
65064
a4abec71279a
Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents:
64284
diff
changeset
|
58 |
unfolding norm_divide norm_mult norm_exp_i_times using \<open>t \<noteq> 0\<close> |
62083 | 59 |
by (simp add: complex_eq_iff norm_mult) |
60 |
also have "\<dots> \<le> abs (t * (b - a)) / abs t" |
|
61 |
using iexp_approx1 [of "t * (b - a)" 0] |
|
62 |
by (intro divide_right_mono) (auto simp add: field_simps eval_nat_numeral) |
|
63 |
also have "\<dots> = b - a" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
64 |
using assms by (auto simp add: abs_mult) |
62083 | 65 |
finally show ?thesis . |
66 |
qed |
|
67 |
||
68 |
(* TODO: refactor! *) |
|
69 |
theorem (in real_distribution) Levy_Inversion: |
|
70 |
fixes a b :: real |
|
71 |
assumes "a \<le> b" |
|
72 |
defines "\<mu> \<equiv> measure M" and "\<phi> \<equiv> char M" |
|
73 |
assumes "\<mu> {a} = 0" and "\<mu> {b} = 0" |
|
63589 | 74 |
shows "(\<lambda>T. 1 / (2 * pi) * (CLBINT t=-T..T. (iexp (-(t * a)) - iexp (-(t * b))) / (\<i> * t) * \<phi> t)) |
62083 | 75 |
\<longlonglongrightarrow> \<mu> {a<..b}" |
76 |
(is "(\<lambda>T. 1 / (2 * pi) * (CLBINT t=-T..T. ?F t * \<phi> t)) \<longlonglongrightarrow> of_real (\<mu> {a<..b})") |
|
77 |
proof - |
|
78 |
interpret P: pair_sigma_finite lborel M .. |
|
79 |
from bounded_Si obtain B where Bprop: "\<And>T. abs (Si T) \<le> B" by auto |
|
80 |
from Bprop [of 0] have [simp]: "B \<ge> 0" by auto |
|
63589 | 81 |
let ?f = "\<lambda>t x :: real. (iexp (t * (x - a)) - iexp(t * (x - b))) / (\<i> * t)" |
62083 | 82 |
{ fix T :: real |
83 |
assume "T \<ge> 0" |
|
84 |
let ?f' = "\<lambda>(t, x). indicator {-T<..<T} t *\<^sub>R ?f t x" |
|
85 |
{ fix x |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
86 |
have int: "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\<lambda>t. 2 * (sin (t * (x-w)) / t))" for w |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
87 |
using integrable_sinc' interval_lebesgue_integrable_mult_right by blast |
62083 | 88 |
have 1: "complex_interval_lebesgue_integrable lborel u v (\<lambda>t. ?f t x)" for u v :: real |
89 |
using Levy_Inversion_aux2[of "x - b" "x - a"] |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
90 |
apply (simp add: interval_lebesgue_integrable_def set_integrable_def del: times_divide_eq_left) |
62083 | 91 |
apply (intro integrableI_bounded_set_indicator[where B="b - a"] conjI impI) |
92 |
apply (auto intro!: AE_I [of _ _ "{0}"] simp: assms) |
|
93 |
done |
|
94 |
have "(CLBINT t. ?f' (t, x)) = (CLBINT t=-T..T. ?f t x)" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
95 |
using \<open>T \<ge> 0\<close> by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def) |
62083 | 96 |
also have "\<dots> = (CLBINT t=-T..(0 :: real). ?f t x) + (CLBINT t=(0 :: real)..T. ?f t x)" |
97 |
(is "_ = _ + ?t") |
|
98 |
using 1 by (intro interval_integral_sum[symmetric]) (simp add: min_absorb1 max_absorb2 \<open>T \<ge> 0\<close>) |
|
99 |
also have "(CLBINT t=-T..(0 :: real). ?f t x) = (CLBINT t=(0::real)..T. ?f (-t) x)" |
|
100 |
by (subst interval_integral_reflect) auto |
|
101 |
also have "\<dots> + ?t = (CLBINT t=(0::real)..T. ?f (-t) x + ?f t x)" |
|
102 |
using 1 |
|
103 |
by (intro interval_lebesgue_integral_add(2) [symmetric] interval_integrable_mirror[THEN iffD2]) simp_all |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
104 |
also have "\<dots> = (CLBINT t=(0::real)..T. ((iexp(t * (x - a)) - iexp (-(t * (x - a)))) - |
63589 | 105 |
(iexp(t * (x - b)) - iexp (-(t * (x - b))))) / (\<i> * t))" |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
67977
diff
changeset
|
106 |
using \<open>T \<ge> 0\<close> by (intro interval_integral_cong) (auto simp add: field_split_simps) |
62083 | 107 |
also have "\<dots> = (CLBINT t=(0::real)..T. complex_of_real( |
108 |
2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t)))" |
|
63167 | 109 |
using \<open>T \<ge> 0\<close> |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
110 |
by (intro interval_integral_cong) (simp add: divide_simps Im_divide Re_divide Im_exp Re_exp complex_eq_iff) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
111 |
also have "\<dots> = complex_of_real (LBINT t=(0::real)..T. |
62083 | 112 |
2 * (sin (t * (x - a)) / t) - 2 * (sin (t * (x - b)) / t))" |
113 |
by (rule interval_lebesgue_integral_of_real) |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
114 |
also have "\<dots> = complex_of_real ((LBINT t=ereal 0..ereal T. 2 * (sin (t * (x - a)) / t)) - |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
115 |
(LBINT t=ereal 0..ereal T. 2 * (sin (t * (x - b)) / t)))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
116 |
unfolding interval_lebesgue_integral_diff |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
117 |
using int by auto |
62083 | 118 |
also have "\<dots> = complex_of_real (2 * (sgn (x - a) * Si (T * abs (x - a)) - |
119 |
sgn (x - b) * Si (T * abs (x - b))))" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
120 |
unfolding interval_lebesgue_integral_mult_right |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
121 |
by (simp add: zero_ereal_def[symmetric] LBINT_I0c_sin_scale_divide[OF \<open>T \<ge> 0\<close>]) |
62083 | 122 |
finally have "(CLBINT t. ?f' (t, x)) = |
123 |
2 * (sgn (x - a) * Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))" . |
|
124 |
} note main_eq = this |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
125 |
have "(CLBINT t=-T..T. ?F t * \<phi> t) = |
62083 | 126 |
(CLBINT t. (CLINT x | M. ?F t * iexp (t * x) * indicator {-T<..<T} t))" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
127 |
using \<open>T \<ge> 0\<close> unfolding \<phi>_def char_def interval_lebesgue_integral_def set_lebesgue_integral_def |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
128 |
by (auto split: split_indicator intro!: Bochner_Integration.integral_cong) |
62083 | 129 |
also have "\<dots> = (CLBINT t. (CLINT x | M. ?f' (t, x)))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
130 |
by (auto intro!: Bochner_Integration.integral_cong simp: field_simps exp_diff exp_minus split: split_indicator) |
62083 | 131 |
also have "\<dots> = (CLINT x | M. (CLBINT t. ?f' (t, x)))" |
132 |
proof (intro P.Fubini_integral [symmetric] integrableI_bounded_set [where B="b - a"]) |
|
133 |
show "emeasure (lborel \<Otimes>\<^sub>M M) ({- T<..<T} \<times> space M) < \<infinity>" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
134 |
using \<open>T \<ge> 0\<close> |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
135 |
by (subst emeasure_pair_measure_Times) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
136 |
(auto simp: ennreal_mult_less_top not_less top_unique) |
62083 | 137 |
show "AE x\<in>{- T<..<T} \<times> space M in lborel \<Otimes>\<^sub>M M. cmod (case x of (t, x) \<Rightarrow> ?f' (t, x)) \<le> b - a" |
63167 | 138 |
using Levy_Inversion_aux2[of "x - b" "x - a" for x] \<open>a \<le> b\<close> |
62083 | 139 |
by (intro AE_I [of _ _ "{0} \<times> UNIV"]) (force simp: emeasure_pair_measure_Times)+ |
140 |
qed (auto split: split_indicator split_indicator_asm) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
141 |
also have "\<dots> = (CLINT x | M. (complex_of_real (2 * (sgn (x - a) * |
62083 | 142 |
Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
143 |
using main_eq by (intro Bochner_Integration.integral_cong, auto) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
144 |
also have "\<dots> = complex_of_real (LINT x | M. (2 * (sgn (x - a) * |
62083 | 145 |
Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))" |
146 |
by (rule integral_complex_of_real) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
147 |
finally have "(CLBINT t=-T..T. ?F t * \<phi> t) = |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
148 |
complex_of_real (LINT x | M. (2 * (sgn (x - a) * |
62083 | 149 |
Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b)))))" . |
150 |
} note main_eq2 = this |
|
151 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
152 |
have "(\<lambda>T :: nat. LINT x | M. (2 * (sgn (x - a) * |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
153 |
Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow> |
62083 | 154 |
(LINT x | M. 2 * pi * indicator {a<..b} x)" |
155 |
proof (rule integral_dominated_convergence [where w="\<lambda>x. 4 * B"]) |
|
156 |
show "integrable M (\<lambda>x. 4 * B)" |
|
157 |
by (rule integrable_const_bound [of _ "4 * B"]) auto |
|
158 |
next |
|
159 |
let ?S = "\<lambda>n::nat. \<lambda>x. sgn (x - a) * Si (n * \<bar>x - a\<bar>) - sgn (x - b) * Si (n * \<bar>x - b\<bar>)" |
|
160 |
{ fix n x |
|
161 |
have "norm (?S n x) \<le> norm (sgn (x - a) * Si (n * \<bar>x - a\<bar>)) + norm (sgn (x - b) * Si (n * \<bar>x - b\<bar>))" |
|
162 |
by (rule norm_triangle_ineq4) |
|
163 |
also have "\<dots> \<le> B + B" |
|
164 |
using Bprop by (intro add_mono) (auto simp: abs_mult abs_sgn_eq) |
|
165 |
finally have "norm (2 * ?S n x) \<le> 4 * B" |
|
166 |
by simp } |
|
167 |
then show "\<And>n. AE x in M. norm (2 * ?S n x) \<le> 4 * B" |
|
168 |
by auto |
|
169 |
have "AE x in M. x \<noteq> a" "AE x in M. x \<noteq> b" |
|
63167 | 170 |
using prob_eq_0[of "{a}"] prob_eq_0[of "{b}"] \<open>\<mu> {a} = 0\<close> \<open>\<mu> {b} = 0\<close> by (auto simp: \<mu>_def) |
62083 | 171 |
then show "AE x in M. (\<lambda>n. 2 * ?S n x) \<longlonglongrightarrow> 2 * pi * indicator {a<..b} x" |
172 |
proof eventually_elim |
|
173 |
fix x assume x: "x \<noteq> a" "x \<noteq> b" |
|
174 |
then have "(\<lambda>n. 2 * (sgn (x - a) * Si (\<bar>x - a\<bar> * n) - sgn (x - b) * Si (\<bar>x - b\<bar> * n))) |
|
175 |
\<longlonglongrightarrow> 2 * (sgn (x - a) * (pi / 2) - sgn (x - b) * (pi / 2))" |
|
176 |
by (intro tendsto_intros filterlim_compose[OF Si_at_top] |
|
177 |
filterlim_tendsto_pos_mult_at_top[OF tendsto_const] filterlim_real_sequentially) |
|
178 |
auto |
|
179 |
also have "(\<lambda>n. 2 * (sgn (x - a) * Si (\<bar>x - a\<bar> * n) - sgn (x - b) * Si (\<bar>x - b\<bar> * n))) = (\<lambda>n. 2 * ?S n x)" |
|
180 |
by (auto simp: ac_simps) |
|
181 |
also have "2 * (sgn (x - a) * (pi / 2) - sgn (x - b) * (pi / 2)) = 2 * pi * indicator {a<..b} x" |
|
63167 | 182 |
using x \<open>a \<le> b\<close> by (auto split: split_indicator) |
62083 | 183 |
finally show "(\<lambda>n. 2 * ?S n x) \<longlonglongrightarrow> 2 * pi * indicator {a<..b} x" . |
184 |
qed |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
185 |
qed simp_all |
62083 | 186 |
also have "(LINT x | M. 2 * pi * indicator {a<..b} x) = 2 * pi * \<mu> {a<..b}" |
187 |
by (simp add: \<mu>_def) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
188 |
finally have "(\<lambda>T. LINT x | M. (2 * (sgn (x - a) * |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
189 |
Si (T * abs (x - a)) - sgn (x - b) * Si (T * abs (x - b))))) \<longlonglongrightarrow> |
62083 | 190 |
2 * pi * \<mu> {a<..b}" . |
191 |
with main_eq2 show ?thesis |
|
192 |
by (auto intro!: tendsto_eq_intros) |
|
193 |
qed |
|
194 |
||
195 |
theorem Levy_uniqueness: |
|
196 |
fixes M1 M2 :: "real measure" |
|
197 |
assumes "real_distribution M1" "real_distribution M2" and |
|
198 |
"char M1 = char M2" |
|
199 |
shows "M1 = M2" |
|
200 |
proof - |
|
201 |
interpret M1: real_distribution M1 by (rule assms) |
|
202 |
interpret M2: real_distribution M2 by (rule assms) |
|
203 |
have "countable ({x. measure M1 {x} \<noteq> 0} \<union> {x. measure M2 {x} \<noteq> 0})" |
|
204 |
by (intro countable_Un M2.countable_support M1.countable_support) |
|
205 |
then have count: "countable {x. measure M1 {x} \<noteq> 0 \<or> measure M2 {x} \<noteq> 0}" |
|
206 |
by (simp add: Un_def) |
|
207 |
||
208 |
have "cdf M1 = cdf M2" |
|
209 |
proof (rule ext) |
|
210 |
fix x |
|
63393 | 211 |
let ?D = "\<lambda>x. \<bar>cdf M1 x - cdf M2 x\<bar>" |
62083 | 212 |
|
213 |
{ fix \<epsilon> :: real |
|
214 |
assume "\<epsilon> > 0" |
|
63393 | 215 |
have "(?D \<longlongrightarrow> 0) at_bot" |
67682
00c436488398
tuned proofs -- prefer explicit names for facts from 'interpret';
wenzelm
parents:
66447
diff
changeset
|
216 |
using M1.cdf_lim_at_bot M2.cdf_lim_at_bot by (intro tendsto_eq_intros) auto |
63393 | 217 |
then have "eventually (\<lambda>y. ?D y < \<epsilon> / 2 \<and> y \<le> x) at_bot" |
218 |
using \<open>\<epsilon> > 0\<close> by (simp only: tendsto_iff dist_real_def diff_0_right eventually_conj eventually_le_at_bot abs_idempotent) |
|
219 |
then obtain M where "\<And>y. y \<le> M \<Longrightarrow> ?D y < \<epsilon> / 2" "M \<le> x" |
|
62083 | 220 |
unfolding eventually_at_bot_linorder by auto |
221 |
with open_minus_countable[OF count, of "{..< M}"] obtain a where |
|
63393 | 222 |
"measure M1 {a} = 0" "measure M2 {a} = 0" "a < M" "a \<le> x" and 1: "?D a < \<epsilon> / 2" |
62083 | 223 |
by auto |
224 |
||
63393 | 225 |
have "(?D \<longlongrightarrow> ?D x) (at_right x)" |
226 |
using M1.cdf_is_right_cont [of x] M2.cdf_is_right_cont [of x] |
|
227 |
by (intro tendsto_intros) (auto simp add: continuous_within) |
|
228 |
then have "eventually (\<lambda>y. \<bar>?D y - ?D x\<bar> < \<epsilon> / 2) (at_right x)" |
|
229 |
using \<open>\<epsilon> > 0\<close> by (simp only: tendsto_iff dist_real_def eventually_conj eventually_at_right_less) |
|
230 |
then obtain N where "N > x" "\<And>y. x < y \<Longrightarrow> y < N \<Longrightarrow> \<bar>?D y - ?D x\<bar> < \<epsilon> / 2" |
|
62083 | 231 |
by (auto simp add: eventually_at_right[OF less_add_one]) |
232 |
with open_minus_countable[OF count, of "{x <..< N}"] obtain b where "x < b" "b < N" |
|
63393 | 233 |
"measure M1 {b} = 0" "measure M2 {b} = 0" and 2: "\<bar>?D b - ?D x\<bar> < \<epsilon> / 2" |
62083 | 234 |
by (auto simp: abs_minus_commute) |
63167 | 235 |
from \<open>a \<le> x\<close> \<open>x < b\<close> have "a < b" "a \<le> b" by auto |
62083 | 236 |
|
63167 | 237 |
from \<open>char M1 = char M2\<close> |
63393 | 238 |
M1.Levy_Inversion [OF \<open>a \<le> b\<close> \<open>measure M1 {a} = 0\<close> \<open>measure M1 {b} = 0\<close>] |
63167 | 239 |
M2.Levy_Inversion [OF \<open>a \<le> b\<close> \<open>measure M2 {a} = 0\<close> \<open>measure M2 {b} = 0\<close>] |
62083 | 240 |
have "complex_of_real (measure M1 {a<..b}) = complex_of_real (measure M2 {a<..b})" |
241 |
by (intro LIMSEQ_unique) auto |
|
63393 | 242 |
then have "?D a = ?D b" |
243 |
unfolding of_real_eq_iff M1.cdf_diff_eq [OF \<open>a < b\<close>, symmetric] M2.cdf_diff_eq [OF \<open>a < b\<close>, symmetric] by simp |
|
244 |
then have "?D x = \<bar>(?D b - ?D x) - ?D a\<bar>" |
|
62083 | 245 |
by simp |
63393 | 246 |
also have "\<dots> \<le> \<bar>?D b - ?D x\<bar> + \<bar>?D a\<bar>" |
247 |
by (rule abs_triangle_ineq4) |
|
248 |
also have "\<dots> \<le> \<epsilon> / 2 + \<epsilon> / 2" |
|
249 |
using 1 2 by (intro add_mono) auto |
|
250 |
finally have "?D x \<le> \<epsilon>" by simp } |
|
62083 | 251 |
then show "cdf M1 x = cdf M2 x" |
252 |
by (metis abs_le_zero_iff dense_ge eq_iff_diff_eq_0) |
|
253 |
qed |
|
254 |
thus ?thesis |
|
63167 | 255 |
by (rule cdf_unique [OF \<open>real_distribution M1\<close> \<open>real_distribution M2\<close>]) |
62083 | 256 |
qed |
257 |
||
258 |
||
259 |
subsection \<open>The Levy continuity theorem\<close> |
|
260 |
||
261 |
theorem levy_continuity1: |
|
262 |
fixes M :: "nat \<Rightarrow> real measure" and M' :: "real measure" |
|
263 |
assumes "\<And>n. real_distribution (M n)" "real_distribution M'" "weak_conv_m M M'" |
|
264 |
shows "(\<lambda>n. char (M n) t) \<longlonglongrightarrow> char M' t" |
|
265 |
unfolding char_def using assms by (rule weak_conv_imp_integral_bdd_continuous_conv) auto |
|
266 |
||
267 |
theorem levy_continuity: |
|
268 |
fixes M :: "nat \<Rightarrow> real measure" and M' :: "real measure" |
|
269 |
assumes real_distr_M : "\<And>n. real_distribution (M n)" |
|
270 |
and real_distr_M': "real_distribution M'" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
271 |
and char_conv: "\<And>t. (\<lambda>n. char (M n) t) \<longlonglongrightarrow> char M' t" |
62083 | 272 |
shows "weak_conv_m M M'" |
273 |
proof - |
|
274 |
interpret Mn: real_distribution "M n" for n by fact |
|
275 |
interpret M': real_distribution M' by fact |
|
276 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
277 |
have *: "\<And>u x. u > 0 \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> (CLBINT t:{-u..u}. 1 - iexp (t * x)) = |
62083 | 278 |
2 * (u - sin (u * x) / x)" |
279 |
proof - |
|
280 |
fix u :: real and x :: real |
|
281 |
assume "u > 0" and "x \<noteq> 0" |
|
282 |
hence "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = (CLBINT t=-u..u. 1 - iexp (t * x))" |
|
283 |
by (subst interval_integral_Icc, auto) |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
284 |
also have "\<dots> = (CLBINT t=-u..ereal 0. 1 - iexp (t * x)) + (CLBINT t= ereal 0..u. 1 - iexp (t * x))" |
63167 | 285 |
using \<open>u > 0\<close> |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
286 |
by (subst interval_integral_sum; force simp add: interval_integrable_isCont) |
62083 | 287 |
also have "\<dots> = (CLBINT t=ereal 0..u. 1 - iexp (t * -x)) + (CLBINT t=ereal 0..u. 1 - iexp (t * x))" |
288 |
by (subst interval_integral_reflect, auto) |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
289 |
also have "... = CLBINT xa=ereal 0..ereal u. 1 - iexp (xa * - x) + (1 - iexp (xa * x))" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
290 |
by (subst interval_lebesgue_integral_add (2) [symmetric]) (auto simp: interval_integrable_isCont) |
62083 | 291 |
also have "\<dots> = (LBINT t=ereal 0..u. 2 - 2 * cos (t * x))" |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
292 |
unfolding exp_Euler cos_of_real by (simp flip: interval_lebesgue_integral_of_real) |
62083 | 293 |
also have "\<dots> = 2 * u - 2 * sin (u * x) / x" |
294 |
by (subst interval_lebesgue_integral_diff) |
|
295 |
(auto intro!: interval_integrable_isCont |
|
63167 | 296 |
simp: interval_lebesgue_integral_of_real integral_cos [OF \<open>x \<noteq> 0\<close>] mult.commute[of _ x]) |
62083 | 297 |
finally show "(CLBINT t:{-u..u}. 1 - iexp (t * x)) = 2 * (u - sin (u * x) / x)" |
298 |
by (simp add: field_simps) |
|
299 |
qed |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
300 |
have main_bound: "\<And>u n. u > 0 \<Longrightarrow> Re (CLBINT t:{-u..u}. 1 - char (M n) t) \<ge> |
62083 | 301 |
u * measure (M n) {x. abs x \<ge> 2 / u}" |
302 |
proof - |
|
303 |
fix u :: real and n |
|
304 |
assume "u > 0" |
|
305 |
interpret P: pair_sigma_finite "M n" lborel .. |
|
306 |
(* TODO: put this in the real_distribution locale as a simp rule? *) |
|
307 |
have Mn1 [simp]: "measure (M n) UNIV = 1" by (metis Mn.prob_space Mn.space_eq_univ) |
|
308 |
(* TODO: make this automatic somehow? *) |
|
309 |
have Mn2 [simp]: "\<And>x. complex_integrable (M n) (\<lambda>t. exp (\<i> * complex_of_real (x * t)))" |
|
310 |
by (rule Mn.integrable_const_bound [where B = 1], auto) |
|
311 |
have Mn3: "set_integrable (M n \<Otimes>\<^sub>M lborel) (UNIV \<times> {- u..u}) (\<lambda>a. 1 - exp (\<i> * complex_of_real (snd a * fst a)))" |
|
63167 | 312 |
using \<open>0 < u\<close> |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
313 |
unfolding set_integrable_def |
62083 | 314 |
by (intro integrableI_bounded_set_indicator [where B="2"]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
315 |
(auto simp: lborel.emeasure_pair_measure_Times ennreal_mult_less_top not_less top_unique |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
316 |
split: split_indicator |
62083 | 317 |
intro!: order_trans [OF norm_triangle_ineq4]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
318 |
have "(CLBINT t:{-u..u}. 1 - char (M n) t) = |
62083 | 319 |
(CLBINT t:{-u..u}. (CLINT x | M n. 1 - iexp (t * x)))" |
320 |
unfolding char_def by (rule set_lebesgue_integral_cong, auto simp del: of_real_mult) |
|
321 |
also have "\<dots> = (CLBINT t. (CLINT x | M n. indicator {-u..u} t *\<^sub>R (1 - iexp (t * x))))" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
322 |
unfolding set_lebesgue_integral_def |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
323 |
by (rule Bochner_Integration.integral_cong) (auto split: split_indicator) |
62083 | 324 |
also have "\<dots> = (CLINT x | M n. (CLBINT t:{-u..u}. 1 - iexp (t * x)))" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
325 |
using Mn3 by (subst P.Fubini_integral) (auto simp: indicator_times split_beta' set_integrable_def set_lebesgue_integral_def) |
62083 | 326 |
also have "\<dots> = (CLINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" |
63886
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents:
63589
diff
changeset
|
327 |
using \<open>u > 0\<close> by (intro Bochner_Integration.integral_cong, auto simp add: * simp del: of_real_mult) |
62083 | 328 |
also have "\<dots> = (LINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" |
329 |
by (rule integral_complex_of_real) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
330 |
finally have "Re (CLBINT t:{-u..u}. 1 - char (M n) t) = |
62083 | 331 |
(LINT x | M n. (if x = 0 then 0 else 2 * (u - sin (u * x) / x)))" by simp |
332 |
also have "\<dots> \<ge> (LINT x : {x. abs x \<ge> 2 / u} | M n. u)" |
|
333 |
proof - |
|
334 |
have "complex_integrable (M n) (\<lambda>x. CLBINT t:{-u..u}. 1 - iexp (snd (x, t) * fst (x, t)))" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
335 |
using Mn3 unfolding set_integrable_def set_lebesgue_integral_def |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
336 |
by (intro P.integrable_fst) (simp add: indicator_times split_beta') |
62083 | 337 |
hence "complex_integrable (M n) (\<lambda>x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
338 |
using \<open>u > 0\<close> |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
339 |
unfolding set_integrable_def |
75463
8e2285baadba
qualified name to fix integrable_cong ambiguity
paulson <lp15@cam.ac.uk>
parents:
74362
diff
changeset
|
340 |
by (subst Bochner_Integration.integrable_cong) (auto simp add: * simp del: of_real_mult) |
62083 | 341 |
hence **: "integrable (M n) (\<lambda>x. if x = 0 then 0 else 2 * (u - sin (u * x) / x))" |
342 |
unfolding complex_of_real_integrable_eq . |
|
343 |
have "2 * sin x \<le> x" if "2 \<le> x" for x :: real |
|
344 |
by (rule order_trans[OF _ \<open>2 \<le> x\<close>]) auto |
|
345 |
moreover have "x \<le> 2 * sin x" if "x \<le> - 2" for x :: real |
|
346 |
by (rule order_trans[OF \<open>x \<le> - 2\<close>]) auto |
|
347 |
moreover have "x < 0 \<Longrightarrow> x \<le> sin x" for x :: real |
|
348 |
using sin_x_le_x[of "-x"] by simp |
|
349 |
ultimately show ?thesis |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
350 |
using \<open>u > 0\<close> unfolding set_lebesgue_integral_def |
62083 | 351 |
by (intro integral_mono [OF _ **]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
352 |
(auto simp: divide_simps sin_x_le_x mult.commute[of u] mult_neg_pos top_unique less_top[symmetric] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
353 |
split: split_indicator) |
62083 | 354 |
qed |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
355 |
also (xtrans) have "(LINT x : {x. abs x \<ge> 2 / u} | M n. u) = u * measure (M n) {x. abs x \<ge> 2 / u}" |
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
356 |
unfolding set_lebesgue_integral_def |
62083 | 357 |
by (simp add: Mn.emeasure_eq_measure) |
358 |
finally show "Re (CLBINT t:{-u..u}. 1 - char (M n) t) \<ge> u * measure (M n) {x. abs x \<ge> 2 / u}" . |
|
359 |
qed |
|
360 |
||
361 |
have tight_aux: "\<And>\<epsilon>. \<epsilon> > 0 \<Longrightarrow> \<exists>a b. a < b \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {a<..b})" |
|
362 |
proof - |
|
363 |
fix \<epsilon> :: real |
|
364 |
assume "\<epsilon> > 0" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
365 |
with M'.isCont_char [of 0] |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
366 |
obtain d where d0: "d>0" and "\<forall>x'. dist x' 0 < d \<longrightarrow> dist (char M' x') (char M' 0) < \<epsilon>/4" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
367 |
unfolding continuous_at_eps_delta by (metis \<open>0 < \<epsilon>\<close> divide_pos_pos zero_less_numeral) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
368 |
then have d1: "\<And>t. abs t < d \<Longrightarrow> cmod (char M' t - 1) < \<epsilon> / 4" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
369 |
by (simp add: M'.char_zero dist_norm) |
62083 | 370 |
have 1: "\<And>x. cmod (1 - char M' x) \<le> 2" |
371 |
by (rule order_trans [OF norm_triangle_ineq4], auto simp add: M'.cmod_char_le_1) |
|
372 |
then have 2: "\<And>u v. complex_set_integrable lborel {u..v} (\<lambda>x. 1 - char M' x)" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
373 |
unfolding set_integrable_def |
62083 | 374 |
by (intro integrableI_bounded_set_indicator[where B=2]) (auto simp: emeasure_lborel_Icc_eq) |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
375 |
have 3: "\<And>u v. integrable lborel (\<lambda>x. indicat_real {u..v} x *\<^sub>R cmod (1 - char M' x))" |
62083 | 376 |
by (intro borel_integrable_compact[OF compact_Icc] continuous_at_imp_continuous_on |
377 |
continuous_intros ballI M'.isCont_char continuous_intros) |
|
378 |
have "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \<le> LBINT t:{-d/2..d/2}. cmod (1 - char M' t)" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
379 |
unfolding set_lebesgue_integral_def |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63886
diff
changeset
|
380 |
using integral_norm_bound[of _ "\<lambda>x. indicator {u..v} x *\<^sub>R (1 - char M' x)" for u v] by simp |
63540 | 381 |
also have 4: "\<dots> \<le> LBINT t:{-d/2..d/2}. \<epsilon> / 4" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
382 |
unfolding set_lebesgue_integral_def |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
383 |
proof (rule integral_mono [OF 3]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
384 |
|
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
385 |
show "indicat_real {- d / 2..d / 2} x *\<^sub>R cmod (1 - char M' x) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
386 |
\<le> indicat_real {- d / 2..d / 2} x *\<^sub>R (\<epsilon> / 4)" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
387 |
if "x \<in> space lborel" for x |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
388 |
proof (cases "x \<in> {-d/2..d/2}") |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
389 |
case True |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
390 |
show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
391 |
using d0 d1 that True [simplified] |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
392 |
by (smt (verit, best) field_sum_of_halves minus_diff_eq norm_minus_cancel indicator_pos_le scaleR_left_mono) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
393 |
qed auto |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
394 |
qed (simp add: emeasure_lborel_Icc_eq) |
63540 | 395 |
also from d0 4 have "\<dots> = d * \<epsilon> / 4" |
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
396 |
unfolding set_lebesgue_integral_def by simp |
62083 | 397 |
finally have bound: "cmod (CLBINT t:{-d/2..d/2}. 1 - char M' t) \<le> d * \<epsilon> / 4" . |
63540 | 398 |
have "cmod (1 - char (M n) x) \<le> 2" for n x |
399 |
by (rule order_trans [OF norm_triangle_ineq4], auto simp add: Mn.cmod_char_le_1) |
|
400 |
then have "(\<lambda>n. CLBINT t:{-d/2..d/2}. 1 - char (M n) t) \<longlonglongrightarrow> (CLBINT t:{-d/2..d/2}. 1 - char M' t)" |
|
67977
557ea2740125
Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents:
67682
diff
changeset
|
401 |
unfolding set_lebesgue_integral_def |
62083 | 402 |
apply (intro integral_dominated_convergence[where w="\<lambda>x. indicator {-d/2..d/2} x *\<^sub>R 2"]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
403 |
apply (auto intro!: char_conv tendsto_intros |
62083 | 404 |
simp: emeasure_lborel_Icc_eq |
405 |
split: split_indicator) |
|
406 |
done |
|
407 |
hence "eventually (\<lambda>n. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - |
|
408 |
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4) sequentially" |
|
63167 | 409 |
using d0 \<open>\<epsilon> > 0\<close> apply (subst (asm) tendsto_iff) |
62083 | 410 |
by (subst (asm) dist_complex_def, drule spec, erule mp, auto) |
411 |
hence "\<exists>N. \<forall>n \<ge> N. cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - |
|
412 |
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4" by (simp add: eventually_sequentially) |
|
74362 | 413 |
then obtain N |
414 |
where "\<forall>n\<ge>N. cmod ((CLBINT t:{- d / 2..d / 2}. 1 - char (M n) t) - |
|
415 |
(CLBINT t:{- d / 2..d / 2}. 1 - char M' t)) < d * \<epsilon> / 4" .. |
|
62083 | 416 |
hence N: "\<And>n. n \<ge> N \<Longrightarrow> cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - |
417 |
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) < d * \<epsilon> / 4" by auto |
|
418 |
{ fix n |
|
419 |
assume "n \<ge> N" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
420 |
have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) = |
62083 | 421 |
cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - (CLBINT t:{-d/2..d/2}. 1 - char M' t) |
422 |
+ (CLBINT t:{-d/2..d/2}. 1 - char M' t))" by simp |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
423 |
also have "\<dots> \<le> cmod ((CLBINT t:{-d/2..d/2}. 1 - char (M n) t) - |
62083 | 424 |
(CLBINT t:{-d/2..d/2}. 1 - char M' t)) + cmod(CLBINT t:{-d/2..d/2}. 1 - char M' t)" |
425 |
by (rule norm_triangle_ineq) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
426 |
also have "\<dots> < d * \<epsilon> / 4 + d * \<epsilon> / 4" |
63167 | 427 |
by (rule add_less_le_mono [OF N [OF \<open>n \<ge> N\<close>] bound]) |
62083 | 428 |
also have "\<dots> = d * \<epsilon> / 2" by auto |
429 |
finally have "cmod (CLBINT t:{-d/2..d/2}. 1 - char (M n) t) < d * \<epsilon> / 2" . |
|
430 |
hence "d * \<epsilon> / 2 > Re (CLBINT t:{-d/2..d/2}. 1 - char (M n) t)" |
|
431 |
by (rule order_le_less_trans [OF complex_Re_le_cmod]) |
|
432 |
hence "d * \<epsilon> / 2 > Re (CLBINT t:{-(d/2)..d/2}. 1 - char (M n) t)" (is "_ > ?lhs") by simp |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
433 |
also have "?lhs \<ge> (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}" |
62083 | 434 |
using d0 by (intro main_bound, simp) |
435 |
finally (xtrans) have "d * \<epsilon> / 2 > (d / 2) * measure (M n) {x. abs x \<ge> 2 / (d / 2)}" . |
|
63167 | 436 |
with d0 \<open>\<epsilon> > 0\<close> have "\<epsilon> > measure (M n) {x. abs x \<ge> 2 / (d / 2)}" by (simp add: field_simps) |
62083 | 437 |
hence "\<epsilon> > 1 - measure (M n) (UNIV - {x. abs x \<ge> 2 / (d / 2)})" |
438 |
apply (subst Mn.borel_UNIV [symmetric]) |
|
439 |
by (subst Mn.prob_compl, auto) |
|
440 |
also have "UNIV - {x. abs x \<ge> 2 / (d / 2)} = {x. -(4 / d) < x \<and> x < (4 / d)}" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
441 |
using d0 by (simp add: set_eq_iff divide_simps abs_if) (smt (verit, best) mult_less_0_iff) |
62083 | 442 |
finally have "measure (M n) {x. -(4 / d) < x \<and> x < (4 / d)} > 1 - \<epsilon>" |
443 |
by auto |
|
444 |
} note 6 = this |
|
445 |
{ fix n :: nat |
|
446 |
have *: "(UN (k :: nat). {- real k<..real k}) = UNIV" |
|
447 |
by (auto, metis leI le_less_trans less_imp_le minus_less_iff reals_Archimedean2) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
448 |
have "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow> |
62083 | 449 |
measure (M n) (UN (k :: nat). {- real k<..real k})" |
450 |
by (rule Mn.finite_Lim_measure_incseq, auto simp add: incseq_def) |
|
451 |
hence "(\<lambda>k. measure (M n) {- real k<..real k}) \<longlonglongrightarrow> 1" |
|
452 |
using Mn.prob_space unfolding * Mn.borel_UNIV by simp |
|
453 |
hence "eventually (\<lambda>k. measure (M n) {- real k<..real k} > 1 - \<epsilon>) sequentially" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
454 |
using \<open>\<epsilon> > 0\<close> order_tendstoD by fastforce |
62083 | 455 |
} note 7 = this |
456 |
{ fix n :: nat |
|
457 |
have "eventually (\<lambda>k. \<forall>m < n. measure (M m) {- real k<..real k} > 1 - \<epsilon>) sequentially" |
|
458 |
(is "?P n") |
|
459 |
proof (induct n) |
|
460 |
case (Suc n) with 7[of n] show ?case |
|
461 |
by eventually_elim (auto simp add: less_Suc_eq) |
|
462 |
qed simp |
|
463 |
} note 8 = this |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
464 |
from 8 [of N] have "\<exists>K :: nat. \<forall>k \<ge> K. \<forall>m<N. 1 - \<epsilon> < |
62083 | 465 |
Sigma_Algebra.measure (M m) {- real k<..real k}" |
466 |
by (auto simp add: eventually_sequentially) |
|
467 |
hence "\<exists>K :: nat. \<forall>m<N. 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}" by auto |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
468 |
then obtain K :: nat where |
62083 | 469 |
"\<forall>m<N. 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}" .. |
470 |
hence K: "\<And>m. m < N \<Longrightarrow> 1 - \<epsilon> < Sigma_Algebra.measure (M m) {- real K<..real K}" |
|
471 |
by auto |
|
472 |
let ?K' = "max K (4 / d)" |
|
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
473 |
have "1 - \<epsilon> < measure (M n) {- max (real K) (4 / d)<..max (real K) (4 / d)}" for n |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
474 |
proof (cases "n < N") |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
475 |
case True |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
476 |
then show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
477 |
by (force intro: order_less_le_trans [OF K Mn.finite_measure_mono]) |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
478 |
next |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
479 |
case False |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
480 |
then show ?thesis |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
481 |
by (force intro: order_less_le_trans [OF 6 Mn.finite_measure_mono]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
482 |
qed |
78517
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
483 |
then have "-?K' < ?K' \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {-?K'<..?K'})" |
28c1f4f5335f
Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents:
75463
diff
changeset
|
484 |
using d0 by (simp add: less_max_iff_disj minus_less_iff) |
62083 | 485 |
thus "\<exists>a b. a < b \<and> (\<forall>n. 1 - \<epsilon> < measure (M n) {a<..b})" by (intro exI) |
486 |
qed |
|
487 |
have tight: "tight M" |
|
488 |
by (auto simp: tight_def intro: assms tight_aux) |
|
489 |
show ?thesis |
|
490 |
proof (rule tight_subseq_weak_converge [OF real_distr_M real_distr_M' tight]) |
|
491 |
fix s \<nu> |
|
66447
a1f5c5c26fa6
Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents:
65064
diff
changeset
|
492 |
assume s: "strict_mono (s :: nat \<Rightarrow> nat)" |
62083 | 493 |
assume nu: "weak_conv_m (M \<circ> s) \<nu>" |
494 |
assume *: "real_distribution \<nu>" |
|
495 |
have 2: "\<And>n. real_distribution ((M \<circ> s) n)" unfolding comp_def by (rule assms) |
|
496 |
have 3: "\<And>t. (\<lambda>n. char ((M \<circ> s) n) t) \<longlonglongrightarrow> char \<nu> t" by (intro levy_continuity1 [OF 2 * nu]) |
|
497 |
have 4: "\<And>t. (\<lambda>n. char ((M \<circ> s) n) t) = ((\<lambda>n. char (M n) t) \<circ> s)" by (rule ext, simp) |
|
498 |
have 5: "\<And>t. (\<lambda>n. char ((M \<circ> s) n) t) \<longlonglongrightarrow> char M' t" |
|
62397
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents:
62083
diff
changeset
|
499 |
by (subst 4, rule LIMSEQ_subseq_LIMSEQ [OF _ s], rule assms) |
62083 | 500 |
hence "char \<nu> = char M'" by (intro ext, intro LIMSEQ_unique [OF 3 5]) |
63167 | 501 |
hence "\<nu> = M'" by (rule Levy_uniqueness [OF * \<open>real_distribution M'\<close>]) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
502 |
thus "weak_conv_m (M \<circ> s) M'" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62397
diff
changeset
|
503 |
by (elim subst) (rule nu) |
62083 | 504 |
qed |
505 |
qed |
|
506 |
||
507 |
end |