src/HOL/Probability/Sinc_Integral.thy
author wenzelm
Fri, 07 Oct 2016 21:09:43 +0200
changeset 64091 f8dfba90e73f
parent 63886 685fb01256af
child 65578 e4997c181cce
permissions -rw-r--r--
more liberal parsing for old AFP logs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63167
diff changeset
     1
(*  Title:     HOL/Probability/Sinc_Integral.thy
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63167
diff changeset
     2
    Authors:   Jeremy Avigad (CMU), Luke Serafin (CMU)
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63167
diff changeset
     3
    Authors:   Johannes Hölzl, TU München
62083
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>Integral of sinc\<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
theory Sinc_Integral
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
     9
  imports Distributions
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    10
begin
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    11
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    12
subsection \<open>Various preparatory integrals\<close>
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
text \<open> Naming convention
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    15
The theorem name consists of the following parts:
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62975
diff changeset
    16
  \<^item> Kind of integral: \<open>has_bochner_integral\<close> / \<open>integrable\<close> / \<open>LBINT\<close>
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    17
  \<^item> Interval: Interval (0 / infinity / open / closed) (infinity / open / closed)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    18
  \<^item> Name of the occurring constants: power, exp, m (for minus), scale, sin, $\ldots$
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    19
\<close>
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
lemma has_bochner_integral_I0i_power_exp_m':
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    22
  "has_bochner_integral lborel (\<lambda>x. x^k * exp (-x) * indicator {0 ..} x::real) (fact k)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    23
  using nn_intergal_power_times_exp_Ici[of k]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    24
  by (intro has_bochner_integral_nn_integral)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    25
     (auto simp: nn_integral_set_ennreal split: split_indicator)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    26
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    27
lemma has_bochner_integral_I0i_power_exp_m:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    28
  "has_bochner_integral lborel (\<lambda>x. x^k * exp (-x) * indicator {0 <..} x::real) (fact k)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    29
  using AE_lborel_singleton[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    30
  by (intro has_bochner_integral_cong_AE[THEN iffD1, OF _ _ _ has_bochner_integral_I0i_power_exp_m'])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    31
     (auto split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    32
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    33
lemma integrable_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> set_integrable lborel {0 <..} (\<lambda>x. exp (-(x * u)))"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    34
  using lborel_integrable_real_affine_iff[of u "\<lambda>x. indicator {0 <..} x *\<^sub>R exp (- x)" 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    35
        has_bochner_integral_I0i_power_exp_m[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    36
  by (simp add: indicator_def zero_less_mult_iff mult_ac integrable.intros)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    37
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    38
lemma LBINT_I0i_exp_mscale: "0 < (u::real) \<Longrightarrow> LBINT x=0..\<infinity>. exp (-(x * u)) = 1 / u"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    39
  using lborel_integral_real_affine[of u "\<lambda>x. indicator {0<..} x *\<^sub>R exp (- x)" 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    40
        has_bochner_integral_I0i_power_exp_m[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    41
  by (auto simp: indicator_def zero_less_mult_iff interval_lebesgue_integral_0_infty field_simps
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    42
           dest!: has_bochner_integral_integral_eq)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    43
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    44
lemma LBINT_I0c_exp_mscale_sin:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    45
  "LBINT x=0..t. exp (-(u * x)) * sin x =
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    46
    (1 / (1 + u^2)) * (1 - exp (-(u * t)) * (u * sin t + cos t))" (is "_ = ?F t")
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    47
  unfolding zero_ereal_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    48
proof (subst interval_integral_FTC_finite)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    49
  show "(?F has_vector_derivative exp (- (u * x)) * sin x) (at x within {min 0 t..max 0 t})" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    50
    by (auto intro!: derivative_eq_intros
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    51
             simp: has_field_derivative_iff_has_vector_derivative[symmetric] power2_eq_square)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    52
       (simp_all add: field_simps add_nonneg_eq_0_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    53
qed (auto intro: continuous_at_imp_continuous_on)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    54
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    55
lemma LBINT_I0i_exp_mscale_sin:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    56
  assumes "0 < x"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    57
  shows "LBINT u=0..\<infinity>. \<bar>exp (-u * x) * sin x\<bar> = \<bar>sin x\<bar> / x"
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    58
proof (subst interval_integral_FTC_nonneg)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    59
  let ?F = "\<lambda>u. 1 / x * (1 - exp (- u * x)) * \<bar>sin x\<bar>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    60
  show "\<And>t. (?F has_real_derivative \<bar>exp (- t * x) * sin x\<bar>) (at t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    61
    using \<open>0 < x\<close> by (auto intro!: derivative_eq_intros simp: abs_mult)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    62
  show "((?F \<circ> real_of_ereal) \<longlongrightarrow> 0) (at_right 0)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    63
    using \<open>0 < x\<close> by (auto simp: zero_ereal_def ereal_tendsto_simps intro!: tendsto_eq_intros)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    64
  have *: "((\<lambda>t. exp (- t * x)) \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    65
    using \<open>0 < x\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    66
    by (auto intro!: exp_at_bot[THEN filterlim_compose] filterlim_tendsto_pos_mult_at_top filterlim_ident
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    67
             simp: filterlim_uminus_at_bot mult.commute[of _ x])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    68
  show "((?F \<circ> real_of_ereal) \<longlongrightarrow> \<bar>sin x\<bar> / x) (at_left \<infinity>)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    69
    using \<open>0 < x\<close> unfolding ereal_tendsto_simps
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    70
    by (intro filterlim_compose[OF _ *]) (auto intro!: tendsto_eq_intros filterlim_ident)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    71
qed auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    72
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    73
lemma
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    74
  shows integrable_inverse_1_plus_square:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    75
      "set_integrable lborel (einterval (-\<infinity>) \<infinity>) (\<lambda>x. inverse (1 + x^2))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    76
  and LBINT_inverse_1_plus_square:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    77
      "LBINT x=-\<infinity>..\<infinity>. inverse (1 + x^2) = pi"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    78
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    79
  have 1: "- (pi / 2) < x \<Longrightarrow> x * 2 < pi \<Longrightarrow> (tan has_real_derivative 1 + (tan x)\<^sup>2) (at x)" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    80
    using cos_gt_zero_pi[of x] by (subst tan_sec) (auto intro!: DERIV_tan simp: power_inverse)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    81
  have 2: "- (pi / 2) < x \<Longrightarrow> x * 2 < pi \<Longrightarrow> isCont (\<lambda>x. 1 + (tan x)\<^sup>2) x" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    82
    using cos_gt_zero_pi[of x] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    83
  show "LBINT x=-\<infinity>..\<infinity>. inverse (1 + x^2) = pi"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    84
    by (subst interval_integral_substitution_nonneg[of "-pi/2" "pi/2" tan "\<lambda>x. 1 + (tan x)^2"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    85
       (auto intro: derivative_eq_intros 1 2 filterlim_tan_at_right
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    86
             simp add: ereal_tendsto_simps filterlim_tan_at_left add_nonneg_eq_0_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    87
  show "set_integrable lborel (einterval (-\<infinity>) \<infinity>) (\<lambda>x. inverse (1 + x^2))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    88
    by (subst interval_integral_substitution_nonneg[of "-pi/2" "pi/2" tan "\<lambda>x. 1 + (tan x)^2"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    89
       (auto intro: derivative_eq_intros 1 2 filterlim_tan_at_right
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    90
             simp add: ereal_tendsto_simps filterlim_tan_at_left add_nonneg_eq_0_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    91
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    92
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
    93
lemma
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    94
  shows integrable_I0i_1_div_plus_square:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    95
    "interval_lebesgue_integrable lborel 0 \<infinity> (\<lambda>x. 1 / (1 + x^2))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    96
  and LBINT_I0i_1_div_plus_square:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    97
    "LBINT x=0..\<infinity>. 1 / (1 + x^2) = pi / 2"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    98
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
    99
  have 1: "0 < x \<Longrightarrow> x * 2 < pi \<Longrightarrow> (tan has_real_derivative 1 + (tan x)\<^sup>2) (at x)" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   100
    using cos_gt_zero_pi[of x] by (subst tan_sec) (auto intro!: DERIV_tan simp: power_inverse)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   101
  have 2: "0 < x \<Longrightarrow> x * 2 < pi \<Longrightarrow> isCont (\<lambda>x. 1 + (tan x)\<^sup>2) x" for x
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   102
    using cos_gt_zero_pi[of x] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   103
  show "LBINT x=0..\<infinity>. 1 / (1 + x^2) = pi / 2"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   104
    by (subst interval_integral_substitution_nonneg[of "0" "pi/2" tan "\<lambda>x. 1 + (tan x)^2"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   105
       (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   106
             simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   107
  show "interval_lebesgue_integrable lborel 0 \<infinity> (\<lambda>x. 1 / (1 + x^2))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   108
    unfolding interval_lebesgue_integrable_def
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   109
    by (subst interval_integral_substitution_nonneg[of "0" "pi/2" tan "\<lambda>x. 1 + (tan x)^2"])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   110
       (auto intro: derivative_eq_intros 1 2 tendsto_eq_intros
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   111
             simp add: ereal_tendsto_simps filterlim_tan_at_left zero_ereal_def add_nonneg_eq_0_iff)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   112
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   113
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   114
section \<open>The sinc function, and the sine integral (Si)\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   115
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   116
abbreviation sinc :: "real \<Rightarrow> real" where
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   117
  "sinc \<equiv> (\<lambda>x. if x = 0 then 1 else sin x / x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   118
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   119
lemma sinc_at_0: "((\<lambda>x. sin x / x::real) \<longlongrightarrow> 1) (at 0)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   120
  using DERIV_sin [of 0] by (auto simp add: has_field_derivative_def field_has_derivative_at)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   121
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   122
lemma isCont_sinc: "isCont sinc x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   123
proof cases
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   124
  assume "x = 0" then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   125
    using LIM_equal [where g = "\<lambda>x. sin x / x" and a=0 and f=sinc and l=1]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   126
    by (auto simp: isCont_def sinc_at_0)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   127
next
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   128
  assume "x \<noteq> 0" show ?thesis
62087
44841d07ef1d revisions to limits and derivatives, plus new lemmas
paulson
parents: 62083
diff changeset
   129
    by (rule continuous_transform_within [where d = "abs x" and f = "\<lambda>x. sin x / x"])
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   130
       (auto simp add: dist_real_def \<open>x \<noteq> 0\<close>)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   131
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   132
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   133
lemma continuous_on_sinc[continuous_intros]:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   134
  "continuous_on S f \<Longrightarrow> continuous_on S (\<lambda>x. sinc (f x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   135
  using continuous_on_compose[of S f sinc, OF _ continuous_at_imp_continuous_on]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   136
  by (auto simp: isCont_sinc)
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
lemma borel_measurable_sinc[measurable]: "sinc \<in> borel_measurable borel"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   139
  by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_sinc)
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
lemma sinc_AE: "AE x in lborel. sin x / x = sinc x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   142
  by (rule AE_I [where N = "{0}"], auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   143
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   144
definition Si :: "real \<Rightarrow> real" where "Si t \<equiv> LBINT x=0..t. sin x / x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   145
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   146
lemma sinc_neg [simp]: "sinc (- x) = sinc x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   147
  by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   148
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   149
(* TODO: Determine whether this can reasonably be eliminated by redefining Si. *)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   150
lemma Si_alt_def : "Si t = LBINT x=0..t. sinc x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   151
proof cases
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   152
  assume "0 \<le> t" then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   153
    using AE_lborel_singleton[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   154
    by (auto simp: Si_def intro!: interval_lebesgue_integral_cong_AE)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   155
next
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   156
  assume "\<not> 0 \<le> t" then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   157
    unfolding Si_def using AE_lborel_singleton[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   158
    by (subst (1 2) interval_integral_endpoints_reverse)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   159
       (auto simp: Si_def intro!: interval_lebesgue_integral_cong_AE)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   160
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   161
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   162
lemma Si_neg:
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   163
  assumes "T \<ge> 0" shows "Si (- T) = - Si T"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   164
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   165
  have "LBINT x=ereal 0..T. -1 *\<^sub>R sinc (- x) = LBINT y= ereal (- 0)..ereal (- T). sinc y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   166
    by (rule interval_integral_substitution_finite [OF assms])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   167
       (auto intro: derivative_intros continuous_at_imp_continuous_on isCont_sinc)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   168
  also have "(LBINT x=ereal 0..T. -1 *\<^sub>R sinc (- x)) = -(LBINT x=ereal 0..T. sinc x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   169
    by (subst sinc_neg) (simp_all add: interval_lebesgue_integral_uminus)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   170
  finally have *: "-(LBINT x=ereal 0..T. sinc x) = LBINT y= ereal 0..ereal (- T). sinc y"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   171
    by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   172
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   173
    using assms unfolding Si_alt_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   174
    by (subst zero_ereal_def)+ (auto simp add: * [symmetric])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   175
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   176
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   177
lemma integrable_sinc':
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   178
  "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\<lambda>t. sin (t * \<theta>) / t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   179
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   180
  have *: "interval_lebesgue_integrable lborel (ereal 0) (ereal T) (\<lambda>t. \<theta> * sinc (t * \<theta>))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   181
    by (intro interval_lebesgue_integrable_mult_right interval_integrable_isCont continuous_within_compose3 [OF isCont_sinc])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   182
       auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   183
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   184
    by (rule interval_lebesgue_integrable_cong_AE[THEN iffD1, OF _ _ _ *])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   185
       (insert AE_lborel_singleton[of 0], auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   186
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   187
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   188
(* TODO: need a better version of FTC2 *)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   189
lemma DERIV_Si: "(Si has_real_derivative sinc x) (at x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   190
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   191
  have "(at x within {min 0 (x - 1)..max 0 (x + 1)}) = at x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   192
    by (intro at_within_interior) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   193
  moreover have "min 0 (x - 1) \<le> x" "x \<le> max 0 (x + 1)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   194
    by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   195
  ultimately show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   196
    using interval_integral_FTC2[of "min 0 (x - 1)" 0 "max 0 (x + 1)" sinc x]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   197
    by (auto simp: continuous_at_imp_continuous_on isCont_sinc Si_alt_def[abs_def] zero_ereal_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   198
                   has_field_derivative_iff_has_vector_derivative
62390
842917225d56 more canonical names
nipkow
parents: 62087
diff changeset
   199
             split del: if_split)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   200
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   201
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   202
lemma isCont_Si: "isCont Si x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   203
  using DERIV_Si by (rule DERIV_isCont)
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
lemma borel_measurable_Si[measurable]: "Si \<in> borel_measurable borel"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   206
  by (auto intro: isCont_Si continuous_at_imp_continuous_on borel_measurable_continuous_on1)
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
lemma Si_at_top_LBINT:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   209
  "((\<lambda>t. (LBINT x=0..\<infinity>. exp (-(x * t)) * (x * sin t + cos t) / (1 + x^2))) \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   210
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   211
  let ?F = "\<lambda>x t. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2) :: real"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   212
  have int: "set_integrable lborel {0<..} (\<lambda>x. exp (- x) * (x + 1) :: real)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   213
    unfolding distrib_left
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   214
    using has_bochner_integral_I0i_power_exp_m[of 0] has_bochner_integral_I0i_power_exp_m[of 1]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   215
    by (intro set_integral_add) (auto dest!: integrable.intros simp: ac_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   216
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   217
  have "((\<lambda>t::real. LBINT x:{0<..}. ?F x t) \<longlongrightarrow> LBINT x::real:{0<..}. 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   218
  proof (rule integral_dominated_convergence_at_top[OF _ _ int], simp_all del: abs_divide split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   219
    have *: "0 < x \<Longrightarrow> \<bar>x * sin t + cos t\<bar> / (1 + x\<^sup>2) \<le> (x * 1 + 1) / 1" for x t :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   220
      by (intro frac_le abs_triangle_ineq[THEN order_trans] add_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   221
         (auto simp add: abs_mult simp del: mult_1_right intro!: mult_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   222
    then have **: "1 \<le> t \<Longrightarrow> 0 < x \<Longrightarrow> \<bar>?F x t\<bar> \<le> exp (- x) * (x + 1)" for x t :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   223
        by (auto simp add: abs_mult times_divide_eq_right[symmetric] simp del: times_divide_eq_right
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   224
                 intro!:  mult_mono)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   225
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   226
    show "\<forall>\<^sub>F i in at_top. AE x in lborel. 0 < x \<longrightarrow> \<bar>?F x i\<bar> \<le> exp (- x) * (x + 1)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   227
      using eventually_ge_at_top[of "1::real"] ** by (auto elim: eventually_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   228
    show "AE x in lborel. 0 < x \<longrightarrow> (?F x \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   229
    proof (intro always_eventually impI allI)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   230
      fix x :: real assume "0 < x"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   231
      show "((\<lambda>t. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2)) \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   232
      proof (rule Lim_null_comparison)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   233
        show "\<forall>\<^sub>F t in at_top. norm (?F x t) \<le> exp (- (x * t)) * (x + 1)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   234
          using eventually_ge_at_top[of "1::real"] * \<open>0 < x\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   235
          by (auto simp add: abs_mult times_divide_eq_right[symmetric] simp del: times_divide_eq_right
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   236
                   intro!: mult_mono elim: eventually_mono)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   237
        show "((\<lambda>t. exp (- (x * t)) * (x + 1)) \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   238
          by (auto simp: filterlim_uminus_at_top [symmetric]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   239
                   intro!: filterlim_tendsto_pos_mult_at_top[OF tendsto_const] \<open>0<x\<close> filterlim_ident
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   240
                           tendsto_mult_left_zero filterlim_compose[OF exp_at_bot])
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
    qed
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
  then show "((\<lambda>t. (LBINT x=0..\<infinity>. exp (-(x * t)) * (x * sin t + cos t) / (1 + x^2))) \<longlongrightarrow> 0) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   245
    by (simp add: interval_lebesgue_integral_0_infty)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   246
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   247
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   248
lemma Si_at_top_integrable:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   249
  assumes "t \<ge> 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   250
  shows "interval_lebesgue_integrable lborel 0 \<infinity> (\<lambda>x. exp (- (x * t)) * (x * sin t + cos t) / (1 + x\<^sup>2))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   251
  using \<open>0 \<le> t\<close> unfolding le_less
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   252
proof
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   253
  assume "0 = t" then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   254
    using integrable_I0i_1_div_plus_square by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   255
next
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   256
  assume [arith]: "0 < t"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   257
  have *: "0 \<le> a \<Longrightarrow> 0 < x \<Longrightarrow> a / (1 + x\<^sup>2) \<le> a" for a x :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   258
    using zero_le_power2[of x, arith] using divide_left_mono[of 1 "1+x\<^sup>2" a] by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   259
  have "set_integrable lborel {0<..} (\<lambda>x. (exp (- x) * x) * (sin t/t) + exp (- x) * cos t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   260
    using has_bochner_integral_I0i_power_exp_m[of 0] has_bochner_integral_I0i_power_exp_m[of 1]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   261
    by (intro set_integral_add set_integrable_mult_left)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   262
       (auto dest!: integrable.intros simp: ac_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   263
  from lborel_integrable_real_affine[OF this, of t 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   264
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   265
    unfolding interval_lebesgue_integral_0_infty
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63329
diff changeset
   266
    by (rule Bochner_Integration.integrable_bound) (auto simp: field_simps * split: split_indicator)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   267
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   268
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   269
lemma Si_at_top: "(Si \<longlongrightarrow> pi / 2) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   270
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   271
  have "\<forall>\<^sub>F t in at_top. pi / 2 - (LBINT u=0..\<infinity>. exp (-(u * t)) * (u * sin t + cos t) / (1 + u^2)) = Si t"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   272
    using eventually_ge_at_top[of 0]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   273
  proof eventually_elim
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   274
    fix t :: real assume "t \<ge> 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   275
    have "Si t = LBINT x=0..t. sin x * (LBINT u=0..\<infinity>. exp (-(u * x)))"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62975
diff changeset
   276
      unfolding Si_def using \<open>0 \<le> t\<close>
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   277
      by (intro interval_integral_discrete_difference[where X="{0}"]) (auto simp: LBINT_I0i_exp_mscale)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   278
    also have "\<dots> = LBINT x. (LBINT u=ereal 0..\<infinity>. indicator {0 <..< t} x *\<^sub>R sin x * exp (-(u * x)))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   279
      using \<open>0\<le>t\<close> by (simp add: zero_ereal_def interval_lebesgue_integral_le_eq mult_ac)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   280
    also have "\<dots> = LBINT x. (LBINT u. indicator ({0<..} \<times> {0 <..< t}) (u, x) *\<^sub>R (sin x * exp (-(u * x))))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   281
      by (subst interval_integral_Ioi) (simp_all add: indicator_times ac_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   282
    also have "\<dots> = LBINT u. (LBINT x. indicator ({0<..} \<times> {0 <..< t}) (u, x) *\<^sub>R (sin x * exp (-(u * x))))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   283
    proof (intro lborel_pair.Fubini_integral[symmetric] lborel_pair.Fubini_integrable)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   284
      show "(\<lambda>(x, y). indicator ({0<..} \<times> {0<..<t}) (y, x) *\<^sub>R (sin x * exp (- (y * x))))
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   285
          \<in> borel_measurable (lborel \<Otimes>\<^sub>M lborel)" (is "?f \<in> borel_measurable _")
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   286
        by measurable
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   287
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   288
      have "AE x in lborel. indicator {0..t} x *\<^sub>R abs (sinc x) = LBINT y. norm (?f (x, y))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   289
        using AE_lborel_singleton[of 0] AE_lborel_singleton[of t]
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   290
      proof eventually_elim
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   291
        fix x :: real assume x: "x \<noteq> 0" "x \<noteq> t"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   292
        have "LBINT y. \<bar>indicator ({0<..} \<times> {0<..<t}) (y, x) *\<^sub>R (sin x * exp (- (y * x)))\<bar> =
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   293
            LBINT y. \<bar>sin x\<bar> * exp (- (y * x)) * indicator {0<..} y * indicator {0<..<t} x"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63329
diff changeset
   294
          by (intro Bochner_Integration.integral_cong) (auto split: split_indicator simp: abs_mult)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   295
        also have "\<dots> = \<bar>sin x\<bar> * indicator {0<..<t} x * (LBINT y=0..\<infinity>.  exp (- (y * x)))"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   296
          by (cases "x > 0")
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   297
             (auto simp: field_simps interval_lebesgue_integral_0_infty split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   298
        also have "\<dots> = \<bar>sin x\<bar> * indicator {0<..<t} x * (1 / x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   299
          by (cases "x > 0") (auto simp add: LBINT_I0i_exp_mscale)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   300
        also have "\<dots> = indicator {0..t} x *\<^sub>R \<bar>sinc x\<bar>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   301
          using x by (simp add: field_simps split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   302
        finally show "indicator {0..t} x *\<^sub>R abs (sinc x) = LBINT y. norm (?f (x, y))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   303
          by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   304
      qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   305
      moreover have "set_integrable lborel {0 .. t} (\<lambda>x. abs (sinc x))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   306
        by (auto intro!: borel_integrable_compact continuous_intros simp del: real_scaleR_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   307
      ultimately show "integrable lborel (\<lambda>x. LBINT y. norm (?f (x, y)))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   308
        by (rule integrable_cong_AE_imp[rotated 2]) simp
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   309
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   310
      have "0 < x \<Longrightarrow> set_integrable lborel {0<..} (\<lambda>y. sin x * exp (- (y * x)))" for x :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   311
          by (intro set_integrable_mult_right integrable_I0i_exp_mscale)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   312
      then show "AE x in lborel. integrable lborel (\<lambda>y. ?f (x, y))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   313
        by (intro AE_I2) (auto simp: indicator_times split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   314
    qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   315
    also have "... = LBINT u=0..\<infinity>. (LBINT x=0..t. exp (-(u * x)) * sin x)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   316
      using \<open>0\<le>t\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   317
      by (auto simp: interval_lebesgue_integral_def zero_ereal_def ac_simps
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63329
diff changeset
   318
               split: split_indicator intro!: Bochner_Integration.integral_cong)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   319
    also have "\<dots> = LBINT u=0..\<infinity>. 1 / (1 + u\<^sup>2) - 1 / (1 + u\<^sup>2) * (exp (- (u * t)) * (u * sin t + cos t))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   320
      by (auto simp: divide_simps LBINT_I0c_exp_mscale_sin intro!: interval_integral_cong)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   321
    also have "... = pi / 2 - (LBINT u=0..\<infinity>. exp (- (u * t)) * (u * sin t + cos t) / (1 + u^2))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   322
      using Si_at_top_integrable[OF \<open>0\<le>t\<close>] by (simp add: integrable_I0i_1_div_plus_square LBINT_I0i_1_div_plus_square)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   323
    finally show "pi / 2 - (LBINT u=0..\<infinity>. exp (-(u * t)) * (u * sin t + cos t) / (1 + u^2)) = Si t" ..
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   324
  qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   325
  then show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   326
    by (rule Lim_transform_eventually)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   327
       (auto intro!: tendsto_eq_intros Si_at_top_LBINT)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   328
qed
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
subsection \<open>The final theorems: boundedness and scalability\<close>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   331
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   332
lemma bounded_Si: "\<exists>B. \<forall>T. \<bar>Si T\<bar> \<le> B"
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 *: "0 \<le> y \<Longrightarrow> dist x y < z \<Longrightarrow> abs x \<le> y + z" for x y z :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   335
    by (simp add: dist_real_def)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   336
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   337
  have "eventually (\<lambda>T. dist (Si T) (pi / 2) < 1) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   338
    using Si_at_top by (elim tendstoD) simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   339
  then have "eventually (\<lambda>T. 0 \<le> T \<and> \<bar>Si T\<bar> \<le> pi / 2 + 1) at_top"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   340
    using eventually_ge_at_top[of 0] by eventually_elim (insert *[of "pi/2" "Si _" 1], auto)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   341
  then have "\<exists>T. 0 \<le> T \<and> (\<forall>t \<ge> T. \<bar>Si t\<bar> \<le> pi / 2 + 1)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   342
    by (auto simp add: eventually_at_top_linorder)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   343
  then obtain T where T: "0 \<le> T" "\<And>t. t \<ge> T \<Longrightarrow> \<bar>Si t\<bar> \<le> pi / 2 + 1"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   344
    by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   345
  moreover have "t \<le> - T \<Longrightarrow> \<bar>Si t\<bar> \<le> pi / 2 + 1" for t
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   346
    using T(1) T(2)[of "-t"] Si_neg[of "- t"] by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   347
  moreover have "\<exists>M. \<forall>t. -T \<le> t \<and> t \<le> T \<longrightarrow> \<bar>Si t\<bar> \<le> M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   348
    by (rule isCont_bounded) (auto intro!: isCont_Si continuous_intros \<open>0 \<le> T\<close>)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   349
  then obtain M where M: "\<And>t. -T \<le> t \<and> t \<le> T \<Longrightarrow> \<bar>Si t\<bar> \<le> M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   350
    by auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   351
  ultimately show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   352
    by (intro exI[of _ "max M (pi/2 + 1)"]) (meson le_max_iff_disj linorder_not_le order_le_less)
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   353
qed
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   354
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   355
lemma LBINT_I0c_sin_scale_divide:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   356
  assumes "T \<ge> 0"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   357
  shows "LBINT t=0..T. sin (t * \<theta>) / t = sgn \<theta> * Si (T * \<bar>\<theta>\<bar>)"
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
  { assume "0 < \<theta>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   360
    have "(LBINT t=ereal 0..T. sin (t * \<theta>) / t) = (LBINT t=ereal 0..T. \<theta> *\<^sub>R sinc (t * \<theta>))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   361
      by (rule interval_integral_discrete_difference[of "{0}"]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   362
    also have "\<dots> = (LBINT t=ereal (0 * \<theta>)..T * \<theta>. sinc t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   363
      apply (rule interval_integral_substitution_finite [OF assms])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   364
      apply (subst mult.commute, rule DERIV_subset)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   365
      by (auto intro!: derivative_intros continuous_at_imp_continuous_on isCont_sinc)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   366
    also have "\<dots> = (LBINT t=ereal (0 * \<theta>)..T * \<theta>. sin t / t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   367
      by (rule interval_integral_discrete_difference[of "{0}"]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   368
    finally have "(LBINT t=ereal 0..T. sin (t * \<theta>) / t) = (LBINT t=ereal 0..T * \<theta>. sin t / t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   369
      by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   370
    hence "LBINT x. indicator {0<..<T} x * sin (x * \<theta>) / x =
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   371
        LBINT x. indicator {0<..<T * \<theta>} x * sin x / x"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62975
diff changeset
   372
      using assms \<open>0 < \<theta>\<close> unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   373
        by (auto simp: ac_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   374
  } note aux1 = this
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   375
  { assume "0 > \<theta>"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   376
    have [simp]: "integrable lborel (\<lambda>x. sin (x * \<theta>) * indicator {0<..<T} x / x)"
62975
1d066f6ab25d Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents: 62390
diff changeset
   377
      using integrable_sinc' [of T \<theta>] assms
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   378
      by (simp add: interval_lebesgue_integrable_def ac_simps)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   379
    have "(LBINT t=ereal 0..T. sin (t * -\<theta>) / t) = (LBINT t=ereal 0..T. -\<theta> *\<^sub>R sinc (t * -\<theta>))"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   380
      by (rule interval_integral_discrete_difference[of "{0}"]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   381
    also have "\<dots> = (LBINT t=ereal (0 * -\<theta>)..T * -\<theta>. sinc t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   382
      apply (rule interval_integral_substitution_finite [OF assms])
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   383
      apply (subst mult.commute, rule DERIV_subset)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   384
      by (auto intro!: derivative_intros continuous_at_imp_continuous_on isCont_sinc)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   385
    also have "\<dots> = (LBINT t=ereal (0 * -\<theta>)..T * -\<theta>. sin t / t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   386
      by (rule interval_integral_discrete_difference[of "{0}"]) auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   387
    finally have "(LBINT t=ereal 0..T. sin (t * -\<theta>) / t) = (LBINT t=ereal 0..T * -\<theta>. sin t / t)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   388
      by simp
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   389
    hence "LBINT x. indicator {0<..<T} x * sin (x * \<theta>) / x =
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   390
       - (LBINT x. indicator {0<..<- (T * \<theta>)} x * sin x / x)"
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62975
diff changeset
   391
      using assms \<open>0 > \<theta>\<close> unfolding interval_lebesgue_integral_def einterval_eq zero_ereal_def
62390
842917225d56 more canonical names
nipkow
parents: 62087
diff changeset
   392
        by (auto simp add: field_simps mult_le_0_iff split: if_split_asm)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   393
  } note aux2 = this
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   394
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   395
    using assms unfolding Si_def interval_lebesgue_integral_def sgn_real_def einterval_eq zero_ereal_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   396
    apply auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   397
    apply (erule aux1)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   398
    apply (rule aux2)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   399
    apply auto
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   400
    done
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   401
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   402
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents:
diff changeset
   403
end