| author | wenzelm | 
| Mon, 06 Feb 2023 15:04:21 +0100 | |
| changeset 77208 | a3f67a4459e1 | 
| parent 73253 | f6bb31879698 | 
| permissions | -rw-r--r-- | 
| 63329 | 1 | (* Title: HOL/Probability/Central_Limit_Theorem.thy | 
| 2 | Authors: Jeremy Avigad (CMU), Luke Serafin (CMU) | |
| 62083 | 3 | *) | 
| 4 | ||
| 5 | section \<open>The Central Limit Theorem\<close> | |
| 6 | ||
| 7 | theory Central_Limit_Theorem | |
| 8 | imports Levy | |
| 9 | begin | |
| 10 | ||
| 64009 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 11 | theorem (in prob_space) central_limit_theorem_zero_mean: | 
| 62083 | 12 | fixes X :: "nat \<Rightarrow> 'a \<Rightarrow> real" | 
| 13 | and \<mu> :: "real measure" | |
| 14 | and \<sigma> :: real | |
| 15 | and S :: "nat \<Rightarrow> 'a \<Rightarrow> real" | |
| 16 | assumes X_indep: "indep_vars (\<lambda>i. borel) X UNIV" | |
| 17 | and X_mean_0: "\<And>n. expectation (X n) = 0" | |
| 18 | and \<sigma>_pos: "\<sigma> > 0" | |
| 19 | and X_square_integrable: "\<And>n. integrable M (\<lambda>x. (X n x)\<^sup>2)" | |
| 20 | and X_variance: "\<And>n. variance (X n) = \<sigma>\<^sup>2" | |
| 21 | and X_distrib: "\<And>n. distr M borel (X n) = \<mu>" | |
| 22 | defines "S n \<equiv> \<lambda>x. \<Sum>i<n. X i x" | |
| 23 | shows "weak_conv_m (\<lambda>n. distr M borel (\<lambda>x. S n x / sqrt (n * \<sigma>\<^sup>2))) std_normal_distribution" | |
| 24 | proof - | |
| 25 | let ?S' = "\<lambda>n x. S n x / sqrt (real n * \<sigma>\<^sup>2)" | |
| 63040 | 26 | define \<phi> where "\<phi> n = char (distr M borel (?S' n))" for n | 
| 27 | define \<psi> where "\<psi> n t = char \<mu> (t / sqrt (\<sigma>\<^sup>2 * n))" for n t | |
| 62083 | 28 | |
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 29 | have X_rv [simp, measurable]: "random_variable borel (X n)" for n | 
| 62083 | 30 | using X_indep unfolding indep_vars_def2 by simp | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 31 | have X_integrable [simp, intro]: "integrable M (X n)" for n | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 32 | by (rule square_integrable_imp_integrable[OF _ X_square_integrable]) simp_all | 
| 62083 | 33 | interpret \<mu>: real_distribution \<mu> | 
| 34 | by (subst X_distrib [symmetric, of 0], rule real_distribution_distr, simp) | |
| 35 | ||
| 36 | (* these are equivalent to the hypotheses on X, given X_distr *) | |
| 37 | have \<mu>_integrable [simp]: "integrable \<mu> (\<lambda>x. x)" | |
| 38 | and \<mu>_mean_integrable [simp]: "\<mu>.expectation (\<lambda>x. x) = 0" | |
| 39 | and \<mu>_square_integrable [simp]: "integrable \<mu> (\<lambda>x. x^2)" | |
| 40 | and \<mu>_variance [simp]: "\<mu>.expectation (\<lambda>x. x^2) = \<sigma>\<^sup>2" | |
| 41 | using assms by (simp_all add: X_distrib [symmetric, of 0] integrable_distr_eq integral_distr) | |
| 42 | ||
| 43 | have main: "\<forall>\<^sub>F n in sequentially. | |
| 44 | cmod (\<phi> n t - (1 + (-(t^2) / 2) / n)^n) \<le> | |
| 45 | t\<^sup>2 / (6 * \<sigma>\<^sup>2) * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>t / sqrt (\<sigma>\<^sup>2 * n)\<bar> * \<bar>x\<bar> ^ 3))" for t | |
| 46 | proof (rule eventually_sequentiallyI) | |
| 47 | fix n :: nat | |
| 48 | assume "n \<ge> nat (ceiling (t^2 / 4))" | |
| 49 | hence n: "n \<ge> t^2 / 4" by (subst nat_ceiling_le_eq [symmetric]) | |
| 50 | let ?t = "t / sqrt (\<sigma>\<^sup>2 * n)" | |
| 51 | ||
| 63040 | 52 | define \<psi>' where "\<psi>' n i = char (distr M borel (\<lambda>x. X i x / sqrt (\<sigma>\<^sup>2 * n)))" for n i | 
| 62083 | 53 | have *: "\<And>n i t. \<psi>' n i t = \<psi> n t" | 
| 54 | unfolding \<psi>_def \<psi>'_def char_def | |
| 55 | by (subst X_distrib [symmetric]) (auto simp: integral_distr) | |
| 56 | ||
| 57 | have "\<phi> n t = char (distr M borel (\<lambda>x. \<Sum>i<n. X i x / sqrt (\<sigma>\<^sup>2 * real n))) t" | |
| 64267 | 58 | by (auto simp: \<phi>_def S_def sum_divide_distrib ac_simps) | 
| 62083 | 59 | also have "\<dots> = (\<Prod> i < n. \<psi>' n i t)" | 
| 60 | unfolding \<psi>'_def | |
| 64267 | 61 | apply (rule char_distr_sum) | 
| 62083 | 62 | apply (rule indep_vars_compose2[where X=X]) | 
| 63 | apply (rule indep_vars_subset) | |
| 64 | apply (rule X_indep) | |
| 65 | apply auto | |
| 66 | done | |
| 67 | also have "\<dots> = (\<psi> n t)^n" | |
| 64272 | 68 | by (auto simp add: * prod_constant) | 
| 62083 | 69 | finally have \<phi>_eq: "\<phi> n t =(\<psi> n t)^n" . | 
| 70 | ||
| 71 | have "norm (\<psi> n t - (1 - ?t^2 * \<sigma>\<^sup>2 / 2)) \<le> ?t\<^sup>2 / 6 * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t\<bar> * \<bar>x\<bar> ^ 3))" | |
| 72 | unfolding \<psi>_def by (rule \<mu>.char_approx3, auto) | |
| 73 | also have "?t^2 * \<sigma>\<^sup>2 = t^2 / n" | |
| 74 | using \<sigma>_pos by (simp add: power_divide) | |
| 75 | also have "t^2 / n / 2 = (t^2 / 2) / n" | |
| 76 | by simp | |
| 63329 | 77 | finally have **: "norm (\<psi> n t - (1 + (-(t^2) / 2) / n)) \<le> | 
| 62083 | 78 | ?t\<^sup>2 / 6 * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t\<bar> * \<bar>x\<bar> ^ 3))" by simp | 
| 79 | ||
| 63329 | 80 | have "norm (\<phi> n t - (complex_of_real (1 + (-(t^2) / 2) / n))^n) \<le> | 
| 62083 | 81 | n * norm (\<psi> n t - (complex_of_real (1 + (-(t^2) / 2) / n)))" | 
| 82 | using n | |
| 83 | by (auto intro!: norm_power_diff \<mu>.cmod_char_le_1 abs_leI | |
| 84 | simp del: of_real_diff simp: of_real_diff[symmetric] divide_le_eq \<phi>_eq \<psi>_def) | |
| 85 | also have "\<dots> \<le> n * (?t\<^sup>2 / 6 * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t\<bar> * \<bar>x\<bar> ^ 3)))" | |
| 86 | by (rule mult_left_mono [OF **], simp) | |
| 63329 | 87 | also have "\<dots> = (t\<^sup>2 / (6 * \<sigma>\<^sup>2) * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t\<bar> * \<bar>x\<bar> ^ 3)))" | 
| 62083 | 88 | using \<sigma>_pos by (simp add: field_simps min_absorb2) | 
| 63329 | 89 | finally show "norm (\<phi> n t - (1 + (-(t^2) / 2) / n)^n) \<le> | 
| 90 | (t\<^sup>2 / (6 * \<sigma>\<^sup>2) * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t\<bar> * \<bar>x\<bar> ^ 3)))" | |
| 62083 | 91 | by simp | 
| 92 | qed | |
| 93 | ||
| 94 | show ?thesis | |
| 95 | proof (rule levy_continuity) | |
| 96 | fix t | |
| 97 | let ?t = "\<lambda>n. t / sqrt (\<sigma>\<^sup>2 * n)" | |
| 98 | have "\<And>x. (\<lambda>n. min (6 * x\<^sup>2) (\<bar>t\<bar> * \<bar>x\<bar> ^ 3 / \<bar>sqrt (\<sigma>\<^sup>2 * real n)\<bar>)) \<longlonglongrightarrow> 0" | |
| 63329 | 99 | using \<sigma>_pos | 
| 62083 | 100 | by (auto simp: real_sqrt_mult min_absorb2 | 
| 101 | intro!: tendsto_min[THEN tendsto_eq_rhs] sqrt_at_top[THEN filterlim_compose] | |
| 102 | filterlim_tendsto_pos_mult_at_top filterlim_at_top_imp_at_infinity | |
| 103 | tendsto_divide_0 filterlim_real_sequentially) | |
| 104 | then have "(\<lambda>n. LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t n\<bar> * \<bar>x\<bar> ^ 3)) \<longlonglongrightarrow> (LINT x|\<mu>. 0)" | |
| 105 | by (intro integral_dominated_convergence [where w = "\<lambda>x. 6 * x^2"]) auto | |
| 106 | then have *: "(\<lambda>n. t\<^sup>2 / (6 * \<sigma>\<^sup>2) * (LINT x|\<mu>. min (6 * x\<^sup>2) (\<bar>?t n\<bar> * \<bar>x\<bar> ^ 3))) \<longlonglongrightarrow> 0" | |
| 107 | by (simp only: integral_zero tendsto_mult_right_zero) | |
| 108 | ||
| 109 | have "(\<lambda>n. complex_of_real ((1 + (-(t^2) / 2) / n)^n)) \<longlonglongrightarrow> complex_of_real (exp (-(t^2) / 2))" | |
| 110 | by (rule isCont_tendsto_compose [OF _ tendsto_exp_limit_sequentially]) auto | |
| 111 | then have "(\<lambda>n. \<phi> n t) \<longlonglongrightarrow> complex_of_real (exp (-(t^2) / 2))" | |
| 112 | by (rule Lim_transform) (rule Lim_null_comparison [OF main *]) | |
| 113 | then show "(\<lambda>n. char (distr M borel (?S' n)) t) \<longlonglongrightarrow> char std_normal_distribution t" | |
| 114 | by (simp add: \<phi>_def char_std_normal_distribution) | |
| 115 | qed (auto intro!: real_dist_normal_dist simp: S_def) | |
| 116 | qed | |
| 117 | ||
| 64009 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 118 | theorem (in prob_space) central_limit_theorem: | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 119 | fixes X :: "nat \<Rightarrow> 'a \<Rightarrow> real" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 120 | and \<mu> :: "real measure" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 121 | and \<sigma> :: real | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 122 | and S :: "nat \<Rightarrow> 'a \<Rightarrow> real" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 123 | assumes X_indep: "indep_vars (\<lambda>i. borel) X UNIV" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 124 | and X_mean: "\<And>n. expectation (X n) = m" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 125 | and \<sigma>_pos: "\<sigma> > 0" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 126 | and X_square_integrable: "\<And>n. integrable M (\<lambda>x. (X n x)\<^sup>2)" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 127 | and X_variance: "\<And>n. variance (X n) = \<sigma>\<^sup>2" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 128 | and X_distrib: "\<And>n. distr M borel (X n) = \<mu>" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 129 | defines "X' i x \<equiv> X i x - m" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 130 | shows "weak_conv_m (\<lambda>n. distr M borel (\<lambda>x. (\<Sum>i<n. X' i x) / sqrt (n*\<sigma>\<^sup>2))) std_normal_distribution" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 131 | proof (intro central_limit_theorem_zero_mean) | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 132 | show "indep_vars (\<lambda>i. borel) X' UNIV" | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 133 | unfolding X'_def[abs_def] using X_indep by (rule indep_vars_compose2) auto | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 134 | have X_rv [simp, measurable]: "random_variable borel (X n)" for n | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 135 | using X_indep unfolding indep_vars_def2 by simp | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 136 | have X_integrable [simp, intro]: "integrable M (X n)" for n | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 137 | by (rule square_integrable_imp_integrable[OF _ X_square_integrable]) simp_all | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 138 | show "expectation (X' n) = 0" for n | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
64272diff
changeset | 139 | using X_mean by (auto simp: X'_def[abs_def] prob_space) | 
| 64009 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 140 | show "\<sigma> > 0" "integrable M (\<lambda>x. (X' n x)\<^sup>2)" "variance (X' n) = \<sigma>\<^sup>2" for n | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 141 | using \<open>0 < \<sigma>\<close> X_integrable X_mean X_square_integrable X_variance unfolding X'_def | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 142 | by (auto simp: prob_space power2_diff) | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 143 | show "distr M borel (X' n) = distr \<mu> borel (\<lambda>x. x - m)" for n | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 144 | unfolding X_distrib[of n, symmetric] using X_integrable | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 145 | by (subst distr_distr) (auto simp: X'_def[abs_def] comp_def) | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 146 | qed | 
| 
a5d293f1af80
Probability: variant of central limit theorem with non-zero mean
 hoelzl parents: 
63329diff
changeset | 147 | |
| 62083 | 148 | end |