src/HOL/Probability/Helly_Selection.thy
author hoelzl
Wed, 06 Jan 2016 12:18:53 +0100
changeset 62083 7582b39f51ed
child 62397 5ae24f33d343
permissions -rw-r--r--
add the proof of the central limit theorem
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: Helly_Selection.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>Helly's selection theorem\<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>The set of bounded, monotone, right continuous functions is sequentially compact\<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 Helly_Selection
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    11
  imports "~~/src/HOL/Library/Diagonal_Subsequence" Weak_Convergence
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
lemma minus_one_less: "x - 1 < (x::real)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    15
  by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    16
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    17
theorem Helly_selection:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    18
  fixes f :: "nat \<Rightarrow> real \<Rightarrow> real"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    19
  assumes rcont: "\<And>n x. continuous (at_right x) (f n)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    20
  assumes mono: "\<And>n. mono (f n)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    21
  assumes bdd: "\<And>n x. \<bar>f n x\<bar> \<le> M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    22
  shows "\<exists>s. subseq s \<and> (\<exists>F. (\<forall>x. continuous (at_right x) F) \<and> mono F \<and> (\<forall>x. \<bar>F x\<bar> \<le> M) \<and>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    23
    (\<forall>x. continuous (at x) F \<longrightarrow> (\<lambda>n. f (s n) x) \<longlonglongrightarrow> F x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    24
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    25
  obtain m :: "real \<Rightarrow> nat" where "bij_betw m \<rat> UNIV"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    26
    using countable_rat Rats_infinite by (erule countableE_infinite)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    27
  then obtain r :: "nat \<Rightarrow> real" where bij: "bij_betw r UNIV \<rat>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    28
    using bij_betw_inv by blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    29
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    30
  have dense_r: "\<And>x y. x < y \<Longrightarrow> \<exists>n. x < r n \<and> r n < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    31
    by (metis Rats_dense_in_real bij f_the_inv_into_f bij_betw_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    32
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    33
  let ?P = "\<lambda>n. \<lambda>s. convergent (\<lambda>k. f (s k) (r n))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    34
  interpret nat: subseqs ?P
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    35
  proof (unfold convergent_def, unfold subseqs_def, auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    36
    fix n :: nat and s :: "nat \<Rightarrow> nat" assume s: "subseq s"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    37
    have "bounded {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    38
      using bounded_closed_interval by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    39
    moreover have "\<And>k. f (s k) (r n) \<in> {-M..M}" 
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    40
      using bdd by (simp add: abs_le_iff minus_le_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    41
    ultimately have "\<exists>l s'. subseq s' \<and> ((\<lambda>k. f (s k) (r n)) \<circ> s') \<longlonglongrightarrow> l"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    42
      using compact_Icc compact_imp_seq_compact seq_compactE by metis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    43
    thus "\<exists>s'. subseq s' \<and> (\<exists>l. (\<lambda>k. f (s (s' k)) (r n)) \<longlonglongrightarrow> l)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    44
      by (auto simp: comp_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    45
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    46
  def d \<equiv> "nat.diagseq"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    47
  have subseq: "subseq d"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    48
    unfolding d_def using nat.subseq_diagseq by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    49
  have rat_cnv: "?P n d" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    50
  proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    51
    have Pn_seqseq: "?P n (nat.seqseq (Suc n))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    52
      by (rule nat.seqseq_holds)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    53
    have 1: "(\<lambda>k. f ((nat.seqseq (Suc n) \<circ> (\<lambda>k. nat.fold_reduce (Suc n) k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    54
      (Suc n + k))) k) (r n)) = (\<lambda>k. f (nat.seqseq (Suc n) k) (r n)) \<circ>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    55
      (\<lambda>k. nat.fold_reduce (Suc n) k (Suc n + k))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    56
      by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    57
    have 2: "?P n (d \<circ> (op + (Suc n)))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    58
      unfolding d_def nat.diagseq_seqseq 1
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    59
      by (intro convergent_subseq_convergent Pn_seqseq nat.subseq_diagonal_rest)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    60
    then obtain L where 3: "(\<lambda>na. f (d (na + Suc n)) (r n)) \<longlonglongrightarrow> L"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    61
      by (auto simp: add.commute dest: convergentD)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    62
    then have "(\<lambda>k. f (d k) (r n)) \<longlonglongrightarrow> L"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    63
      by (rule LIMSEQ_offset)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    64
    then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    65
      by (auto simp: convergent_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    66
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    67
  let ?f = "\<lambda>n. \<lambda>k. f (d k) (r n)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    68
  have lim_f: "?f n \<longlonglongrightarrow> lim (?f n)" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    69
    using rat_cnv convergent_LIMSEQ_iff by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    70
  have lim_bdd: "lim (?f n) \<in> {-M..M}" for n
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    71
  proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    72
    have "closed {-M..M}" using closed_real_atLeastAtMost by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    73
    hence "(\<forall>i. ?f n i \<in> {-M..M}) \<and> ?f n \<longlonglongrightarrow> lim (?f n) \<longrightarrow> lim (?f n) \<in> {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    74
      unfolding closed_sequential_limits by (drule_tac x = "\<lambda>k. f (d k) (r n)" in spec) blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    75
    moreover have "\<forall>i. ?f n i \<in> {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    76
      using bdd by (simp add: abs_le_iff minus_le_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    77
    ultimately show "lim (?f n) \<in> {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    78
      using lim_f by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    79
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    80
  then have limset_bdd: "\<And>x. {lim (?f n) |n. x < r n} \<subseteq> {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    81
    by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    82
  then have bdd_below: "bdd_below {lim (?f n) |n. x < r n}" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    83
    by (metis (mono_tags) bdd_below_Icc bdd_below_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    84
  have r_unbdd: "\<exists>n. x < r n" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    85
    using dense_r[OF less_add_one, of x] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    86
  then have nonempty: "{lim (?f n) |n. x < r n} \<noteq> {}" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    87
    by auto
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
  def F \<equiv> "\<lambda>x. Inf {lim (?f n) |n. x < r n}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    90
  have F_eq: "ereal (F x) = (INF n:{n. x < r n}. ereal (lim (?f n)))" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    91
    unfolding F_def by (subst ereal_Inf'[OF bdd_below nonempty]) (simp add: setcompr_eq_image)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    92
  have mono_F: "mono F"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    93
    using nonempty by (auto intro!: cInf_superset_mono simp: F_def bdd_below mono_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    94
  moreover have "\<And>x. continuous (at_right x) F"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    95
    unfolding continuous_within order_tendsto_iff eventually_at_right[OF less_add_one]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    96
  proof safe
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    97
    show "F x < u \<Longrightarrow> \<exists>b>x. \<forall>y>x. y < b \<longrightarrow> F y < u" for x u
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    98
      unfolding F_def cInf_less_iff[OF nonempty bdd_below] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    99
  next
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   100
    show "\<exists>b>x. \<forall>y>x. y < b \<longrightarrow> l < F y" if l: "l < F x" for x l
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   101
      using less_le_trans[OF l mono_F[THEN monoD, of x]] by (auto intro: less_add_one)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   102
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   103
  moreover
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   104
  { fix x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   105
    have "F x \<in> {-M..M}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   106
      unfolding F_def using limset_bdd bdd_below r_unbdd by (intro closed_subset_contains_Inf) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   107
    then have "\<bar>F x\<bar> \<le> M" by auto }
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   108
  moreover have "(\<lambda>n. f (d n) x) \<longlonglongrightarrow> F x" if cts: "continuous (at x) F" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   109
  proof (rule limsup_le_liminf_real)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   110
    show "limsup (\<lambda>n. f (d n) x) \<le> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   111
    proof (rule tendsto_le_const)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   112
      show "(F \<longlongrightarrow> ereal (F x)) (at_right x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   113
        using cts unfolding continuous_at_split by (auto simp: continuous_within)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   114
      show "\<forall>\<^sub>F i in at_right x. limsup (\<lambda>n. f (d n) x) \<le> F i"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   115
        unfolding eventually_at_right[OF less_add_one]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   116
      proof (rule, rule, rule less_add_one, safe)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   117
        fix y assume y: "x < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   118
        with dense_r obtain N where "x < r N" "r N < y" by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   119
        have *: "y < r n' \<Longrightarrow> lim (?f N) \<le> lim (?f n')" for n'
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   120
          using \<open>r N < y\<close> by (intro LIMSEQ_le[OF lim_f lim_f]) (auto intro!: mono[THEN monoD])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   121
        have "limsup (\<lambda>n. f (d n) x) \<le> limsup (?f N)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   122
          using \<open>x < r N\<close> by (auto intro!: Limsup_mono always_eventually mono[THEN monoD])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   123
        also have "\<dots> = lim (\<lambda>n. ereal (?f N n))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   124
          using rat_cnv[of N] by (force intro!: convergent_limsup_cl simp: convergent_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   125
        also have "\<dots> \<le> F y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   126
          by (auto intro!: INF_greatest * simp: convergent_real_imp_convergent_ereal rat_cnv F_eq)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   127
        finally show "limsup (\<lambda>n. f (d n) x) \<le> F y" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   128
      qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   129
    qed simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   130
    show "F x \<le> liminf (\<lambda>n. f (d n) x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   131
    proof (rule tendsto_ge_const)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   132
      show "(F \<longlongrightarrow> ereal (F x)) (at_left x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   133
        using cts unfolding continuous_at_split by (auto simp: continuous_within)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   134
      show "\<forall>\<^sub>F i in at_left x. F i \<le> liminf (\<lambda>n. f (d n) x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   135
        unfolding eventually_at_left[OF minus_one_less]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   136
      proof (rule, rule, rule minus_one_less, safe)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   137
        fix y assume y: "y < x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   138
        with dense_r obtain N where "y < r N" "r N < x" by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   139
        have "F y \<le> liminf (?f N)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   140
          using \<open>y < r N\<close> by (auto simp: F_eq convergent_real_imp_convergent_ereal
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   141
            rat_cnv convergent_liminf_cl intro!: INF_lower2)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   142
        also have "\<dots> \<le> liminf (\<lambda>n. f (d n) x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   143
          using \<open>r N < x\<close> by (auto intro!: Liminf_mono monoD[OF mono] always_eventually)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   144
        finally show "F y \<le> liminf (\<lambda>n. f (d n) x)" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   145
      qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   146
    qed simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   147
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   148
  ultimately show ?thesis using subseq by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   149
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   150
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   151
(** Weak convergence corollaries to Helly's theorem. **)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   152
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   153
definition
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   154
  tight :: "(nat \<Rightarrow> real measure) \<Rightarrow> bool"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   155
where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   156
  "tight \<mu> \<equiv> (\<forall>n. real_distribution (\<mu> n)) \<and> (\<forall>(\<epsilon>::real)>0. \<exists>a b::real. a < b \<and> (\<forall>n. measure (\<mu> n) {a<..b} > 1 - \<epsilon>))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   157
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   158
(* Can strengthen to equivalence. *)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   159
theorem tight_imp_convergent_subsubsequence:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   160
  assumes \<mu>: "tight \<mu>" "subseq s"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   161
  shows "\<exists>r M. subseq r \<and> real_distribution M \<and> weak_conv_m (\<mu> \<circ> s \<circ> r) M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   162
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   163
  def f \<equiv> "\<lambda>k. cdf (\<mu> (s k))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   164
  interpret \<mu>: real_distribution "\<mu> k" for k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   165
    using \<mu> unfolding tight_def by auto
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 rcont: "\<And>x. continuous (at_right x) (f k)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   168
    and mono: "mono (f k)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   169
    and top: "(f k \<longlongrightarrow> 1) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   170
    and bot: "(f k \<longlongrightarrow> 0) at_bot" for k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   171
    unfolding f_def mono_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   172
    using \<mu>.cdf_nondecreasing \<mu>.cdf_is_right_cont \<mu>.cdf_lim_at_top_prob \<mu>.cdf_lim_at_bot by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   173
  have bdd: "\<bar>f k x\<bar> \<le> 1" for k x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   174
    by (auto simp add: abs_le_iff minus_le_iff f_def \<mu>.cdf_nonneg \<mu>.cdf_bounded_prob)
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
  from Helly_selection[OF rcont mono bdd, of "\<lambda>x. x"] obtain r F
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   177
    where F: "subseq r" "\<And>x. continuous (at_right x) F" "mono F" "\<And>x. \<bar>F x\<bar> \<le> 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   178
    and lim_F: "\<And>x. continuous (at x) F \<Longrightarrow> (\<lambda>n. f (r n) x) \<longlonglongrightarrow> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   179
    by blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   180
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   181
  have "0 \<le> f n x" for n x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   182
    unfolding f_def by (rule \<mu>.cdf_nonneg)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   183
  have F_nonneg: "0 \<le> F x" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   184
  proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   185
    obtain y where "y < x" "isCont F y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   186
      using open_minus_countable[OF mono_ctble_discont[OF \<open>mono F\<close>], of "{..< x}"] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   187
    then have "0 \<le> F y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   188
      by (intro LIMSEQ_le_const[OF lim_F]) (auto simp: f_def \<mu>.cdf_nonneg)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   189
    also have "\<dots> \<le> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   190
      using \<open>y < x\<close> by (auto intro!: monoD[OF \<open>mono F\<close>])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   191
    finally show "0 \<le> F x" .
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   192
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   193
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   194
  have Fab: "\<exists>a b. (\<forall>x\<ge>b. F x \<ge> 1 - \<epsilon>) \<and> (\<forall>x\<le>a. F x \<le> \<epsilon>)" if \<epsilon>: "0 < \<epsilon>" for \<epsilon>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   195
  proof auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   196
    obtain a' b' where a'b': "a' < b'" "\<And>k. measure (\<mu> k) {a'<..b'} > 1 - \<epsilon>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   197
      using \<epsilon> \<mu> by (auto simp: tight_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   198
    obtain a where a: "a < a'" "isCont F a"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   199
      using open_minus_countable[OF mono_ctble_discont[OF \<open>mono F\<close>], of "{..< a'}"] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   200
    obtain b where b: "b' < b" "isCont F b"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   201
      using open_minus_countable[OF mono_ctble_discont[OF \<open>mono F\<close>], of "{b' <..}"] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   202
    have "a < b"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   203
      using a b a'b' by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   204
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   205
    let ?\<mu> = "\<lambda>k. measure (\<mu> (s (r k)))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   206
    have ab: "?\<mu> k {a<..b} > 1 - \<epsilon>" for k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   207
    proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   208
      have "?\<mu> k {a'<..b'} \<le> ?\<mu> k {a<..b}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   209
        using a b by (intro \<mu>.finite_measure_mono) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   210
      then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   211
        using a'b'(2) by (metis less_eq_real_def less_trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   212
    qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   213
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   214
    have "(\<lambda>k. ?\<mu> k {..b}) \<longlonglongrightarrow> F b"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   215
      using b(2) lim_F unfolding f_def cdf_def o_def by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   216
    then have "1 - \<epsilon> \<le> F b"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   217
    proof (rule tendsto_le_const[OF sequentially_bot], intro always_eventually allI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   218
      fix k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   219
      have "1 - \<epsilon> < ?\<mu> k {a<..b}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   220
        using ab by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   221
      also have "\<dots> \<le> ?\<mu> k {..b}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   222
        by (auto intro!: \<mu>.finite_measure_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   223
      finally show "1 - \<epsilon> \<le> ?\<mu> k {..b}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   224
        by (rule less_imp_le)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   225
    qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   226
    then show "\<exists>b. \<forall>x\<ge>b. 1 - \<epsilon> \<le> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   227
      using F unfolding mono_def by (metis order.trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   228
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   229
    have "(\<lambda>k. ?\<mu> k {..a}) \<longlonglongrightarrow> F a"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   230
      using a(2) lim_F unfolding f_def cdf_def o_def by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   231
    then have Fa: "F a \<le> \<epsilon>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   232
    proof (rule tendsto_ge_const[OF sequentially_bot], intro always_eventually allI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   233
      fix k
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   234
      have "?\<mu> k {..a} + ?\<mu> k {a<..b} \<le> 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   235
        by (subst \<mu>.finite_measure_Union[symmetric]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   236
      then show "?\<mu> k {..a} \<le> \<epsilon>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   237
        using ab[of k] by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   238
    qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   239
    then show "\<exists>a. \<forall>x\<le>a. F x \<le> \<epsilon>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   240
      using F unfolding mono_def by (metis order.trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   241
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   242
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   243
  have "(F \<longlongrightarrow> 1) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   244
  proof (rule order_tendstoI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   245
    show "1 < y \<Longrightarrow> \<forall>\<^sub>F x in at_top. F x < y" for y
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   246
      using \<open>\<And>x. \<bar>F x\<bar> \<le> 1\<close> \<open>\<And>x. 0 \<le> F x\<close> by (auto intro: le_less_trans always_eventually)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   247
    fix y :: real assume "y < 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   248
    then obtain z where "y < z" "z < 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   249
      using dense[of y 1] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   250
    with Fab[of "1 - z"] show "\<forall>\<^sub>F x in at_top. y < F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   251
      by (auto simp: eventually_at_top_linorder intro: less_le_trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   252
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   253
  moreover
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   254
  have "(F \<longlongrightarrow> 0) at_bot"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   255
  proof (rule order_tendstoI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   256
    show "y < 0 \<Longrightarrow> \<forall>\<^sub>F x in at_bot. y < F x" for y
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   257
      using \<open>\<And>x. 0 \<le> F x\<close> by (auto intro: less_le_trans always_eventually)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   258
    fix y :: real assume "0 < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   259
    then obtain z where "0 < z" "z < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   260
      using dense[of 0 y] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   261
    with Fab[of z] show "\<forall>\<^sub>F x in at_bot. F x < y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   262
      by (auto simp: eventually_at_bot_linorder intro: le_less_trans)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   263
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   264
  ultimately have M: "real_distribution (interval_measure F)" "cdf (interval_measure F) = F"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   265
    using F by (auto intro!: real_distribution_interval_measure cdf_interval_measure simp: mono_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   266
  with lim_F lim_subseq M have "weak_conv_m (\<mu> \<circ> s \<circ> r) (interval_measure F)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   267
    by (auto simp: weak_conv_def weak_conv_m_def f_def comp_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   268
  then show "\<exists>r M. subseq r \<and> (real_distribution M \<and> weak_conv_m (\<mu> \<circ> s \<circ> r) M)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   269
    using F M by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   270
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   271
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   272
corollary tight_subseq_weak_converge:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   273
  fixes \<mu> :: "nat \<Rightarrow> real measure" and M :: "real measure"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   274
  assumes "\<And>n. real_distribution (\<mu> n)" "real_distribution M" and tight: "tight \<mu>" and
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   275
    subseq: "\<And>s \<nu>. subseq s \<Longrightarrow> real_distribution \<nu> \<Longrightarrow> weak_conv_m (\<mu> \<circ> s) \<nu> \<Longrightarrow> weak_conv_m (\<mu> \<circ> s) M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   276
  shows "weak_conv_m \<mu> M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   277
proof (rule ccontr)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   278
  def f \<equiv> "\<lambda>n. cdf (\<mu> n)" and F \<equiv> "cdf M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   279
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   280
  assume "\<not> weak_conv_m \<mu> M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   281
  then obtain x where x: "isCont F x" "\<not> (\<lambda>n. f n x) \<longlonglongrightarrow> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   282
    by (auto simp: weak_conv_m_def weak_conv_def f_def F_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   283
  then obtain \<epsilon> where "\<epsilon> > 0" and "infinite {n. \<not> dist (f n x) (F x) < \<epsilon>}"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   284
    by (auto simp: tendsto_iff not_eventually INFM_iff_infinite cofinite_eq_sequentially[symmetric])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   285
  then obtain s where s: "\<And>n. \<not> dist (f (s n) x) (F x) < \<epsilon>" and "subseq s"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   286
    using enumerate_in_set enumerate_mono by (fastforce simp: subseq_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   287
  then obtain r \<nu> where r: "subseq r" "real_distribution \<nu>" "weak_conv_m (\<mu> \<circ> s \<circ> r) \<nu>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   288
    using tight_imp_convergent_subsubsequence[OF tight] by blast
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   289
  then have "weak_conv_m (\<mu> \<circ> (s \<circ> r)) M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   290
    using \<open>subseq s\<close> r by (intro subseq subseq_o) (auto simp: comp_assoc)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   291
  then have "(\<lambda>n. f (s (r n)) x) \<longlonglongrightarrow> F x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   292
    using x by (auto simp: weak_conv_m_def weak_conv_def F_def f_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   293
  then show False
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   294
    using s \<open>\<epsilon> > 0\<close> by (auto dest: tendstoD)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   295
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   296
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   297
end