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