src/HOL/Probability/Weak_Convergence.thy
author hoelzl
Tue, 09 Feb 2016 09:21:10 +0100
changeset 62375 670063003ad3
parent 62083 7582b39f51ed
child 62975 1d066f6ab25d
permissions -rw-r--r--
add extended nonnegative real numbers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     1
(*
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     2
  Theory: Weak_Convergence.thy
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     3
  Authors: Jeremy Avigad, Luke Serafin
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     4
*)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     5
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     6
section \<open>Weak Convergence of Functions and Distributions\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     7
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     8
text \<open>Properties of weak convergence of functions and measures, including the portmanteau theorem.\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     9
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    10
theory Weak_Convergence
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    11
  imports Distribution_Functions
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    12
begin
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    13
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    14
section \<open>Weak Convergence of Functions\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    15
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    16
definition
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    17
  weak_conv :: "(nat \<Rightarrow> (real \<Rightarrow> real)) \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    18
where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    19
  "weak_conv F_seq F \<equiv> \<forall>x. isCont F x \<longrightarrow> (\<lambda>n. F_seq n x) \<longlonglongrightarrow> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    20
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    21
section \<open>Weak Convergence of Distributions\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    22
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    23
definition
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    24
  weak_conv_m :: "(nat \<Rightarrow> real measure) \<Rightarrow> real measure \<Rightarrow> bool"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    25
where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    26
  "weak_conv_m M_seq M \<equiv> weak_conv (\<lambda>n. cdf (M_seq n)) (cdf M)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    27
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    28
section \<open>Skorohod's theorem\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    29
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
    30
locale right_continuous_mono =
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    31
  fixes f :: "real \<Rightarrow> real" and a b :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    32
  assumes cont: "\<And>x. continuous (at_right x) f"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    33
  assumes mono: "mono f"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
    34
  assumes bot: "(f \<longlongrightarrow> a) at_bot"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    35
  assumes top: "(f \<longlongrightarrow> b) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    36
begin
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    37
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    38
abbreviation I :: "real \<Rightarrow> real" where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    39
  "I \<omega> \<equiv> Inf {x. \<omega> \<le> f x}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    40
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    41
lemma pseudoinverse: assumes "a < \<omega>" "\<omega> < b" shows "\<omega> \<le> f x \<longleftrightarrow> I \<omega> \<le> x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    42
proof
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    43
  let ?F = "{x. \<omega> \<le> f x}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    44
  obtain y where "f y < \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    45
    by (metis eventually_happens' trivial_limit_at_bot_linorder order_tendstoD(2) bot \<open>a < \<omega>\<close>)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    46
  with mono have bdd: "bdd_below ?F"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    47
    by (auto intro!: bdd_belowI[of _ y] elim: mono_invE[OF _ less_le_trans])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    48
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    49
  have ne: "?F \<noteq> {}"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
    50
    using order_tendstoD(1)[OF top \<open>\<omega> < b\<close>]
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    51
    by (auto dest!: eventually_happens'[OF trivial_limit_at_top_linorder] intro: less_imp_le)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    52
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    53
  show "\<omega> \<le> f x \<Longrightarrow> I \<omega> \<le> x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    54
    by (auto intro!: cInf_lower bdd)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    55
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    56
  { assume *: "I \<omega> \<le> x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    57
    have "\<omega> \<le> (INF s:{x. \<omega> \<le> f x}. f s)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    58
      by (rule cINF_greatest[OF ne]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    59
    also have "\<dots> = f (I \<omega>)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    60
      using continuous_at_Inf_mono[OF mono cont ne bdd] ..
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    61
    also have "\<dots> \<le> f x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    62
      using * by (rule monoD[OF \<open>mono f\<close>])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    63
    finally show "\<omega> \<le> f x" . }
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    64
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    65
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    66
lemma pseudoinverse': "\<forall>\<omega>\<in>{a<..<b}. \<forall>x. \<omega> \<le> f x \<longleftrightarrow> I \<omega> \<le> x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    67
  by (intro ballI allI impI pseudoinverse) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    68
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    69
lemma mono_I: "mono_on I {a <..< b}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    70
  unfolding mono_on_def by (metis order.trans order.refl pseudoinverse')
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    71
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    72
end
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    73
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    74
locale cdf_distribution = real_distribution
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    75
begin
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    76
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    77
abbreviation "C \<equiv> cdf M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    78
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    79
sublocale right_continuous_mono C 0 1
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    80
  by standard
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    81
     (auto intro: cdf_nondecreasing cdf_is_right_cont cdf_lim_at_top_prob cdf_lim_at_bot monoI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    82
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    83
lemma measurable_C[measurable]: "C \<in> borel_measurable borel"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    84
  by (intro borel_measurable_mono mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    85
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    86
lemma measurable_CI[measurable]: "I \<in> borel_measurable (restrict_space borel {0<..<1})"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    87
  by (intro borel_measurable_mono_on_fnc mono_I)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    88
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    89
lemma emeasure_distr_I: "emeasure (distr (restrict_space lborel {0<..<1::real}) borel I) UNIV = 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    90
  by (simp add: emeasure_distr space_restrict_space emeasure_restrict_space )
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    91
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    92
lemma distr_I_eq_M: "distr (restrict_space lborel {0<..<1::real}) borel I = M" (is "?I = _")
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    93
proof (intro cdf_unique ext)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    94
  let ?\<Omega> = "restrict_space lborel {0<..<1}::real measure"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    95
  interpret \<Omega>: prob_space ?\<Omega>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    96
    by (auto simp add: emeasure_restrict_space space_restrict_space intro!: prob_spaceI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    97
  show "real_distribution ?I"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    98
    by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    99
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   100
  fix x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   101
  have "cdf ?I x = measure lborel {\<omega>\<in>{0<..<1}. \<omega> \<le> C x}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   102
    by (subst cdf_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   103
       (auto simp: pseudoinverse[symmetric] measure_distr space_restrict_space measure_restrict_space
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   104
             intro!: arg_cong2[where f="measure"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   105
  also have "\<dots> = measure lborel {0 <..< C x}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   106
    using cdf_bounded_prob[of x] AE_lborel_singleton[of "C x"]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   107
    by (auto intro!: arg_cong[where f=real_of_ereal] emeasure_eq_AE simp: measure_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   108
  also have "\<dots> = C x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   109
    by (simp add: cdf_nonneg)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   110
  finally show "cdf (distr ?\<Omega> borel I) x = C x" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   111
qed standard
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   112
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   113
end
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   114
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   115
context
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   116
  fixes \<mu> :: "nat \<Rightarrow> real measure"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   117
    and M :: "real measure"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   118
  assumes \<mu>: "\<And>n. real_distribution (\<mu> n)"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   119
  assumes M: "real_distribution M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   120
  assumes \<mu>_to_M: "weak_conv_m \<mu> M"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   121
begin
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   122
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   123
(* state using obtains? *)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   124
theorem Skorohod:
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   125
 "\<exists> (\<Omega> :: real measure) (Y_seq :: nat \<Rightarrow> real \<Rightarrow> real) (Y :: real \<Rightarrow> real).
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   126
    prob_space \<Omega> \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   127
    (\<forall>n. Y_seq n \<in> measurable \<Omega> borel) \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   128
    (\<forall>n. distr \<Omega> borel (Y_seq n) = \<mu> n) \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   129
    Y \<in> measurable \<Omega> lborel \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   130
    distr \<Omega> borel Y = M \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   131
    (\<forall>x \<in> space \<Omega>. (\<lambda>n. Y_seq n x) \<longlonglongrightarrow> Y x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   132
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   133
  interpret \<mu>: cdf_distribution "\<mu> n" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   134
    unfolding cdf_distribution_def by (rule \<mu>)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   135
  interpret M: cdf_distribution M
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   136
    unfolding cdf_distribution_def by (rule M)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   137
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   138
  have conv: "measure M {x} = 0 \<Longrightarrow> (\<lambda>n. \<mu>.C n x) \<longlonglongrightarrow> M.C x" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   139
    using \<mu>_to_M M.isCont_cdf by (auto simp: weak_conv_m_def weak_conv_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   140
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   141
  let ?\<Omega> = "restrict_space lborel {0<..<1} :: real measure"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   142
  have "prob_space ?\<Omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   143
    by (auto simp: space_restrict_space emeasure_restrict_space intro!: prob_spaceI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   144
  interpret \<Omega>: prob_space ?\<Omega>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   145
    by fact
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   146
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   147
  have Y_distr: "distr ?\<Omega> borel M.I = M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   148
    by (rule M.distr_I_eq_M)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   149
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   150
  have Y_cts_cnv: "(\<lambda>n. \<mu>.I n \<omega>) \<longlonglongrightarrow> M.I \<omega>"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   151
    if \<omega>: "\<omega> \<in> {0<..<1}" "isCont M.I \<omega>" for \<omega> :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   152
  proof (intro limsup_le_liminf_real)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   153
    show "liminf (\<lambda>n. \<mu>.I n \<omega>) \<ge> M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   154
      unfolding le_Liminf_iff
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   155
    proof safe
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   156
      fix B :: ereal assume B: "B < M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   157
      then show "\<forall>\<^sub>F n in sequentially. B < \<mu>.I n \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   158
      proof (cases B)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   159
        case (real r)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   160
        with B have r: "r < M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   161
          by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   162
        then obtain x where x: "r < x" "x < M.I \<omega>" "measure M {x} = 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   163
          using open_minus_countable[OF M.countable_support, of "{r<..<M.I \<omega>}"] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   164
        then have Fx_less: "M.C x < \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   165
          using M.pseudoinverse' \<omega> not_less by blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   166
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   167
        have "\<forall>\<^sub>F n in sequentially. \<mu>.C n x < \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   168
          using order_tendstoD(2)[OF conv[OF x(3)] Fx_less] .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   169
        then have "\<forall>\<^sub>F n in sequentially. x < \<mu>.I n \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   170
          by eventually_elim (insert \<omega> \<mu>.pseudoinverse[symmetric], simp add: not_le[symmetric])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   171
        then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   172
          by eventually_elim (insert x(1), simp add: real)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   173
      qed auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   174
    qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   175
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   176
    have *: "limsup (\<lambda>n. \<mu>.I n \<omega>) \<le> M.I \<omega>'"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   177
      if \<omega>': "0 < \<omega>'" "\<omega>' < 1" "\<omega> < \<omega>'" for \<omega>' :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   178
    proof (rule dense_ge_bounded)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   179
      fix B' assume "ereal (M.I \<omega>') < B'" "B' < ereal (M.I \<omega>' + 1)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   180
      then obtain B where "M.I \<omega>' < B" and [simp]: "B' = ereal B"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   181
        by (cases B') auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   182
      then obtain y where y: "M.I \<omega>' < y" "y < B" "measure M {y} = 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   183
        using open_minus_countable[OF M.countable_support, of "{M.I \<omega>'<..<B}"] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   184
      then have "\<omega>' \<le> M.C (M.I \<omega>')"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   185
        using M.pseudoinverse' \<omega>' by (metis greaterThanLessThan_iff order_refl)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   186
      also have "... \<le> M.C y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   187
        using M.mono y unfolding mono_def by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   188
      finally have Fy_gt: "\<omega> < M.C y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   189
        using \<omega>'(3) by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   190
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   191
      have "\<forall>\<^sub>F n in sequentially. \<omega> \<le> \<mu>.C n y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   192
        using order_tendstoD(1)[OF conv[OF y(3)] Fy_gt] by eventually_elim (rule less_imp_le)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   193
      then have 2: "\<forall>\<^sub>F n in sequentially. \<mu>.I n \<omega> \<le> ereal y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   194
        by simp (subst \<mu>.pseudoinverse'[rule_format, OF \<omega>(1), symmetric])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   195
      then show "limsup (\<lambda>n. \<mu>.I n \<omega>) \<le> B'"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   196
        using \<open>y < B\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   197
        by (intro Limsup_bounded[rotated]) (auto intro: le_less_trans elim: eventually_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   198
    qed simp
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   199
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   200
    have **: "(M.I \<longlongrightarrow> ereal (M.I \<omega>)) (at_right \<omega>)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   201
      using \<omega>(2) by (auto intro: tendsto_within_subset simp: continuous_at)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   202
    show "limsup (\<lambda>n. \<mu>.I n \<omega>) \<le> M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   203
      using \<omega>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   204
      by (intro tendsto_le_const[OF trivial_limit_at_right_real **])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   205
         (auto intro!: exI[of _ 1] * simp: eventually_at_right[of _ 1])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   206
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   207
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   208
  let ?D = "{\<omega>\<in>{0<..<1}. \<not> isCont M.I \<omega>}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   209
  have D_countable: "countable ?D"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   210
    using mono_on_ctble_discont[OF M.mono_I] by (simp add: at_within_open[of _ "{0 <..< 1}"] cong: conj_cong)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   211
  hence D: "emeasure ?\<Omega> ?D = 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   212
    using emeasure_lborel_countable[OF D_countable]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   213
    by (subst emeasure_restrict_space) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   214
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   215
  def Y' \<equiv> "\<lambda>\<omega>. if \<omega> \<in> ?D then 0 else M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   216
  have Y'_AE: "AE \<omega> in ?\<Omega>. Y' \<omega> = M.I \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   217
    by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y'_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   218
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   219
  def Y_seq' \<equiv> "\<lambda>n \<omega>. if \<omega> \<in> ?D then 0 else \<mu>.I n \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   220
  have Y_seq'_AE: "\<And>n. AE \<omega> in ?\<Omega>. Y_seq' n \<omega> = \<mu>.I n \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   221
    by (rule AE_I [OF _ D]) (auto simp: space_restrict_space sets_restrict_space_iff Y_seq'_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   222
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   223
  have Y'_cnv: "\<forall>\<omega>\<in>{0<..<1}. (\<lambda>n. Y_seq' n \<omega>) \<longlonglongrightarrow> Y' \<omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   224
    by (auto simp: Y'_def Y_seq'_def Y_cts_cnv)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   225
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   226
  have [simp]: "Y_seq' n \<in> borel_measurable ?\<Omega>" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   227
    by (rule measurable_discrete_difference[of "\<mu>.I n" _ _ ?D])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   228
       (insert \<mu>.measurable_CI[of n] D_countable, auto simp: sets_restrict_space Y_seq'_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   229
  moreover have "distr ?\<Omega> borel (Y_seq' n) = \<mu> n" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   230
    using \<mu>.distr_I_eq_M [of n] Y_seq'_AE [of n]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   231
    by (subst distr_cong_AE[where f = "Y_seq' n" and g = "\<mu>.I n"], auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   232
  moreover have [simp]: "Y' \<in> borel_measurable ?\<Omega>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   233
    by (rule measurable_discrete_difference[of M.I _ _ ?D])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   234
       (insert M.measurable_CI D_countable, auto simp: sets_restrict_space Y'_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   235
  moreover have "distr ?\<Omega> borel Y' = M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   236
    using M.distr_I_eq_M Y'_AE
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   237
    by (subst distr_cong_AE[where f = Y' and g = M.I], auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   238
  ultimately have "prob_space ?\<Omega> \<and> (\<forall>n. Y_seq' n \<in> borel_measurable ?\<Omega>) \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   239
    (\<forall>n. distr ?\<Omega> borel (Y_seq' n) = \<mu> n) \<and> Y' \<in> measurable ?\<Omega> lborel \<and> distr ?\<Omega> borel Y' = M \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   240
    (\<forall>x\<in>space ?\<Omega>. (\<lambda>n. Y_seq' n x) \<longlonglongrightarrow> Y' x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   241
    using Y'_cnv \<open>prob_space ?\<Omega>\<close> by (auto simp: space_restrict_space)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   242
  thus ?thesis by metis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   243
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   244
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   245
text \<open>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   246
  The Portmanteau theorem, that is, the equivalence of various definitions of weak convergence.
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   247
\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   248
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   249
theorem weak_conv_imp_bdd_ae_continuous_conv:
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   250
  fixes
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   251
    f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   252
  assumes
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   253
    discont_null: "M ({x. \<not> isCont f x}) = 0" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   254
    f_bdd: "\<And>x. norm (f x) \<le> B" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   255
    [measurable]: "f \<in> borel_measurable borel"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   256
  shows
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   257
    "(\<lambda> n. integral\<^sup>L (\<mu> n) f) \<longlonglongrightarrow> integral\<^sup>L M f"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   258
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   259
  have "0 \<le> B"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   260
    using norm_ge_zero f_bdd by (rule order_trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   261
  note Skorohod
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   262
  then obtain Omega Y_seq Y where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   263
    ps_Omega [simp]: "prob_space Omega" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   264
    Y_seq_measurable [measurable]: "\<And>n. Y_seq n \<in> borel_measurable Omega" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   265
    distr_Y_seq: "\<And>n. distr Omega borel (Y_seq n) = \<mu> n" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   266
    Y_measurable [measurable]: "Y \<in> borel_measurable Omega" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   267
    distr_Y: "distr Omega borel Y = M" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   268
    YnY: "\<And>x :: real. x \<in> space Omega \<Longrightarrow> (\<lambda>n. Y_seq n x) \<longlonglongrightarrow> Y x"  by force
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   269
  interpret prob_space Omega by fact
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   270
  have *: "emeasure Omega (Y -` {x. \<not> isCont f x} \<inter> space Omega) = 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   271
    by (subst emeasure_distr [symmetric, where N=borel]) (auto simp: distr_Y discont_null)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   272
  have *: "AE x in Omega. (\<lambda>n. f (Y_seq n x)) \<longlonglongrightarrow> f (Y x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   273
    by (rule AE_I [OF _ *]) (auto intro: isCont_tendsto_compose YnY)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   274
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   275
    by (auto intro!: integral_dominated_convergence[where w="\<lambda>x. B"]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   276
             simp: f_bdd * integral_distr distr_Y_seq [symmetric] distr_Y [symmetric])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   277
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   278
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   279
theorem weak_conv_imp_integral_bdd_continuous_conv:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   280
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   281
  assumes
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   282
    "\<And>x. isCont f x" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   283
    "\<And>x. norm (f x) \<le> B"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   284
  shows
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   285
    "(\<lambda> n. integral\<^sup>L (\<mu> n) f) \<longlonglongrightarrow> integral\<^sup>L M f"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   286
  using assms
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   287
  by (intro weak_conv_imp_bdd_ae_continuous_conv)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   288
     (auto intro!: borel_measurable_continuous_on1 continuous_at_imp_continuous_on)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   289
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   290
theorem weak_conv_imp_continuity_set_conv:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   291
  fixes f :: "real \<Rightarrow> real"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   292
  assumes [measurable]: "A \<in> sets borel" and "M (frontier A) = 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   293
  shows "(\<lambda>n. measure (\<mu> n) A) \<longlonglongrightarrow> measure M A"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   294
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   295
  interpret M: real_distribution M by fact
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   296
  interpret \<mu>: real_distribution "\<mu> n" for n by fact
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   297
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   298
  have "(\<lambda>n. (\<integral>x. indicator A x \<partial>\<mu> n) :: real) \<longlonglongrightarrow> (\<integral>x. indicator A x \<partial>M)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   299
    by (intro weak_conv_imp_bdd_ae_continuous_conv[where B=1])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   300
       (auto intro: assms simp: isCont_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   301
  then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   302
    by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   303
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   304
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   305
end
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   306
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   307
definition
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   308
  cts_step :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   309
where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   310
  "cts_step a b x \<equiv> if x \<le> a then 1 else if x \<ge> b then 0 else (b - x) / (b - a)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   311
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   312
lemma cts_step_uniformly_continuous:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   313
  assumes [arith]: "a < b"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   314
  shows "uniformly_continuous_on UNIV (cts_step a b)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   315
  unfolding uniformly_continuous_on_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   316
proof clarsimp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   317
  fix e :: real assume [arith]: "0 < e"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   318
  let ?d = "min (e * (b - a)) (b - a)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   319
  have "?d > 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   320
    by (auto simp add: field_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   321
  moreover have "dist x' x < ?d \<Longrightarrow> dist (cts_step a b x') (cts_step a b x) < e" for x x'
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   322
    by (auto simp: dist_real_def divide_simps cts_step_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   323
  ultimately show "\<exists>d > 0. \<forall>x x'. dist x' x < d \<longrightarrow> dist (cts_step a b x') (cts_step a b x) < e"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   324
    by blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   325
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   326
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   327
lemma (in real_distribution) integrable_cts_step: "a < b \<Longrightarrow> integrable M (cts_step a b)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   328
  by (rule integrable_const_bound [of _ 1]) (auto simp: cts_step_def[abs_def])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   329
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   330
lemma (in real_distribution) cdf_cts_step:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   331
  assumes [arith]: "x < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   332
  shows "cdf M x \<le> integral\<^sup>L M (cts_step x y)" and "integral\<^sup>L M (cts_step x y) \<le> cdf M y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   333
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   334
  have "cdf M x = integral\<^sup>L M (indicator {..x})"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   335
    by (simp add: cdf_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   336
  also have "\<dots> \<le> expectation (cts_step x y)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   337
    by (intro integral_mono integrable_cts_step) (auto simp: cts_step_def split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   338
  finally show "cdf M x \<le> expectation (cts_step x y)" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   339
next
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   340
  have "expectation (cts_step x y) \<le> integral\<^sup>L M (indicator {..y})"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   341
    by (intro integral_mono integrable_cts_step) (auto simp: cts_step_def split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   342
  also have "\<dots> = cdf M y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   343
    by (simp add: cdf_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   344
  finally show "expectation (cts_step x y) \<le> cdf M y" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   345
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   346
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   347
context
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   348
  fixes M_seq :: "nat \<Rightarrow> real measure"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   349
    and M :: "real measure"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   350
  assumes distr_M_seq [simp]: "\<And>n. real_distribution (M_seq n)"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   351
  assumes distr_M [simp]: "real_distribution M"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   352
begin
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   353
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   354
theorem continuity_set_conv_imp_weak_conv:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   355
  fixes f :: "real \<Rightarrow> real"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   356
  assumes *: "\<And>A. A \<in> sets borel \<Longrightarrow> M (frontier A) = 0 \<Longrightarrow> (\<lambda> n. (measure (M_seq n) A)) \<longlonglongrightarrow> measure M A"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   357
  shows "weak_conv_m M_seq M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   358
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   359
  interpret real_distribution M by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   360
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   361
    by (auto intro!: * simp: frontier_real_Iic isCont_cdf emeasure_eq_measure weak_conv_m_def weak_conv_def cdf_def2)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   362
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   363
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   364
theorem integral_cts_step_conv_imp_weak_conv:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   365
  assumes integral_conv: "\<And>x y. x < y \<Longrightarrow> (\<lambda>n. integral\<^sup>L (M_seq n) (cts_step x y)) \<longlonglongrightarrow> integral\<^sup>L M (cts_step x y)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   366
  shows "weak_conv_m M_seq M"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   367
  unfolding weak_conv_m_def weak_conv_def
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   368
proof (clarsimp)
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   369
  interpret real_distribution M by (rule distr_M)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   370
  fix x assume "isCont (cdf M) x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   371
  hence left_cont: "continuous (at_left x) (cdf M)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   372
    unfolding continuous_at_split ..
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   373
  { fix y :: real assume [arith]: "x < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   374
    have "limsup (\<lambda>n. cdf (M_seq n) x) \<le> limsup (\<lambda>n. integral\<^sup>L (M_seq n) (cts_step x y))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   375
      by (auto intro!: Limsup_mono always_eventually real_distribution.cdf_cts_step)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   376
    also have "\<dots> = integral\<^sup>L M (cts_step x y)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   377
      by (intro lim_imp_Limsup) (auto intro: integral_conv)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   378
    also have "\<dots> \<le> cdf M y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   379
      by (simp add: cdf_cts_step)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   380
    finally have "limsup (\<lambda>n. cdf (M_seq n) x) \<le> cdf M y" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   381
  } note * = this
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   382
  { fix y :: real assume [arith]: "x > y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   383
    have "cdf M y \<le> ereal (integral\<^sup>L M (cts_step y x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   384
      by (simp add: cdf_cts_step)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   385
    also have "\<dots> = liminf (\<lambda>n. integral\<^sup>L (M_seq n) (cts_step y x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   386
      by (intro lim_imp_Liminf[symmetric]) (auto intro: integral_conv)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   387
    also have "\<dots> \<le> liminf (\<lambda>n. cdf (M_seq n) x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   388
      by (auto intro!: Liminf_mono always_eventually real_distribution.cdf_cts_step)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   389
    finally have "liminf (\<lambda>n. cdf (M_seq n) x) \<ge> cdf M y" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   390
  } note ** = this
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   391
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   392
  have "limsup (\<lambda>n. cdf (M_seq n) x) \<le> cdf M x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   393
  proof (rule tendsto_le_const)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   394
    show "\<forall>\<^sub>F i in at_right x. limsup (\<lambda>xa. ereal (cdf (M_seq xa) x)) \<le> ereal (cdf M i)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   395
      by (subst eventually_at_right[of _ "x + 1"]) (auto simp: * intro: exI [of _ "x+1"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   396
  qed (insert cdf_is_right_cont, auto simp: continuous_within)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   397
  moreover have "cdf M x \<le> liminf (\<lambda>n. cdf (M_seq n) x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   398
  proof (rule tendsto_ge_const)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   399
    show "\<forall>\<^sub>F i in at_left x. ereal (cdf M i) \<le> liminf (\<lambda>xa. ereal (cdf (M_seq xa) x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   400
      by (subst eventually_at_left[of "x - 1"]) (auto simp: ** intro: exI [of _ "x-1"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   401
  qed (insert left_cont, auto simp: continuous_within)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   402
  ultimately show "(\<lambda>n. cdf (M_seq n) x) \<longlonglongrightarrow> cdf M x"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   403
    by (elim limsup_le_liminf_real)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   404
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   405
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   406
theorem integral_bdd_continuous_conv_imp_weak_conv:
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   407
  assumes
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   408
    "\<And>f. (\<And>x. isCont f x) \<Longrightarrow> (\<And>x. abs (f x) \<le> 1) \<Longrightarrow> (\<lambda>n. integral\<^sup>L (M_seq n) f::real) \<longlonglongrightarrow> integral\<^sup>L M f"
62375
670063003ad3 add extended nonnegative real numbers
hoelzl
parents: 62083
diff changeset
   409
  shows
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   410
    "weak_conv_m M_seq M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   411
  apply (rule integral_cts_step_conv_imp_weak_conv [OF assms])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   412
  apply (rule continuous_on_interior)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   413
  apply (rule uniformly_continuous_imp_continuous)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   414
  apply (rule cts_step_uniformly_continuous)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   415
  apply (auto simp: cts_step_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   416
  done
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   417
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   418
end
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   419
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   420
end