| author | wenzelm | 
| Sat, 20 Feb 2021 23:01:35 +0100 | |
| changeset 73264 | 440546ea20e6 | 
| parent 70817 | dd675800469d | 
| child 75455 | 91c16c5ad3e9 | 
| permissions | -rw-r--r-- | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
1  | 
(* Title: HOL/Probability/Distributions.thy  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
2  | 
Author: Sudeep Kanav, TU München  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
3  | 
Author: Johannes Hölzl, TU München  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
4  | 
Author: Jeremy Avigad, CMU *)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
5  | 
|
| 61808 | 6  | 
section \<open>Properties of Various Distributions\<close>  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
7  | 
|
| 50419 | 8  | 
theory Distributions  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
9  | 
imports Convolution Information  | 
| 50419 | 10  | 
begin  | 
11  | 
||
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
12  | 
lemma (in prob_space) distributed_affine:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
13  | 
fixes f :: "real \<Rightarrow> ennreal"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
14  | 
assumes f: "distributed M lborel X f"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
15  | 
assumes c: "c \<noteq> 0"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
16  | 
shows "distributed M lborel (\<lambda>x. t + c * X x) (\<lambda>x. f ((x - t) / c) / \<bar>c\<bar>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
17  | 
unfolding distributed_def  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
18  | 
proof safe  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
19  | 
have [measurable]: "f \<in> borel_measurable borel"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
20  | 
using f by (simp add: distributed_def)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
21  | 
have [measurable]: "X \<in> borel_measurable M"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
22  | 
using f by (simp add: distributed_def)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
23  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
24  | 
show "(\<lambda>x. f ((x - t) / c) / \<bar>c\<bar>) \<in> borel_measurable lborel"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
25  | 
by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
26  | 
show "random_variable lborel (\<lambda>x. t + c * X x)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
27  | 
by simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
28  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
29  | 
have eq: "ennreal \<bar>c\<bar> * (f x / ennreal \<bar>c\<bar>) = f x" for x  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
30  | 
using c  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
31  | 
by (cases "f x")  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
32  | 
(auto simp: divide_ennreal ennreal_mult[symmetric] ennreal_top_divide ennreal_mult_top)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
33  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
34  | 
have "density lborel f = distr M lborel X"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
35  | 
using f by (simp add: distributed_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
36  | 
with c show "distr M lborel (\<lambda>x. t + c * X x) = density lborel (\<lambda>x. f ((x - t) / c) / ennreal \<bar>c\<bar>)"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
37  | 
by (subst (2) lborel_real_affine[where c="c" and t="t"])  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
38  | 
(simp_all add: density_density_eq density_distr distr_distr field_simps eq cong: distr_cong)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
39  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
40  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
41  | 
lemma (in prob_space) distributed_affineI:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
42  | 
fixes f :: "real \<Rightarrow> ennreal" and c :: real  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
43  | 
assumes f: "distributed M lborel (\<lambda>x. (X x - t) / c) (\<lambda>x. \<bar>c\<bar> * f (x * c + t))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
44  | 
assumes c: "c \<noteq> 0"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
45  | 
shows "distributed M lborel X f"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
46  | 
proof -  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
47  | 
have eq: "f x * ennreal \<bar>c\<bar> / ennreal \<bar>c\<bar> = f x" for x  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
48  | 
using c by (simp add: ennreal_times_divide[symmetric])  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
49  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
50  | 
show ?thesis  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
51  | 
using distributed_affine[OF f c, where t=t] c  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
52  | 
by (simp add: field_simps eq)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
53  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
54  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
55  | 
lemma (in prob_space) distributed_AE2:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
56  | 
assumes [measurable]: "distributed M N X f" "Measurable.pred N P"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
57  | 
shows "(AE x in M. P (X x)) \<longleftrightarrow> (AE x in N. 0 < f x \<longrightarrow> P x)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
58  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
59  | 
have "(AE x in M. P (X x)) \<longleftrightarrow> (AE x in distr M N X. P x)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
60  | 
by (simp add: AE_distr_iff)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
61  | 
also have "\<dots> \<longleftrightarrow> (AE x in density N f. P x)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
62  | 
unfolding distributed_distr_eq_density[OF assms(1)] ..  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
63  | 
also have "\<dots> \<longleftrightarrow> (AE x in N. 0 < f x \<longrightarrow> P x)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
64  | 
by (rule AE_density) simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
65  | 
finally show ?thesis .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
66  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
67  | 
|
| 61808 | 68  | 
subsection \<open>Erlang\<close>  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
69  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
70  | 
lemma nn_intergal_power_times_exp_Icc:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
71  | 
assumes [arith]: "0 \<le> a"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
72  | 
  shows "(\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x \<partial>lborel) =
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
73  | 
(1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n)) * fact k" (is "?I = _")  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
74  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
75  | 
let ?f = "\<lambda>k x. x^k * exp (-x) / fact k"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
76  | 
let ?F = "\<lambda>k x. - (\<Sum>n\<le>k. (x^n * exp (-x)) / fact n)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
77  | 
have "?I * (inverse (real_of_nat (fact k))) =  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
78  | 
      (\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 .. a} x * (inverse (real_of_nat (fact k))) \<partial>lborel)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
79  | 
by (intro nn_integral_multc[symmetric]) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
80  | 
  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (?f k x) * indicator {0 .. a} x \<partial>lborel)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
81  | 
by (intro nn_integral_cong)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
82  | 
(simp add: field_simps ennreal_mult'[symmetric] indicator_mult_ennreal)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
83  | 
also have "\<dots> = ennreal (?F k a - ?F k 0)"  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
84  | 
proof (rule nn_integral_FTC_Icc)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
85  | 
    fix x assume "x \<in> {0..a}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
86  | 
show "DERIV (?F k) x :> ?f k x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
87  | 
proof(induction k)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
88  | 
case 0 show ?case by (auto intro!: derivative_eq_intros)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
89  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
90  | 
case (Suc k)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
91  | 
have "DERIV (\<lambda>x. ?F k x - (x^Suc k * exp (-x)) / fact (Suc k)) x :>  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
92  | 
?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k))"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
93  | 
by (intro DERIV_diff Suc)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
94  | 
(auto intro!: derivative_eq_intros simp del: fact_Suc power_Suc  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
95  | 
simp add: field_simps power_Suc[symmetric])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
96  | 
also have "(\<lambda>x. ?F k x - (x^Suc k * exp (-x)) / fact (Suc k)) = ?F (Suc k)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
97  | 
by simp  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
98  | 
also have "?f k x - ((real (Suc k) - x) * x ^ k * exp (- x)) / (fact (Suc k)) = ?f (Suc k) x"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
99  | 
by (auto simp: field_simps simp del: fact_Suc)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
100  | 
(simp_all add: of_nat_Suc field_simps)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
101  | 
finally show ?case .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
102  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
103  | 
qed auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
104  | 
also have "\<dots> = ennreal (1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n))"  | 
| 64267 | 105  | 
by (auto simp: power_0_left if_distrib[where f="\<lambda>x. x / a" for a] sum.If_cases)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
106  | 
also have "\<dots> = ennreal ((1 - (\<Sum>n\<le>k. (a^n * exp (-a)) / fact n)) * fact k) * ennreal (inverse (fact k))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
107  | 
by (subst ennreal_mult''[symmetric]) (auto intro!: arg_cong[where f=ennreal])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
108  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
109  | 
by (auto simp add: mult_right_ennreal_cancel le_less)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
110  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
111  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
112  | 
lemma nn_intergal_power_times_exp_Ici:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
113  | 
  shows "(\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel) = real_of_nat (fact k)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
114  | 
proof (rule LIMSEQ_unique)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
115  | 
  let ?X = "\<lambda>n. \<integral>\<^sup>+ x. ennreal (x^k * exp (-x)) * indicator {0 .. real n} x \<partial>lborel"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
116  | 
  show "?X \<longlonglongrightarrow> (\<integral>\<^sup>+x. ennreal (x^k * exp (-x)) * indicator {0 ..} x \<partial>lborel)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
117  | 
apply (intro nn_integral_LIMSEQ)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
118  | 
apply (auto simp: incseq_def le_fun_def eventually_sequentially  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
67977 
diff
changeset
 | 
119  | 
split: split_indicator intro!: tendsto_eventually)  | 
| 
59587
 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 
nipkow 
parents: 
59000 
diff
changeset
 | 
120  | 
apply (metis nat_ceiling_le_eq)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
121  | 
done  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
122  | 
|
| 61973 | 123  | 
have "((\<lambda>x::real. (1 - (\<Sum>n\<le>k. (x ^ n / exp x) / (fact n))) * fact k) \<longlongrightarrow>  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
124  | 
(1 - (\<Sum>n\<le>k. 0 / (fact n))) * fact k) at_top"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
125  | 
by (intro tendsto_intros tendsto_power_div_exp_0) simp  | 
| 61969 | 126  | 
then show "?X \<longlonglongrightarrow> real_of_nat (fact k)"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
127  | 
by (subst nn_intergal_power_times_exp_Icc)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
128  | 
(auto simp: exp_minus field_simps intro!: filterlim_compose[OF _ filterlim_real_sequentially])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
129  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
130  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
131  | 
definition erlang_density :: "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
132  | 
"erlang_density k l x = (if x < 0 then 0 else (l^(Suc k) * x^k * exp (- l * x)) / fact k)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
133  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
134  | 
definition erlang_CDF :: "nat \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
135  | 
"erlang_CDF k l x = (if x < 0 then 0 else 1 - (\<Sum>n\<le>k. ((l * x)^n * exp (- l * x) / fact n)))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
136  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
137  | 
lemma erlang_density_nonneg[simp]: "0 \<le> l \<Longrightarrow> 0 \<le> erlang_density k l x"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
138  | 
by (simp add: erlang_density_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
139  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
140  | 
lemma borel_measurable_erlang_density[measurable]: "erlang_density k l \<in> borel_measurable borel"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
141  | 
by (auto simp add: erlang_density_def[abs_def])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
142  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
143  | 
lemma erlang_CDF_transform: "0 < l \<Longrightarrow> erlang_CDF k l a = erlang_CDF k 1 (l * a)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
144  | 
by (auto simp add: erlang_CDF_def mult_less_0_iff)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
145  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
146  | 
lemma erlang_CDF_nonneg[simp]: assumes "0 < l" shows "0 \<le> erlang_CDF k l x"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
147  | 
unfolding erlang_CDF_def  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
148  | 
proof (clarsimp simp: not_less)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
149  | 
assume "0 \<le> x"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
150  | 
have "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
151  | 
exp (- (l * x)) * (\<Sum>n\<le>k. (l * x) ^ n / fact n)"  | 
| 64267 | 152  | 
unfolding sum_distrib_left by (intro sum.cong) (auto simp: field_simps)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
153  | 
also have "\<dots> = (\<Sum>n\<le>k. (l * x) ^ n / fact n) / exp (l * x)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
154  | 
by (simp add: exp_minus field_simps)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
155  | 
also have "\<dots> \<le> 1"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
156  | 
proof (subst divide_le_eq_1_pos)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
157  | 
show "(\<Sum>n\<le>k. (l * x) ^ n / fact n) \<le> exp (l * x)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
158  | 
using \<open>0 < l\<close> \<open>0 \<le> x\<close> summable_exp_generic[of "l * x"]  | 
| 64267 | 159  | 
by (auto simp: exp_def divide_inverse ac_simps intro!: sum_le_suminf)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
160  | 
qed simp  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
161  | 
finally show "(\<Sum>n\<le>k. (l * x) ^ n * exp (- (l * x)) / fact n) \<le> 1" .  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
162  | 
qed  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
163  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
164  | 
lemma nn_integral_erlang_density:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
165  | 
assumes [arith]: "0 < l"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
166  | 
  shows "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = erlang_CDF k l a"
 | 
| 63540 | 167  | 
proof (cases "0 \<le> a")  | 
168  | 
case [arith]: True  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
169  | 
  have eq: "\<And>x. indicator {0..a} (x / l) = indicator {0..a*l} x"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
170  | 
by (simp add: field_simps split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
171  | 
  have "(\<integral>\<^sup>+x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) =
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
172  | 
    (\<integral>\<^sup>+x. (l/fact k) * (ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x) \<partial>lborel)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
173  | 
by (intro nn_integral_cong)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
174  | 
(auto simp: erlang_density_def power_mult_distrib ennreal_mult[symmetric] split: split_indicator)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
175  | 
  also have "\<dots> = (l/fact k) * (\<integral>\<^sup>+x. ennreal ((l*x)^k * exp (- (l*x))) * indicator {0 .. a} x \<partial>lborel)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
176  | 
by (intro nn_integral_cmult) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
177  | 
  also have "\<dots> = ennreal (l/fact k) * ((1/l) * (\<integral>\<^sup>+x. ennreal (x^k * exp (- x)) * indicator {0 .. l * a} x \<partial>lborel))"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
178  | 
by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
179  | 
also have "\<dots> = (1 - (\<Sum>n\<le>k. ((l * a)^n * exp (-(l * a))) / fact n))"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
180  | 
by (subst nn_intergal_power_times_exp_Icc) (auto simp: ennreal_mult'[symmetric])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
181  | 
also have "\<dots> = erlang_CDF k l a"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
182  | 
by (auto simp: erlang_CDF_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
183  | 
finally show ?thesis .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
184  | 
next  | 
| 63540 | 185  | 
case False  | 
186  | 
  then have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x) * indicator {.. a} x \<partial>lborel) = (\<integral>\<^sup>+x. 0 \<partial>(lborel::real measure))"
 | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
187  | 
by (intro nn_integral_cong) (auto simp: erlang_density_def)  | 
| 63540 | 188  | 
with False show ?thesis  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
189  | 
by (simp add: erlang_CDF_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
190  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
191  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
192  | 
lemma emeasure_erlang_density:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
193  | 
  "0 < l \<Longrightarrow> emeasure (density lborel (erlang_density k l)) {.. a} = erlang_CDF k l a"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
194  | 
by (simp add: emeasure_density nn_integral_erlang_density)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
195  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
196  | 
lemma nn_integral_erlang_ith_moment:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
197  | 
fixes k i :: nat and l :: real  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
198  | 
assumes [arith]: "0 < l"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
199  | 
shows "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x * x ^ i) \<partial>lborel) = fact (k + i) / (fact k * l ^ i)"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
200  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
201  | 
  have eq: "\<And>x. indicator {0..} (x / l) = indicator {0..} x"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
202  | 
by (simp add: field_simps split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
203  | 
have "(\<integral>\<^sup>+ x. ennreal (erlang_density k l x * x^i) \<partial>lborel) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
204  | 
    (\<integral>\<^sup>+x. (l/(fact k * l^i)) * (ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x) \<partial>lborel)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
205  | 
by (intro nn_integral_cong)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
206  | 
(auto simp: erlang_density_def power_mult_distrib power_add ennreal_mult'[symmetric] split: split_indicator)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
207  | 
  also have "\<dots> = (l/(fact k * l^i)) * (\<integral>\<^sup>+x. ennreal ((l*x)^(k+i) * exp (- (l*x))) * indicator {0 ..} x \<partial>lborel)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
208  | 
by (intro nn_integral_cmult) auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
209  | 
  also have "\<dots> = ennreal (l/(fact k * l^i)) * ((1/l) * (\<integral>\<^sup>+x. ennreal (x^(k+i) * exp (- x)) * indicator {0 ..} x \<partial>lborel))"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
210  | 
by (subst nn_integral_real_affine[where c="1 / l" and t=0]) (auto simp: field_simps eq)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
211  | 
also have "\<dots> = fact (k + i) / (fact k * l ^ i)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
212  | 
by (subst nn_intergal_power_times_exp_Ici) (auto simp: ennreal_mult'[symmetric])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
213  | 
finally show ?thesis .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
214  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
215  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
216  | 
lemma prob_space_erlang_density:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
217  | 
assumes l[arith]: "0 < l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
218  | 
shows "prob_space (density lborel (erlang_density k l))" (is "prob_space ?D")  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
219  | 
proof  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
220  | 
show "emeasure ?D (space ?D) = 1"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
221  | 
using nn_integral_erlang_ith_moment[OF l, where k=k and i=0] by (simp add: emeasure_density)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
222  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
223  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
224  | 
lemma (in prob_space) erlang_distributed_le:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
225  | 
assumes D: "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
226  | 
assumes [simp, arith]: "0 < l" "0 \<le> a"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
227  | 
shows "\<P>(x in M. X x \<le> a) = erlang_CDF k l a"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
228  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
229  | 
  have "emeasure M {x \<in> space M. X x \<le> a } = emeasure (distr M lborel X) {.. a}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
230  | 
using distributed_measurable[OF D]  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
231  | 
by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
232  | 
  also have "\<dots> = emeasure (density lborel (erlang_density k l)) {.. a}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
233  | 
unfolding distributed_distr_eq_density[OF D] ..  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
234  | 
also have "\<dots> = erlang_CDF k l a"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
235  | 
by (auto intro!: emeasure_erlang_density)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
236  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
237  | 
by (auto simp: emeasure_eq_measure measure_nonneg)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
238  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
239  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
240  | 
lemma (in prob_space) erlang_distributed_gt:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
241  | 
assumes D[simp]: "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
242  | 
assumes [arith]: "0 < l" "0 \<le> a"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
243  | 
shows "\<P>(x in M. a < X x ) = 1 - (erlang_CDF k l a)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
244  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
245  | 
have " 1 - (erlang_CDF k l a) = 1 - \<P>(x in M. X x \<le> a)" by (subst erlang_distributed_le) auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
246  | 
  also have "\<dots> = prob (space M - {x \<in> space M. X x \<le> a })"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
247  | 
using distributed_measurable[OF D] by (auto simp: prob_compl)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
248  | 
also have "\<dots> = \<P>(x in M. a < X x )" by (auto intro!: arg_cong[where f=prob] simp: not_le)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
249  | 
finally show ?thesis by simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
250  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
251  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
252  | 
lemma erlang_CDF_at0: "erlang_CDF k l 0 = 0"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
253  | 
by (induction k) (auto simp: erlang_CDF_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
254  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
255  | 
lemma erlang_distributedI:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
256  | 
assumes X[measurable]: "X \<in> borel_measurable M" and [arith]: "0 < l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
257  | 
    and X_distr: "\<And>a. 0 \<le> a \<Longrightarrow> emeasure M {x\<in>space M. X x \<le> a} = erlang_CDF k l a"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
258  | 
shows "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
259  | 
proof (rule distributedI_borel_atMost)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
260  | 
fix a :: real  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
261  | 
  { assume "a \<le> 0"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
262  | 
    with X have "emeasure M {x\<in>space M. X x \<le> a} \<le> emeasure M {x\<in>space M. X x \<le> 0}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
263  | 
by (intro emeasure_mono) auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
264  | 
also have "... = 0" by (auto intro!: erlang_CDF_at0 simp: X_distr[of 0])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
265  | 
    finally have "emeasure M {x\<in>space M. X x \<le> a} \<le> 0" by simp
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
266  | 
    then have "emeasure M {x\<in>space M. X x \<le> a} = 0" by simp
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
267  | 
}  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
268  | 
note eq_0 = this  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
269  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
270  | 
  show "(\<integral>\<^sup>+ x. erlang_density k l x * indicator {..a} x \<partial>lborel) = ennreal (erlang_CDF k l a)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
271  | 
using nn_integral_erlang_density[of l k a]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
272  | 
by (simp add: ennreal_indicator ennreal_mult)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
273  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
274  | 
  show "emeasure M {x\<in>space M. X x \<le> a} = ennreal (erlang_CDF k l a)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
275  | 
using X_distr[of a] eq_0 by (auto simp: one_ennreal_def erlang_CDF_def)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
276  | 
qed simp_all  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
277  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
278  | 
lemma (in prob_space) erlang_distributed_iff:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
279  | 
assumes [arith]: "0<l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
280  | 
shows "distributed M lborel X (erlang_density k l) \<longleftrightarrow>  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
281  | 
(X \<in> borel_measurable M \<and> 0 < l \<and> (\<forall>a\<ge>0. \<P>(x in M. X x \<le> a) = erlang_CDF k l a ))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
282  | 
using  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
283  | 
distributed_measurable[of M lborel X "erlang_density k l"]  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
284  | 
emeasure_erlang_density[of l]  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
285  | 
erlang_distributed_le[of X k l]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
286  | 
by (auto intro!: erlang_distributedI simp: one_ennreal_def emeasure_eq_measure)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
287  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
288  | 
lemma (in prob_space) erlang_distributed_mult_const:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
289  | 
assumes erlX: "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
290  | 
assumes a_pos[arith]: "0 < \<alpha>" "0 < l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
291  | 
shows "distributed M lborel (\<lambda>x. \<alpha> * X x) (erlang_density k (l / \<alpha>))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
292  | 
proof (subst erlang_distributed_iff, safe)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
293  | 
have [measurable]: "random_variable borel X" and [arith]: "0 < l "  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
294  | 
  and  [simp]: "\<And>a. 0 \<le> a \<Longrightarrow> prob {x \<in> space M. X x \<le> a} = erlang_CDF k l a"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
295  | 
by(insert erlX, auto simp: erlang_distributed_iff)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
296  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
297  | 
show "random_variable borel (\<lambda>x. \<alpha> * X x)" "0 < l / \<alpha>" "0 < l / \<alpha>"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
298  | 
by (auto simp:field_simps)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
299  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
300  | 
fix a:: real assume [arith]: "0 \<le> a"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
301  | 
obtain b:: real where [simp, arith]: "b = a/ \<alpha>" by blast  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
302  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
303  | 
have [arith]: "0 \<le> b" by (auto simp: divide_nonneg_pos)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
304  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
305  | 
  have "prob {x \<in> space M. \<alpha> * X x \<le> a}  = prob {x \<in> space M.  X x \<le> b}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
306  | 
by (rule arg_cong[where f= prob]) (auto simp:field_simps)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
307  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
308  | 
  moreover have "prob {x \<in> space M. X x \<le> b} = erlang_CDF k l b" by auto
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
309  | 
moreover have "erlang_CDF k (l / \<alpha>) a = erlang_CDF k l b" unfolding erlang_CDF_def by auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
310  | 
  ultimately show "prob {x \<in> space M. \<alpha> * X x \<le> a} = erlang_CDF k (l / \<alpha>) a" by fastforce
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
311  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
312  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
313  | 
lemma (in prob_space) has_bochner_integral_erlang_ith_moment:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
314  | 
fixes k i :: nat and l :: real  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
315  | 
assumes [arith]: "0 < l" and D: "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
316  | 
shows "has_bochner_integral M (\<lambda>x. X x ^ i) (fact (k + i) / (fact k * l ^ i))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
317  | 
proof (rule has_bochner_integral_nn_integral)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
318  | 
show "AE x in M. 0 \<le> X x ^ i"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
319  | 
by (subst distributed_AE2[OF D]) (auto simp: erlang_density_def)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
320  | 
show "(\<integral>\<^sup>+ x. ennreal (X x ^ i) \<partial>M) = ennreal (fact (k + i) / (fact k * l ^ i))"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
321  | 
using nn_integral_erlang_ith_moment[of l k i]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
322  | 
by (subst distributed_nn_integral[symmetric, OF D]) (auto simp: ennreal_mult')  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
323  | 
qed (insert distributed_measurable[OF D], auto)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
324  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
325  | 
lemma (in prob_space) erlang_ith_moment_integrable:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
326  | 
"0 < l \<Longrightarrow> distributed M lborel X (erlang_density k l) \<Longrightarrow> integrable M (\<lambda>x. X x ^ i)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
327  | 
by rule (rule has_bochner_integral_erlang_ith_moment)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
328  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
329  | 
lemma (in prob_space) erlang_ith_moment:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
330  | 
"0 < l \<Longrightarrow> distributed M lborel X (erlang_density k l) \<Longrightarrow>  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
331  | 
expectation (\<lambda>x. X x ^ i) = fact (k + i) / (fact k * l ^ i)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
332  | 
by (rule has_bochner_integral_integral_eq) (rule has_bochner_integral_erlang_ith_moment)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
333  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
334  | 
lemma (in prob_space) erlang_distributed_variance:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
335  | 
assumes [arith]: "0 < l" and "distributed M lborel X (erlang_density k l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
336  | 
shows "variance X = (k + 1) / l\<^sup>2"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
337  | 
proof (subst variance_eq)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
338  | 
show "integrable M X" "integrable M (\<lambda>x. (X x)\<^sup>2)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
339  | 
using erlang_ith_moment_integrable[OF assms, of 1] erlang_ith_moment_integrable[OF assms, of 2]  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
340  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
341  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
342  | 
show "expectation (\<lambda>x. (X x)\<^sup>2) - (expectation X)\<^sup>2 = real (k + 1) / l\<^sup>2"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
343  | 
using erlang_ith_moment[OF assms, of 1] erlang_ith_moment[OF assms, of 2]  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
344  | 
by simp (auto simp: power2_eq_square field_simps of_nat_Suc)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
345  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
346  | 
|
| 61808 | 347  | 
subsection \<open>Exponential distribution\<close>  | 
| 50419 | 348  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
349  | 
abbreviation exponential_density :: "real \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
350  | 
"exponential_density \<equiv> erlang_density 0"  | 
| 50419 | 351  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
352  | 
lemma exponential_density_def:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
353  | 
"exponential_density l x = (if x < 0 then 0 else l * exp (- x * l))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
354  | 
by (simp add: fun_eq_iff erlang_density_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
355  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
356  | 
lemma erlang_CDF_0: "erlang_CDF 0 l a = (if 0 \<le> a then 1 - exp (- l * a) else 0)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
357  | 
by (simp add: erlang_CDF_def)  | 
| 50419 | 358  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
359  | 
lemma prob_space_exponential_density: "0 < l \<Longrightarrow> prob_space (density lborel (exponential_density l))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
360  | 
by (rule prob_space_erlang_density)  | 
| 50419 | 361  | 
|
362  | 
lemma (in prob_space) exponential_distributedD_le:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
363  | 
assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l"  | 
| 50419 | 364  | 
shows "\<P>(x in M. X x \<le> a) = 1 - exp (- a * l)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
365  | 
using erlang_distributed_le[OF D l a] a by (simp add: erlang_CDF_def)  | 
| 50419 | 366  | 
|
367  | 
lemma (in prob_space) exponential_distributedD_gt:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
368  | 
assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l"  | 
| 50419 | 369  | 
shows "\<P>(x in M. a < X x ) = exp (- a * l)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
370  | 
using erlang_distributed_gt[OF D l a] a by (simp add: erlang_CDF_def)  | 
| 50419 | 371  | 
|
372  | 
lemma (in prob_space) exponential_distributed_memoryless:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
373  | 
assumes D: "distributed M lborel X (exponential_density l)" and a: "0 \<le> a" and l: "0 < l" and t: "0 \<le> t"  | 
| 50419 | 374  | 
shows "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. t < X x)"  | 
375  | 
proof -  | 
|
376  | 
have "\<P>(x in M. a + t < X x \<bar> a < X x) = \<P>(x in M. a + t < X x) / \<P>(x in M. a < X x)"  | 
|
| 67399 | 377  | 
using \<open>0 \<le> t\<close> by (auto simp: cond_prob_def intro!: arg_cong[where f=prob] arg_cong2[where f="(/)"])  | 
| 50419 | 378  | 
also have "\<dots> = exp (- (a + t) * l) / exp (- a * l)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
379  | 
using a t by (simp add: exponential_distributedD_gt[OF D _ l])  | 
| 50419 | 380  | 
also have "\<dots> = exp (- t * l)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
381  | 
using l by (auto simp: field_simps exp_add[symmetric])  | 
| 50419 | 382  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
383  | 
using t by (simp add: exponential_distributedD_gt[OF D _ l])  | 
| 50419 | 384  | 
qed  | 
385  | 
||
386  | 
lemma exponential_distributedI:  | 
|
387  | 
assumes X[measurable]: "X \<in> borel_measurable M" and [arith]: "0 < l"  | 
|
388  | 
    and X_distr: "\<And>a. 0 \<le> a \<Longrightarrow> emeasure M {x\<in>space M. X x \<le> a} = 1 - exp (- a * l)"
 | 
|
389  | 
shows "distributed M lborel X (exponential_density l)"  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
390  | 
proof (rule erlang_distributedI)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
391  | 
  fix a :: real assume "0 \<le> a" then show "emeasure M {x \<in> space M. X x \<le> a} = ennreal (erlang_CDF 0 l a)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
392  | 
using X_distr[of a] by (simp add: erlang_CDF_def ennreal_minus ennreal_1[symmetric] del: ennreal_1)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
393  | 
qed fact+  | 
| 50419 | 394  | 
|
395  | 
lemma (in prob_space) exponential_distributed_iff:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
396  | 
assumes "0 < l"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
397  | 
shows "distributed M lborel X (exponential_density l) \<longleftrightarrow>  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
398  | 
(X \<in> borel_measurable M \<and> (\<forall>a\<ge>0. \<P>(x in M. X x \<le> a) = 1 - exp (- a * l)))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
399  | 
using assms erlang_distributed_iff[of l X 0] by (auto simp: erlang_CDF_0)  | 
| 50419 | 400  | 
|
401  | 
||
402  | 
lemma (in prob_space) exponential_distributed_expectation:  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
403  | 
"0 < l \<Longrightarrow> distributed M lborel X (exponential_density l) \<Longrightarrow> expectation X = 1 / l"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
404  | 
using erlang_ith_moment[of l X 0 1] by simp  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
405  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
406  | 
lemma exponential_density_nonneg: "0 < l \<Longrightarrow> 0 \<le> exponential_density l x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
407  | 
by (auto simp: exponential_density_def)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
408  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
409  | 
lemma (in prob_space) exponential_distributed_min:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
410  | 
assumes "0 < l" "0 < u"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
411  | 
assumes expX: "distributed M lborel X (exponential_density l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
412  | 
assumes expY: "distributed M lborel Y (exponential_density u)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
413  | 
assumes ind: "indep_var borel X borel Y"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
414  | 
shows "distributed M lborel (\<lambda>x. min (X x) (Y x)) (exponential_density (l + u))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
415  | 
proof (subst exponential_distributed_iff, safe)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
416  | 
have randX: "random_variable borel X"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
417  | 
using expX \<open>0 < l\<close> by (simp add: exponential_distributed_iff)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
418  | 
moreover have randY: "random_variable borel Y"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
419  | 
using expY \<open>0 < u\<close> by (simp add: exponential_distributed_iff)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
420  | 
ultimately show "random_variable borel (\<lambda>x. min (X x) (Y x))" by auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
421  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
422  | 
show " 0 < l + u"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
423  | 
using \<open>0 < l\<close> \<open>0 < u\<close> by auto  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
424  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
425  | 
fix a::real assume a[arith]: "0 \<le> a"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
426  | 
have gt1[simp]: "\<P>(x in M. a < X x ) = exp (- a * l)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
427  | 
by (rule exponential_distributedD_gt[OF expX a]) fact  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
428  | 
have gt2[simp]: "\<P>(x in M. a < Y x ) = exp (- a * u)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
429  | 
by (rule exponential_distributedD_gt[OF expY a]) fact  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
430  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
431  | 
have "\<P>(x in M. a < (min (X x) (Y x)) ) = \<P>(x in M. a < (X x) \<and> a < (Y x))" by (auto intro!:arg_cong[where f=prob])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
432  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
433  | 
also have " ... = \<P>(x in M. a < (X x)) * \<P>(x in M. a< (Y x) )"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
434  | 
    using prob_indep_random_variable[OF ind, of "{a <..}" "{a <..}"] by simp
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
435  | 
also have " ... = exp (- a * (l + u))" by (auto simp:field_simps mult_exp_exp)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
436  | 
finally have indep_prob: "\<P>(x in M. a < (min (X x) (Y x)) ) = exp (- a * (l + u))" .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
437  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
438  | 
  have "{x \<in> space M. (min (X x) (Y x)) \<le>a } = (space M - {x \<in> space M. a<(min (X x) (Y x)) })"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
439  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
440  | 
  then have "1 - prob {x \<in> space M. a < (min (X x) (Y x))} = prob {x \<in> space M. (min (X x) (Y x)) \<le> a}"
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
441  | 
using randX randY by (auto simp: prob_compl)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
442  | 
  then show "prob {x \<in> space M. (min (X x) (Y x)) \<le> a} = 1 - exp (- a * (l + u))"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
443  | 
using indep_prob by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
444  | 
qed  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
445  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
446  | 
lemma (in prob_space) exponential_distributed_Min:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
447  | 
assumes finI: "finite I"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
448  | 
  assumes A: "I \<noteq> {}"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
449  | 
assumes l: "\<And>i. i \<in> I \<Longrightarrow> 0 < l i"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
450  | 
assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (exponential_density (l i))"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
451  | 
assumes ind: "indep_vars (\<lambda>i. borel) X I"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
452  | 
shows "distributed M lborel (\<lambda>x. Min ((\<lambda>i. X i x)`I)) (exponential_density (\<Sum>i\<in>I. l i))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
453  | 
using assms  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
454  | 
proof (induct rule: finite_ne_induct)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
455  | 
case (singleton i) then show ?case by simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
456  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
457  | 
case (insert i I)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
458  | 
then have "distributed M lborel (\<lambda>x. min (X i x) (Min ((\<lambda>i. X i x)`I))) (exponential_density (l i + (\<Sum>i\<in>I. l i)))"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
459  | 
by (intro exponential_distributed_min indep_vars_Min insert)  | 
| 64267 | 460  | 
(auto intro: indep_vars_subset sum_pos)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
461  | 
then show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
462  | 
using insert by simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
463  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
464  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
465  | 
lemma (in prob_space) exponential_distributed_variance:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
466  | 
"0 < l \<Longrightarrow> distributed M lborel X (exponential_density l) \<Longrightarrow> variance X = 1 / l\<^sup>2"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
467  | 
using erlang_distributed_variance[of l X 0] by simp  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
468  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
469  | 
lemma nn_integral_zero': "AE x in M. f x = 0 \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>M) = 0"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
470  | 
by (simp cong: nn_integral_cong_AE)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
471  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
472  | 
lemma convolution_erlang_density:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
473  | 
fixes k\<^sub>1 k\<^sub>2 :: nat  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
474  | 
assumes [simp, arith]: "0 < l"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
475  | 
shows "(\<lambda>x. \<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y)) * ennreal (erlang_density k\<^sub>2 l y) \<partial>lborel) =  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
476  | 
(erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
477  | 
(is "?LHS = ?RHS")  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
478  | 
proof  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
479  | 
fix x :: real  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
480  | 
have "x \<le> 0 \<or> 0 < x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
481  | 
by arith  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
482  | 
then show "?LHS x = ?RHS x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
483  | 
proof  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
484  | 
assume "x \<le> 0" then show ?thesis  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
485  | 
apply (subst nn_integral_zero')  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
486  | 
      apply (rule AE_I[where N="{0}"])
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
487  | 
apply (auto simp add: erlang_density_def not_less)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
488  | 
done  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
489  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
490  | 
note zero_le_mult_iff[simp] zero_le_divide_iff[simp]  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
491  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
492  | 
have I_eq1: "integral\<^sup>N lborel (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l) = 1"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
493  | 
using nn_integral_erlang_ith_moment[of l "Suc k\<^sub>1 + Suc k\<^sub>2 - 1" 0] by (simp del: fact_Suc)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
494  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
495  | 
    have 1: "(\<integral>\<^sup>+ x. ennreal (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l x * indicator {0<..} x) \<partial>lborel) = 1"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
496  | 
apply (subst I_eq1[symmetric])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
497  | 
unfolding erlang_density_def  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
498  | 
by (auto intro!: nn_integral_cong split:split_indicator)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
499  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
500  | 
have "prob_space (density lborel ?LHS)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
501  | 
by (intro prob_space_convolution_density)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
502  | 
(auto intro!: prob_space_erlang_density erlang_density_nonneg)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
503  | 
then have 2: "integral\<^sup>N lborel ?LHS = 1"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
504  | 
by (auto dest!: prob_space.emeasure_space_1 simp: emeasure_density)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
505  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
506  | 
    let ?I = "(integral\<^sup>N lborel (\<lambda>y. ennreal ((1 - y)^ k\<^sub>1 * y^k\<^sub>2 * indicator {0..1} y)))"
 | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
507  | 
let ?C = "(fact (Suc (k\<^sub>1 + k\<^sub>2))) / ((fact k\<^sub>1) * (fact k\<^sub>2))"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
508  | 
let ?s = "Suc k\<^sub>1 + Suc k\<^sub>2 - 1"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
509  | 
    let ?L = "(\<lambda>x. \<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x- y) * erlang_density k\<^sub>2 l y * indicator {0..x} y) \<partial>lborel)"
 | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
510  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
511  | 
    { fix x :: real assume [arith]: "0 < x"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
512  | 
have *: "\<And>x y n. (x - y * x::real)^n = x^n * (1 - y)^n"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
513  | 
unfolding power_mult_distrib[symmetric] by (simp add: field_simps)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
514  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
515  | 
have "?LHS x = ?L x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
516  | 
unfolding erlang_density_def  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
517  | 
by (auto intro!: nn_integral_cong simp: ennreal_mult split:split_indicator)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
518  | 
also have "... = (\<lambda>x. ennreal ?C * ?I * erlang_density ?s l x) x"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
519  | 
apply (subst nn_integral_real_affine[where c=x and t = 0])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
520  | 
apply (simp_all add: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] del: fact_Suc)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
521  | 
apply (intro nn_integral_cong)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
522  | 
apply (auto simp add: erlang_density_def mult_less_0_iff exp_minus field_simps exp_diff power_add *  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
523  | 
ennreal_mult[symmetric]  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
524  | 
simp del: fact_Suc split: split_indicator)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
525  | 
done  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
526  | 
finally have "(\<integral>\<^sup>+y. ennreal (erlang_density k\<^sub>1 l (x - y) * erlang_density k\<^sub>2 l y) \<partial>lborel) =  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
527  | 
(\<lambda>x. ennreal ?C * ?I * erlang_density ?s l x) x"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
528  | 
by (simp add: ennreal_mult) }  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
529  | 
note * = this  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
530  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
531  | 
assume [arith]: "0 < x"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
532  | 
    have 3: "1 = integral\<^sup>N lborel (\<lambda>xa. ?LHS xa * indicator {0<..} xa)"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
533  | 
by (subst 2[symmetric])  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
534  | 
         (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"]
 | 
| 62390 | 535  | 
simp: erlang_density_def nn_integral_multc[symmetric] indicator_def split: if_split_asm)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
536  | 
    also have "... = integral\<^sup>N lborel (\<lambda>x. (ennreal (?C) * ?I) * ((erlang_density ?s l x) * indicator {0<..} x))"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
537  | 
by (auto intro!: nn_integral_cong simp: ennreal_mult[symmetric] * split: split_indicator)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
538  | 
also have "... = ennreal (?C) * ?I"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
539  | 
using 1  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
540  | 
by (auto simp: nn_integral_cmult)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
541  | 
finally have " ennreal (?C) * ?I = 1" by presburger  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
542  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
543  | 
then show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
544  | 
using * by (simp add: ennreal_mult)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
545  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
546  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
547  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
548  | 
lemma (in prob_space) sum_indep_erlang:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
549  | 
assumes indep: "indep_var borel X borel Y"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
550  | 
assumes [simp, arith]: "0 < l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
551  | 
assumes erlX: "distributed M lborel X (erlang_density k\<^sub>1 l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
552  | 
assumes erlY: "distributed M lborel Y (erlang_density k\<^sub>2 l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
553  | 
shows "distributed M lborel (\<lambda>x. X x + Y x) (erlang_density (Suc k\<^sub>1 + Suc k\<^sub>2 - 1) l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
554  | 
using assms  | 
| 61808 | 555  | 
apply (subst convolution_erlang_density[symmetric, OF \<open>0<l\<close>])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
556  | 
apply (intro distributed_convolution)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
557  | 
apply auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
558  | 
done  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
559  | 
|
| 64267 | 560  | 
lemma (in prob_space) erlang_distributed_sum:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
561  | 
assumes finI : "finite I"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
562  | 
  assumes A: "I \<noteq> {}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
563  | 
assumes [simp, arith]: "0 < l"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
564  | 
assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (erlang_density (k i) l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
565  | 
assumes ind: "indep_vars (\<lambda>i. borel) X I"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
566  | 
shows "distributed M lborel (\<lambda>x. \<Sum>i\<in>I. X i x) (erlang_density ((\<Sum>i\<in>I. Suc (k i)) - 1) l)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
567  | 
using assms  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
568  | 
proof (induct rule: finite_ne_induct)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
569  | 
case (singleton i) then show ?case by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
570  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
571  | 
case (insert i I)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
572  | 
then have "distributed M lborel (\<lambda>x. (X i x) + (\<Sum>i\<in> I. X i x)) (erlang_density (Suc (k i) + Suc ((\<Sum>i\<in>I. Suc (k i)) - 1) - 1) l)"  | 
| 64267 | 573  | 
by(intro sum_indep_erlang indep_vars_sum) (auto intro!: indep_vars_subset)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
574  | 
also have "(\<lambda>x. (X i x) + (\<Sum>i\<in> I. X i x)) = (\<lambda>x. \<Sum>i\<in>insert i I. X i x)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
575  | 
using insert by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
576  | 
also have "Suc(k i) + Suc ((\<Sum>i\<in>I. Suc (k i)) - 1) - 1 = (\<Sum>i\<in>insert i I. Suc (k i)) - 1"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
577  | 
using insert by (auto intro!: Suc_pred simp: ac_simps)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
578  | 
finally show ?case by fast  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
579  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
580  | 
|
| 64267 | 581  | 
lemma (in prob_space) exponential_distributed_sum:  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
582  | 
assumes finI: "finite I"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
583  | 
  assumes A: "I \<noteq> {}"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
584  | 
assumes l: "0 < l"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
585  | 
assumes expX: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (exponential_density l)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
586  | 
assumes ind: "indep_vars (\<lambda>i. borel) X I"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
587  | 
shows "distributed M lborel (\<lambda>x. \<Sum>i\<in>I. X i x) (erlang_density ((card I) - 1) l)"  | 
| 64267 | 588  | 
using erlang_distributed_sum[OF assms] by simp  | 
| 50419 | 589  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
590  | 
lemma (in information_space) entropy_exponential:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
591  | 
assumes l[simp, arith]: "0 < l"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
592  | 
assumes D: "distributed M lborel X (exponential_density l)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
593  | 
shows "entropy b lborel X = log b (exp 1 / l)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
594  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
595  | 
have [simp]: "integrable lborel (exponential_density l)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
596  | 
using distributed_integrable[OF D, of "\<lambda>_. 1"] by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
597  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
598  | 
have [simp]: "integral\<^sup>L lborel (exponential_density l) = 1"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
599  | 
using distributed_integral[OF D, of "\<lambda>_. 1"] by (simp add: prob_space)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
600  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
601  | 
have [simp]: "integrable lborel (\<lambda>x. exponential_density l x * x)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
602  | 
using erlang_ith_moment_integrable[OF l D, of 1] distributed_integrable[OF D, of "\<lambda>x. x"] by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
603  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
604  | 
have [simp]: "integral\<^sup>L lborel (\<lambda>x. exponential_density l x * x) = 1 / l"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
605  | 
using erlang_ith_moment[OF l D, of 1] distributed_integral[OF D, of "\<lambda>x. x"] by simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
606  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
607  | 
have "entropy b lborel X = - (\<integral> x. exponential_density l x * log b (exponential_density l x) \<partial>lborel)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
608  | 
using D by (rule entropy_distr) simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
609  | 
also have "(\<integral> x. exponential_density l x * log b (exponential_density l x) \<partial>lborel) =  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
610  | 
(\<integral> x. (ln l * exponential_density l x - l * (exponential_density l x * x)) / ln b \<partial>lborel)"  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
611  | 
by (intro Bochner_Integration.integral_cong) (auto simp: log_def ln_mult exponential_density_def field_simps)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
612  | 
also have "\<dots> = (ln l - 1) / ln b"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
613  | 
by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
614  | 
finally show ?thesis  | 
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70365 
diff
changeset
 | 
615  | 
by (simp add: log_def ln_div) (simp add: field_split_simps)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
616  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
617  | 
|
| 61808 | 618  | 
subsection \<open>Uniform distribution\<close>  | 
| 50419 | 619  | 
|
620  | 
lemma uniform_distrI:  | 
|
621  | 
assumes X: "X \<in> measurable M M'"  | 
|
622  | 
and A: "A \<in> sets M'" "emeasure M' A \<noteq> \<infinity>" "emeasure M' A \<noteq> 0"  | 
|
623  | 
assumes distr: "\<And>B. B \<in> sets M' \<Longrightarrow> emeasure M (X -` B \<inter> space M) = emeasure M' (A \<inter> B) / emeasure M' A"  | 
|
624  | 
shows "distr M M' X = uniform_measure M' A"  | 
|
625  | 
unfolding uniform_measure_def  | 
|
626  | 
proof (intro measure_eqI)  | 
|
627  | 
let ?f = "\<lambda>x. indicator A x / emeasure M' A"  | 
|
628  | 
fix B assume B: "B \<in> sets (distr M M' X)"  | 
|
629  | 
with X have "emeasure M (X -` B \<inter> space M) = emeasure M' (A \<inter> B) / emeasure M' A"  | 
|
630  | 
by (simp add: distr[of B] measurable_sets)  | 
|
631  | 
also have "\<dots> = (1 / emeasure M' A) * emeasure M' (A \<inter> B)"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
632  | 
by (simp add: divide_ennreal_def ac_simps)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50419 
diff
changeset
 | 
633  | 
also have "\<dots> = (\<integral>\<^sup>+ x. (1 / emeasure M' A) * indicator (A \<inter> B) x \<partial>M')"  | 
| 50419 | 634  | 
using A B  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
635  | 
by (intro nn_integral_cmult_indicator[symmetric]) (auto intro!: )  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50419 
diff
changeset
 | 
636  | 
also have "\<dots> = (\<integral>\<^sup>+ x. ?f x * indicator B x \<partial>M')"  | 
| 56996 | 637  | 
by (rule nn_integral_cong) (auto split: split_indicator)  | 
| 50419 | 638  | 
finally show "emeasure (distr M M' X) B = emeasure (density M' ?f) B"  | 
639  | 
using A B X by (auto simp add: emeasure_distr emeasure_density)  | 
|
640  | 
qed simp  | 
|
641  | 
||
642  | 
lemma uniform_distrI_borel:  | 
|
643  | 
fixes A :: "real set"  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
644  | 
assumes X[measurable]: "X \<in> borel_measurable M" and A: "emeasure lborel A = ennreal r" "0 < r"  | 
| 50419 | 645  | 
and [measurable]: "A \<in> sets borel"  | 
646  | 
  assumes distr: "\<And>a. emeasure M {x\<in>space M. X x \<le> a} = emeasure lborel (A \<inter> {.. a}) / r"
 | 
|
647  | 
shows "distributed M lborel X (\<lambda>x. indicator A x / measure lborel A)"  | 
|
648  | 
proof (rule distributedI_borel_atMost)  | 
|
649  | 
let ?f = "\<lambda>x. 1 / r * indicator A x"  | 
|
650  | 
fix a  | 
|
651  | 
  have "emeasure lborel (A \<inter> {..a}) \<le> emeasure lborel A"
 | 
|
652  | 
using A by (intro emeasure_mono) auto  | 
|
653  | 
also have "\<dots> < \<infinity>"  | 
|
654  | 
using A by simp  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
655  | 
  finally have fin: "emeasure lborel (A \<inter> {..a}) \<noteq> top"
 | 
| 50419 | 656  | 
by simp  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
657  | 
from emeasure_eq_ennreal_measure[OF this]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
658  | 
  have fin_eq: "emeasure lborel (A \<inter> {..a}) / r = ennreal (measure lborel (A \<inter> {..a}) / r)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
659  | 
using A by (simp add: divide_ennreal measure_nonneg)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
660  | 
  then show "emeasure M {x\<in>space M. X x \<le> a} = ennreal (measure lborel (A \<inter> {..a}) / r)"
 | 
| 50419 | 661  | 
using distr by simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
662  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
663  | 
  have "(\<integral>\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
664  | 
    (\<integral>\<^sup>+ x. ennreal (1 / measure lborel A) * indicator (A \<inter> {..a}) x \<partial>lborel)"
 | 
| 56996 | 665  | 
by (auto intro!: nn_integral_cong split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
666  | 
  also have "\<dots> = ennreal (1 / measure lborel A) * emeasure lborel (A \<inter> {..a})"
 | 
| 61808 | 667  | 
using \<open>A \<in> sets borel\<close>  | 
| 56996 | 668  | 
by (intro nn_integral_cmult_indicator) (auto simp: measure_nonneg)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
669  | 
  also have "\<dots> = ennreal (measure lborel (A \<inter> {..a}) / r)"
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
670  | 
unfolding emeasure_eq_ennreal_measure[OF fin] using A  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
671  | 
by (simp add: measure_def ennreal_mult'[symmetric])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
672  | 
  finally show "(\<integral>\<^sup>+ x. ennreal (indicator A x / measure lborel A * indicator {..a} x) \<partial>lborel) =
 | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
673  | 
    ennreal (measure lborel (A \<inter> {..a}) / r)" .
 | 
| 
56571
 
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
 
hoelzl 
parents: 
56536 
diff
changeset
 | 
674  | 
qed (auto simp: measure_nonneg)  | 
| 50419 | 675  | 
|
676  | 
lemma (in prob_space) uniform_distrI_borel_atLeastAtMost:  | 
|
677  | 
fixes a b :: real  | 
|
678  | 
assumes X: "X \<in> borel_measurable M" and "a < b"  | 
|
679  | 
assumes distr: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> \<P>(x in M. X x \<le> t) = (t - a) / (b - a)"  | 
|
680  | 
  shows "distributed M lborel X (\<lambda>x. indicator {a..b} x / measure lborel {a..b})"
 | 
|
681  | 
proof (rule uniform_distrI_borel)  | 
|
682  | 
fix t  | 
|
683  | 
have "t < a \<or> (a \<le> t \<and> t \<le> b) \<or> b < t"  | 
|
684  | 
by auto  | 
|
685  | 
  then show "emeasure M {x\<in>space M. X x \<le> t} = emeasure lborel ({a .. b} \<inter> {..t}) / (b - a)"
 | 
|
686  | 
proof (elim disjE conjE)  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
687  | 
assume "t < a"  | 
| 50419 | 688  | 
    then have "emeasure M {x\<in>space M. X x \<le> t} \<le> emeasure M {x\<in>space M. X x \<le> a}"
 | 
689  | 
using X by (auto intro!: emeasure_mono measurable_sets)  | 
|
690  | 
also have "\<dots> = 0"  | 
|
| 61808 | 691  | 
using distr[of a] \<open>a < b\<close> by (simp add: emeasure_eq_measure)  | 
| 50419 | 692  | 
    finally have "emeasure M {x\<in>space M. X x \<le> t} = 0"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
693  | 
by (simp add: antisym measure_nonneg)  | 
| 61808 | 694  | 
with \<open>t < a\<close> show ?thesis by simp  | 
| 50419 | 695  | 
next  | 
696  | 
assume bnds: "a \<le> t" "t \<le> b"  | 
|
697  | 
    have "{a..b} \<inter> {..t} = {a..t}"
 | 
|
698  | 
using bnds by auto  | 
|
| 61808 | 699  | 
then show ?thesis using \<open>a \<le> t\<close> \<open>a < b\<close>  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
700  | 
using distr[OF bnds] by (simp add: emeasure_eq_measure divide_ennreal)  | 
| 50419 | 701  | 
next  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
702  | 
assume "b < t"  | 
| 50419 | 703  | 
    have "1 = emeasure M {x\<in>space M. X x \<le> b}"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
704  | 
using distr[of b] \<open>a < b\<close> by (simp add: one_ennreal_def emeasure_eq_measure)  | 
| 50419 | 705  | 
    also have "\<dots> \<le> emeasure M {x\<in>space M. X x \<le> t}"
 | 
| 61808 | 706  | 
using X \<open>b < t\<close> by (auto intro!: emeasure_mono measurable_sets)  | 
| 50419 | 707  | 
    finally have "emeasure M {x\<in>space M. X x \<le> t} = 1"
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
708  | 
by (simp add: antisym emeasure_eq_measure)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
709  | 
with \<open>b < t\<close> \<open>a < b\<close> show ?thesis by (simp add: measure_def divide_ennreal)  | 
| 50419 | 710  | 
qed  | 
| 61808 | 711  | 
qed (insert X \<open>a < b\<close>, auto)  | 
| 50419 | 712  | 
|
713  | 
lemma (in prob_space) uniform_distributed_measure:  | 
|
714  | 
fixes a b :: real  | 
|
715  | 
  assumes D: "distributed M lborel X (\<lambda>x. indicator {a .. b} x / measure lborel {a .. b})"
 | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
716  | 
assumes t: "a \<le> t" "t \<le> b"  | 
| 50419 | 717  | 
shows "\<P>(x in M. X x \<le> t) = (t - a) / (b - a)"  | 
718  | 
proof -  | 
|
719  | 
  have "emeasure M {x \<in> space M. X x \<le> t} = emeasure (distr M lborel X) {.. t}"
 | 
|
720  | 
using distributed_measurable[OF D]  | 
|
721  | 
by (subst emeasure_distr) (auto intro!: arg_cong2[where f=emeasure])  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
722  | 
  also have "\<dots> = (\<integral>\<^sup>+x. ennreal (1 / (b - a)) * indicator {a .. t} x \<partial>lborel)"
 | 
| 61808 | 723  | 
using distributed_borel_measurable[OF D] \<open>a \<le> t\<close> \<open>t \<le> b\<close>  | 
| 50419 | 724  | 
unfolding distributed_distr_eq_density[OF D]  | 
725  | 
by (subst emeasure_density)  | 
|
| 56996 | 726  | 
(auto intro!: nn_integral_cong simp: measure_def split: split_indicator)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
727  | 
also have "\<dots> = ennreal (1 / (b - a)) * (t - a)"  | 
| 61808 | 728  | 
using \<open>a \<le> t\<close> \<open>t \<le> b\<close>  | 
| 56996 | 729  | 
by (subst nn_integral_cmult_indicator) auto  | 
| 50419 | 730  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
731  | 
using t by (simp add: emeasure_eq_measure ennreal_mult''[symmetric] measure_nonneg)  | 
| 50419 | 732  | 
qed  | 
733  | 
||
734  | 
lemma (in prob_space) uniform_distributed_bounds:  | 
|
735  | 
fixes a b :: real  | 
|
736  | 
  assumes D: "distributed M lborel X (\<lambda>x. indicator {a .. b} x / measure lborel {a .. b})"
 | 
|
737  | 
shows "a < b"  | 
|
738  | 
proof (rule ccontr)  | 
|
739  | 
assume "\<not> a < b"  | 
|
740  | 
  then have "{a .. b} = {} \<or> {a .. b} = {a .. a}" by simp
 | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
741  | 
with uniform_distributed_params[OF D] show False  | 
| 50419 | 742  | 
by (auto simp: measure_def)  | 
743  | 
qed  | 
|
744  | 
||
745  | 
lemma (in prob_space) uniform_distributed_iff:  | 
|
746  | 
fixes a b :: real  | 
|
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
747  | 
  shows "distributed M lborel X (\<lambda>x. indicator {a..b} x / measure lborel {a..b}) \<longleftrightarrow>
 | 
| 50419 | 748  | 
    (X \<in> borel_measurable M \<and> a < b \<and> (\<forall>t\<in>{a .. b}. \<P>(x in M. X x \<le> t)= (t - a) / (b - a)))"
 | 
749  | 
using  | 
|
750  | 
uniform_distributed_bounds[of X a b]  | 
|
751  | 
uniform_distributed_measure[of X a b]  | 
|
752  | 
distributed_measurable[of M lborel X]  | 
|
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
753  | 
by (auto intro!: uniform_distrI_borel_atLeastAtMost simp del: content_real_if)  | 
| 50419 | 754  | 
|
755  | 
lemma (in prob_space) uniform_distributed_expectation:  | 
|
756  | 
fixes a b :: real  | 
|
757  | 
  assumes D: "distributed M lborel X (\<lambda>x. indicator {a .. b} x / measure lborel {a .. b})"
 | 
|
758  | 
shows "expectation X = (a + b) / 2"  | 
|
759  | 
proof (subst distributed_integral[OF D, of "\<lambda>x. x", symmetric])  | 
|
760  | 
have "a < b"  | 
|
761  | 
using uniform_distributed_bounds[OF D] .  | 
|
762  | 
||
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
763  | 
  have "(\<integral> x. indicator {a .. b} x / measure lborel {a .. b} * x \<partial>lborel) =
 | 
| 50419 | 764  | 
    (\<integral> x. (x / measure lborel {a .. b}) * indicator {a .. b} x \<partial>lborel)"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
765  | 
by (intro Bochner_Integration.integral_cong) auto  | 
| 50419 | 766  | 
  also have "(\<integral> x. (x / measure lborel {a .. b}) * indicator {a .. b} x \<partial>lborel) = (a + b) / 2"
 | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
767  | 
proof (subst integral_FTC_Icc_real)  | 
| 50419 | 768  | 
fix x  | 
| 53077 | 769  | 
    show "DERIV (\<lambda>x. x\<^sup>2 / (2 * measure lborel {a..b})) x :> x / measure lborel {a..b}"
 | 
| 50419 | 770  | 
using uniform_distributed_params[OF D]  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
771  | 
by (auto intro!: derivative_eq_intros simp del: content_real_if)  | 
| 50419 | 772  | 
    show "isCont (\<lambda>x. x / Sigma_Algebra.measure lborel {a..b}) x"
 | 
773  | 
using uniform_distributed_params[OF D]  | 
|
774  | 
by (auto intro!: isCont_divide)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50419 
diff
changeset
 | 
775  | 
    have *: "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) =
 | 
| 50419 | 776  | 
(b*b - a * a) / (2 * (b - a))"  | 
| 61808 | 777  | 
using \<open>a < b\<close>  | 
| 50419 | 778  | 
by (auto simp: measure_def power2_eq_square diff_divide_distrib[symmetric])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50419 
diff
changeset
 | 
779  | 
    show "b\<^sup>2 / (2 * measure lborel {a..b}) - a\<^sup>2 / (2 * measure lborel {a..b}) = (a + b) / 2"
 | 
| 61808 | 780  | 
using \<open>a < b\<close>  | 
| 50419 | 781  | 
unfolding * square_diff_square_factored by (auto simp: field_simps)  | 
| 61808 | 782  | 
qed (insert \<open>a < b\<close>, simp)  | 
| 50419 | 783  | 
  finally show "(\<integral> x. indicator {a .. b} x / measure lborel {a .. b} * x \<partial>lborel) = (a + b) / 2" .
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
784  | 
qed (auto simp: measure_nonneg)  | 
| 50419 | 785  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
786  | 
lemma (in prob_space) uniform_distributed_variance:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
787  | 
fixes a b :: real  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
788  | 
  assumes D: "distributed M lborel X (\<lambda>x. indicator {a .. b} x / measure lborel {a .. b})"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
789  | 
shows "variance X = (b - a)\<^sup>2 / 12"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
790  | 
proof (subst distributed_variance)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
791  | 
have [arith]: "a < b" using uniform_distributed_bounds[OF D] .  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
792  | 
  let ?\<mu> = "expectation X" let ?D = "\<lambda>x. indicator {a..b} (x + ?\<mu>) / measure lborel {a..b}"
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
793  | 
  have "(\<integral>x. x\<^sup>2 * (?D x) \<partial>lborel) = (\<integral>x. x\<^sup>2 * (indicator {a - ?\<mu> .. b - ?\<mu>} x) / measure lborel {a .. b} \<partial>lborel)"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
794  | 
by (intro Bochner_Integration.integral_cong) (auto split: split_indicator)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
795  | 
also have "\<dots> = (b - a)\<^sup>2 / 12"  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57418 
diff
changeset
 | 
796  | 
by (simp add: integral_power uniform_distributed_expectation[OF D])  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
797  | 
(simp add: eval_nat_numeral field_simps )  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
798  | 
finally show "(\<integral>x. x\<^sup>2 * ?D x \<partial>lborel) = (b - a)\<^sup>2 / 12" .  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
799  | 
qed (auto intro: D simp del: content_real_if)  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
56996 
diff
changeset
 | 
800  | 
|
| 61808 | 801  | 
subsection \<open>Normal distribution\<close>  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
802  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
803  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
804  | 
definition normal_density :: "real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real" where  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
805  | 
"normal_density \<mu> \<sigma> x = 1 / sqrt (2 * pi * \<sigma>\<^sup>2) * exp (-(x - \<mu>)\<^sup>2/ (2 * \<sigma>\<^sup>2))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
806  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
807  | 
abbreviation std_normal_density :: "real \<Rightarrow> real" where  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
808  | 
"std_normal_density \<equiv> normal_density 0 1"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
809  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
810  | 
lemma std_normal_density_def: "std_normal_density x = (1 / sqrt (2 * pi)) * exp (- x\<^sup>2 / 2)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
811  | 
unfolding normal_density_def by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
812  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
813  | 
lemma normal_density_nonneg[simp]: "0 \<le> normal_density \<mu> \<sigma> x"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
814  | 
by (auto simp: normal_density_def)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
815  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
816  | 
lemma normal_density_pos: "0 < \<sigma> \<Longrightarrow> 0 < normal_density \<mu> \<sigma> x"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
817  | 
by (auto simp: normal_density_def)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
818  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
819  | 
lemma borel_measurable_normal_density[measurable]: "normal_density \<mu> \<sigma> \<in> borel_measurable borel"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
820  | 
by (auto simp: normal_density_def[abs_def])  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
821  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
822  | 
lemma gaussian_moment_0:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
823  | 
  "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2)) (sqrt pi / 2)"
 | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
824  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
825  | 
  let ?pI = "\<lambda>f. (\<integral>\<^sup>+s. f (s::real) * indicator {0..} s \<partial>lborel)"
 | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
826  | 
let ?gauss = "\<lambda>x. exp (- x\<^sup>2)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
827  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
828  | 
  let ?I = "indicator {0<..} :: real \<Rightarrow> real"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
829  | 
let ?ff= "\<lambda>x s. x * exp (- x\<^sup>2 * (1 + s\<^sup>2)) :: real"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
830  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
831  | 
have *: "?pI ?gauss = (\<integral>\<^sup>+x. ?gauss x * ?I x \<partial>lborel)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
832  | 
    by (intro nn_integral_cong_AE AE_I[where N="{0}"]) (auto split: split_indicator)
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
833  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
834  | 
have "?pI ?gauss * ?pI ?gauss = (\<integral>\<^sup>+x. \<integral>\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \<partial>lborel \<partial>lborel)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
835  | 
by (auto simp: nn_integral_cmult[symmetric] nn_integral_multc[symmetric] * ennreal_mult[symmetric]  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
836  | 
intro!: nn_integral_cong split: split_indicator)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
837  | 
also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+s. ?ff x s * ?I s * ?I x \<partial>lborel \<partial>lborel)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
838  | 
proof (rule nn_integral_cong, cases)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
839  | 
fix x :: real assume "x \<noteq> 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
840  | 
then show "(\<integral>\<^sup>+s. ?gauss x * ?gauss s * ?I s * ?I x \<partial>lborel) = (\<integral>\<^sup>+s. ?ff x s * ?I s * ?I x \<partial>lborel)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
841  | 
by (subst nn_integral_real_affine[where t="0" and c="x"])  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
842  | 
(auto simp: mult_exp_exp nn_integral_cmult[symmetric] field_simps zero_less_mult_iff ennreal_mult[symmetric]  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
843  | 
intro!: nn_integral_cong split: split_indicator)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
844  | 
qed simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
845  | 
also have "... = \<integral>\<^sup>+s. \<integral>\<^sup>+x. ?ff x s * ?I s * ?I x \<partial>lborel \<partial>lborel"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
846  | 
by (rule lborel_pair.Fubini'[symmetric]) auto  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
847  | 
also have "... = ?pI (\<lambda>s. ?pI (\<lambda>x. ?ff x s))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
848  | 
by (rule nn_integral_cong_AE)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
849  | 
       (auto intro!: nn_integral_cong_AE AE_I[where N="{0}"] split: split_indicator_asm)
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
850  | 
also have "\<dots> = ?pI (\<lambda>s. ennreal (1 / (2 * (1 + s\<^sup>2))))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
851  | 
proof (intro nn_integral_cong ennreal_mult_right_cong)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
852  | 
fix s :: real show "?pI (\<lambda>x. ?ff x s) = ennreal (1 / (2 * (1 + s\<^sup>2)))"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
853  | 
proof (subst nn_integral_FTC_atLeast)  | 
| 61973 | 854  | 
have "((\<lambda>a. - (exp (- (a\<^sup>2 * (1 + s\<^sup>2))) / (2 + 2 * s\<^sup>2))) \<longlongrightarrow> (- (0 / (2 + 2 * s\<^sup>2)))) at_top"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
855  | 
apply (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_compose[OF filterlim_uminus_at_bot_at_top])  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
856  | 
apply (subst mult.commute)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
857  | 
apply (auto intro!: filterlim_tendsto_pos_mult_at_top  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
858  | 
filterlim_at_top_mult_at_top[OF filterlim_ident filterlim_ident]  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
859  | 
simp: add_pos_nonneg power2_eq_square add_nonneg_eq_0_iff)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
860  | 
done  | 
| 61973 | 861  | 
then show "((\<lambda>a. - (exp (- a\<^sup>2 - s\<^sup>2 * a\<^sup>2) / (2 + 2 * s\<^sup>2))) \<longlongrightarrow> 0) at_top"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
862  | 
by (simp add: field_simps)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
863  | 
qed (auto intro!: derivative_eq_intros simp: field_simps add_nonneg_eq_0_iff)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
864  | 
qed  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
865  | 
also have "... = ennreal (pi / 4)"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
866  | 
proof (subst nn_integral_FTC_atLeast)  | 
| 61973 | 867  | 
show "((\<lambda>a. arctan a / 2) \<longlongrightarrow> (pi / 2) / 2 ) at_top"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
868  | 
by (intro tendsto_intros) (simp_all add: tendsto_arctan_at_top)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
869  | 
qed (auto intro!: derivative_eq_intros simp: add_nonneg_eq_0_iff field_simps power2_eq_square)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
870  | 
finally have "?pI ?gauss^2 = pi / 4"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
871  | 
by (simp add: power2_eq_square)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
872  | 
then have "?pI ?gauss = sqrt (pi / 4)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
873  | 
using power_eq_iff_eq_base[of 2 "enn2real (?pI ?gauss)" "sqrt (pi / 4)"]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
874  | 
by (cases "?pI ?gauss") (auto simp: power2_eq_square ennreal_mult[symmetric] ennreal_top_mult)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
875  | 
  also have "?pI ?gauss = (\<integral>\<^sup>+x. indicator {0..} x *\<^sub>R exp (- x\<^sup>2) \<partial>lborel)"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
876  | 
by (intro nn_integral_cong) (simp split: split_indicator)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
877  | 
also have "sqrt (pi / 4) = sqrt pi / 2"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
878  | 
by (simp add: real_sqrt_divide)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
879  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
880  | 
by (rule has_bochner_integral_nn_integral[rotated 3])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
881  | 
auto  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
882  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
883  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
884  | 
lemma gaussian_moment_1:  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
885  | 
  "has_bochner_integral lborel (\<lambda>x::real. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x)) (1 / 2)"
 | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
886  | 
proof -  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
887  | 
  have "(\<integral>\<^sup>+x. indicator {0..} x *\<^sub>R (exp (- x\<^sup>2) * x) \<partial>lborel) =
 | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
888  | 
    (\<integral>\<^sup>+x. ennreal (x * exp (- x\<^sup>2)) * indicator {0..} x \<partial>lborel)"
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
889  | 
by (intro nn_integral_cong)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
890  | 
(auto simp: ac_simps split: split_indicator)  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
891  | 
also have "\<dots> = ennreal (0 - (- exp (- 0\<^sup>2) / 2))"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
892  | 
proof (rule nn_integral_FTC_atLeast)  | 
| 61973 | 893  | 
have "((\<lambda>x::real. - exp (- x\<^sup>2) / 2) \<longlongrightarrow> - 0 / 2) at_top"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
894  | 
by (intro tendsto_divide tendsto_minus filterlim_compose[OF exp_at_bot]  | 
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
895  | 
filterlim_compose[OF filterlim_uminus_at_bot_at_top]  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
896  | 
filterlim_pow_at_top filterlim_ident)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
897  | 
auto  | 
| 61973 | 898  | 
then show "((\<lambda>a::real. - exp (- a\<^sup>2) / 2) \<longlongrightarrow> 0) at_top"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
899  | 
by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
900  | 
qed (auto intro!: derivative_eq_intros)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
901  | 
also have "\<dots> = ennreal (1 / 2)"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
902  | 
by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
903  | 
finally show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
904  | 
by (rule has_bochner_integral_nn_integral[rotated 3])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
905  | 
(auto split: split_indicator)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
906  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
907  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
908  | 
lemma  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
909  | 
fixes k :: nat  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
910  | 
shows gaussian_moment_even_pos:  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
911  | 
    "has_bochner_integral lborel (\<lambda>x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k)))
 | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
912  | 
((sqrt pi / 2) * (fact (2 * k) / (2 ^ (2 * k) * fact k)))"  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
913  | 
(is "?even")  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
914  | 
and gaussian_moment_odd_pos:  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
915  | 
      "has_bochner_integral lborel (\<lambda>x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*x^(2 * k + 1))) (fact k / 2)"
 | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
916  | 
(is "?odd")  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
917  | 
proof -  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
918  | 
let ?M = "\<lambda>k x. exp (- x\<^sup>2) * x^k :: real"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
919  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
920  | 
  { fix k I assume Mk: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R ?M k x) I"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
921  | 
have "2 \<noteq> (0::real)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
922  | 
by linarith  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
923  | 
    let ?f = "\<lambda>b. \<integral>x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..b} x \<partial>lborel"
 | 
| 61973 | 924  | 
    have "((\<lambda>b. (k + 1) / 2 * (\<integral>x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x) \<partial>lborel) - ?M (k + 1) b / 2) \<longlongrightarrow>
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
925  | 
        (k + 1) / 2 * (\<integral>x. indicator {0..} x *\<^sub>R ?M k x \<partial>lborel) - 0 / 2) at_top" (is ?tendsto)
 | 
| 61808 | 926  | 
proof (intro tendsto_intros \<open>2 \<noteq> 0\<close> tendsto_integral_at_top sets_lborel Mk[THEN integrable.intros])  | 
| 61973 | 927  | 
show "(?M (k + 1) \<longlongrightarrow> 0) at_top"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
928  | 
proof cases  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
929  | 
assume "even k"  | 
| 61973 | 930  | 
have "((\<lambda>x. ((x\<^sup>2)^(k div 2 + 1) / exp (x\<^sup>2)) * (1 / x) :: real) \<longlongrightarrow> 0 * 0) at_top"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
931  | 
by (intro tendsto_intros tendsto_divide_0[OF tendsto_const] filterlim_compose[OF tendsto_power_div_exp_0]  | 
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
932  | 
filterlim_at_top_imp_at_infinity filterlim_ident filterlim_pow_at_top filterlim_ident)  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
933  | 
auto  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
934  | 
also have "(\<lambda>x. ((x\<^sup>2)^(k div 2 + 1) / exp (x\<^sup>2)) * (1 / x) :: real) = ?M (k + 1)"  | 
| 61808 | 935  | 
using \<open>even k\<close> by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: evenE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
936  | 
finally show ?thesis by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
937  | 
next  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
938  | 
assume "odd k"  | 
| 61973 | 939  | 
have "((\<lambda>x. ((x\<^sup>2)^((k - 1) div 2 + 1) / exp (x\<^sup>2)) :: real) \<longlongrightarrow> 0) at_top"  | 
| 
57275
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
940  | 
by (intro filterlim_compose[OF tendsto_power_div_exp_0] filterlim_at_top_imp_at_infinity  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
941  | 
filterlim_ident filterlim_pow_at_top)  | 
| 
 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 
hoelzl 
parents: 
57254 
diff
changeset
 | 
942  | 
auto  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
943  | 
also have "(\<lambda>x. ((x\<^sup>2)^((k - 1) div 2 + 1) / exp (x\<^sup>2)) :: real) = ?M (k + 1)"  | 
| 61808 | 944  | 
using \<open>odd k\<close> by (auto simp: fun_eq_iff exp_minus field_simps power2_eq_square power_mult elim: oddE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
945  | 
finally show ?thesis by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
946  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
947  | 
qed  | 
| 61973 | 948  | 
    also have "?tendsto \<longleftrightarrow> ((?f \<longlongrightarrow> (k + 1) / 2 * (\<integral>x. indicator {0..} x *\<^sub>R ?M k x \<partial>lborel) - 0 / 2) at_top)"
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
949  | 
proof (intro filterlim_cong refl eventually_at_top_linorder[THEN iffD2] exI[of _ 0] allI impI)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
950  | 
fix b :: real assume b: "0 \<le> b"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
951  | 
      have "Suc k * (\<integral>x. indicator {0..b} x *\<^sub>R ?M k x \<partial>lborel) = (\<integral>x. indicator {0..b} x *\<^sub>R (exp (- x\<^sup>2) * ((Suc k) * x ^ k)) \<partial>lborel)"
 | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
952  | 
unfolding integral_mult_right_zero[symmetric] by (intro Bochner_Integration.integral_cong) auto  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
953  | 
also have "\<dots> = exp (- b\<^sup>2) * b ^ (Suc k) - exp (- 0\<^sup>2) * 0 ^ (Suc k) -  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
954  | 
          (\<integral>x. indicator {0..b} x *\<^sub>R (- 2 * x * exp (- x\<^sup>2) * x ^ (Suc k)) \<partial>lborel)"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
955  | 
by (rule integral_by_parts')  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
956  | 
(auto intro!: derivative_eq_intros b  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
957  | 
simp: diff_Suc of_nat_Suc field_simps split: nat.split)  | 
| 
67977
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
958  | 
      also have "... = exp (- b\<^sup>2) * b ^ (Suc k) - (\<integral>x. indicator {0..b} x *\<^sub>R (- 2 * (exp (- x\<^sup>2) * x ^ (k + 2))) \<partial>lborel)"
 | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
959  | 
by (auto simp: intro!: Bochner_Integration.integral_cong)  | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
960  | 
      also have "... = exp (- b\<^sup>2) * b ^ (Suc k) + 2 * (\<integral>x. indicator {0..b} x *\<^sub>R ?M (k + 2) x \<partial>lborel)"
 | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
961  | 
unfolding Bochner_Integration.integral_mult_right_zero [symmetric]  | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
962  | 
by (simp del: real_scaleR_def)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
963  | 
      finally have "Suc k * (\<integral>x. indicator {0..b} x *\<^sub>R ?M k x \<partial>lborel) =
 | 
| 
67977
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
964  | 
        exp (- b\<^sup>2) * b ^ (Suc k) + 2 * (\<integral>x. indicator {0..b} x *\<^sub>R ?M (k + 2) x \<partial>lborel)" .
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
965  | 
      then show "(k + 1) / 2 * (\<integral>x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x)\<partial>lborel) - exp (- b\<^sup>2) * b ^ (k + 1) / 2 = ?f b"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
966  | 
by (simp add: field_simps atLeastAtMost_def indicator_inter_arith)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
967  | 
qed  | 
| 61973 | 968  | 
    finally have int_M_at_top: "((?f \<longlongrightarrow> (k + 1) / 2 * (\<integral>x. indicator {0..} x *\<^sub>R ?M k x \<partial>lborel)) at_top)"
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
969  | 
by simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
970  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
971  | 
    have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R ?M (k + 2) x) ((k + 1) / 2 * I)"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
972  | 
proof (rule has_bochner_integral_monotone_convergence_at_top)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
973  | 
fix y :: real  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
974  | 
      have *: "(\<lambda>x. indicator {0..} x *\<^sub>R ?M (k + 2) x * indicator {..y} x::real) =
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
975  | 
            (\<lambda>x. indicator {0..y} x *\<^sub>R ?M (k + 2) x)"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
976  | 
by rule (simp split: split_indicator)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
977  | 
      show "integrable lborel (\<lambda>x. indicator {0..} x *\<^sub>R (?M (k + 2) x) * indicator {..y} x::real)"
 | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
978  | 
unfolding * by (rule borel_integrable_compact) (auto intro!: continuous_intros)  | 
| 61973 | 979  | 
show "((?f \<longlongrightarrow> (k + 1) / 2 * I) at_top)"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
980  | 
using int_M_at_top has_bochner_integral_integral_eq[OF Mk] by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
981  | 
qed (auto split: split_indicator) }  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
982  | 
note step = this  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
983  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
984  | 
show ?even  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
985  | 
proof (induct k)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
986  | 
case (Suc k)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
987  | 
note step[OF this]  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
988  | 
also have "(real (2 * k + 1) / 2 * (sqrt pi / 2 * ((fact (2 * k)) / ((2::real)^(2*k) * fact k)))) =  | 
| 
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
989  | 
sqrt pi / 2 * ((fact (2 * Suc k)) / ((2::real)^(2 * Suc k) * fact (Suc k)))"  | 
| 
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
990  | 
apply (simp add: field_simps del: fact_Suc)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
991  | 
apply (simp add: of_nat_mult field_simps)  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
992  | 
done  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
993  | 
finally show ?case  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
994  | 
by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
995  | 
qed (insert gaussian_moment_0, simp)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
996  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
997  | 
show ?odd  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
998  | 
proof (induct k)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
999  | 
case (Suc k)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1000  | 
note step[OF this]  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
1001  | 
also have "(real (2 * k + 1 + 1) / (2::real) * ((fact k) / 2)) = (fact (Suc k)) / 2"  | 
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70365 
diff
changeset
 | 
1002  | 
by (simp add: field_simps of_nat_Suc field_split_simps del: fact_Suc) (simp add: field_simps)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1003  | 
finally show ?case  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1004  | 
by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1005  | 
qed (insert gaussian_moment_1, simp)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1006  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1007  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1008  | 
context  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1009  | 
fixes k :: nat and \<mu> \<sigma> :: real assumes [arith]: "0 < \<sigma>"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1010  | 
begin  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1011  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1012  | 
lemma normal_moment_even:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1013  | 
"has_bochner_integral lborel (\<lambda>x. normal_density \<mu> \<sigma> x * (x - \<mu>) ^ (2 * k)) (fact (2 * k) / ((2 / \<sigma>\<^sup>2)^k * fact k))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1014  | 
proof -  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1015  | 
have eq: "\<And>x::real. x\<^sup>2^k = (x^k)\<^sup>2"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1016  | 
by (simp add: power_mult[symmetric] ac_simps)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1017  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1018  | 
have "has_bochner_integral lborel (\<lambda>x. exp (-x\<^sup>2)*x^(2 * k))  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1019  | 
(sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k)))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1020  | 
using has_bochner_integral_even_function[OF gaussian_moment_even_pos[where k=k]] by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1021  | 
then have "has_bochner_integral lborel (\<lambda>x. (exp (-x\<^sup>2)*x^(2 * k)) * ((2*\<sigma>\<^sup>2)^k / sqrt (2 * pi * \<sigma>\<^sup>2)))  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1022  | 
((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\<sigma>\<^sup>2)^k / sqrt (2 * pi * \<sigma>\<^sup>2)))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1023  | 
by (rule has_bochner_integral_mult_left)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1024  | 
also have "(\<lambda>x. (exp (-x\<^sup>2)*x^(2 * k)) * ((2*\<sigma>\<^sup>2)^k / sqrt (2 * pi * \<sigma>\<^sup>2))) =  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1025  | 
(\<lambda>x. exp (- ((sqrt 2 * \<sigma>) * x)\<^sup>2 / (2*\<sigma>\<^sup>2)) * ((sqrt 2 * \<sigma>) * x) ^ (2 * k) / sqrt (2 * pi * \<sigma>\<^sup>2))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1026  | 
by (auto simp: fun_eq_iff field_simps real_sqrt_power[symmetric] real_sqrt_mult  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1027  | 
real_sqrt_divide power_mult eq)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1028  | 
also have "((sqrt pi * (fact (2 * k) / (2 ^ (2 * k) * fact k))) * ((2*\<sigma>\<^sup>2)^k / sqrt (2 * pi * \<sigma>\<^sup>2))) =  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
1029  | 
(inverse (sqrt 2 * \<sigma>) * ((fact (2 * k))) / ((2/\<sigma>\<^sup>2) ^ k * (fact k)))"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1030  | 
by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_mult  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1031  | 
power2_eq_square)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1032  | 
finally show ?thesis  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1033  | 
unfolding normal_density_def  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1034  | 
by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \<sigma>" and t=\<mu>]) simp_all  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1035  | 
qed  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1036  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1037  | 
lemma normal_moment_abs_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1038  | 
"has_bochner_integral lborel (\<lambda>x. normal_density \<mu> \<sigma> x * \<bar>x - \<mu>\<bar>^(2 * k + 1)) (2^k * \<sigma>^(2 * k + 1) * fact k * sqrt (2 / pi))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1039  | 
proof -  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
1040  | 
  have "has_bochner_integral lborel (\<lambda>x::real. indicator {0..} x *\<^sub>R (exp (-x\<^sup>2)*\<bar>x\<bar>^(2 * k + 1))) (fact k / 2)"
 | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1041  | 
by (rule has_bochner_integral_cong[THEN iffD1, OF _ _ _ gaussian_moment_odd_pos[of k]]) auto  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1042  | 
from has_bochner_integral_even_function[OF this]  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
1043  | 
have "has_bochner_integral lborel (\<lambda>x::real. exp (-x\<^sup>2)*\<bar>x\<bar>^(2 * k + 1)) (fact k)"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1044  | 
by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1045  | 
then have "has_bochner_integral lborel (\<lambda>x. (exp (-x\<^sup>2)*\<bar>x\<bar>^(2 * k + 1)) * (2^k * \<sigma>^(2 * k + 1) / sqrt (pi * \<sigma>\<^sup>2)))  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1046  | 
(fact k * (2^k * \<sigma>^(2 * k + 1) / sqrt (pi * \<sigma>\<^sup>2)))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1047  | 
by (rule has_bochner_integral_mult_left)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1048  | 
also have "(\<lambda>x. (exp (-x\<^sup>2)*\<bar>x\<bar>^(2 * k + 1)) * (2^k * \<sigma>^(2 * k + 1) / sqrt (pi * \<sigma>\<^sup>2))) =  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1049  | 
(\<lambda>x. exp (- (((sqrt 2 * \<sigma>) * x)\<^sup>2 / (2 * \<sigma>\<^sup>2))) * \<bar>sqrt 2 * \<sigma> * x\<bar> ^ (2 * k + 1) / sqrt (2 * pi * \<sigma>\<^sup>2))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1050  | 
by (simp add: field_simps abs_mult real_sqrt_power[symmetric] power_mult real_sqrt_mult)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1051  | 
also have "(fact k * (2^k * \<sigma>^(2 * k + 1) / sqrt (pi * \<sigma>\<^sup>2))) =  | 
| 
59730
 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 
paulson <lp15@cam.ac.uk> 
parents: 
59587 
diff
changeset
 | 
1052  | 
(inverse (sqrt 2) * inverse \<sigma> * (2 ^ k * (\<sigma> * \<sigma> ^ (2 * k)) * (fact k) * sqrt (2 / pi)))"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1053  | 
by (auto simp: fun_eq_iff power_mult field_simps real_sqrt_power[symmetric] real_sqrt_divide  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1054  | 
real_sqrt_mult)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1055  | 
finally show ?thesis  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1056  | 
unfolding normal_density_def  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1057  | 
by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \<sigma>" and t=\<mu>])  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1058  | 
simp_all  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1059  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1060  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1061  | 
lemma normal_moment_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1062  | 
"has_bochner_integral lborel (\<lambda>x. normal_density \<mu> \<sigma> x * (x - \<mu>)^(2 * k + 1)) 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1063  | 
proof -  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1064  | 
have "has_bochner_integral lborel (\<lambda>x. exp (- x\<^sup>2) * x^(2 * k + 1)::real) 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1065  | 
using gaussian_moment_odd_pos by (rule has_bochner_integral_odd_function) simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1066  | 
then have "has_bochner_integral lborel (\<lambda>x. (exp (-x\<^sup>2)*x^(2 * k + 1)) * (2^k*\<sigma>^(2*k)/sqrt pi))  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1067  | 
(0 * (2^k*\<sigma>^(2*k)/sqrt pi))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1068  | 
by (rule has_bochner_integral_mult_left)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1069  | 
also have "(\<lambda>x. (exp (-x\<^sup>2)*x^(2 * k + 1)) * (2^k*\<sigma>^(2*k)/sqrt pi)) =  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1070  | 
(\<lambda>x. exp (- ((sqrt 2 * \<sigma> * x)\<^sup>2 / (2 * \<sigma>\<^sup>2))) *  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1071  | 
(sqrt 2 * \<sigma> * x * (sqrt 2 * \<sigma> * x) ^ (2 * k)) /  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1072  | 
sqrt (2 * pi * \<sigma>\<^sup>2))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1073  | 
unfolding real_sqrt_mult  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1074  | 
by (simp add: field_simps abs_mult real_sqrt_power[symmetric] power_mult fun_eq_iff)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1075  | 
finally show ?thesis  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1076  | 
unfolding normal_density_def  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1077  | 
by (subst lborel_has_bochner_integral_real_affine_iff[where c="sqrt 2 * \<sigma>" and t=\<mu>]) simp_all  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1078  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1079  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1080  | 
lemma integral_normal_moment_even:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1081  | 
"integral\<^sup>L lborel (\<lambda>x. normal_density \<mu> \<sigma> x * (x - \<mu>)^(2 * k)) = fact (2 * k) / ((2 / \<sigma>\<^sup>2)^k * fact k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1082  | 
using normal_moment_even by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1083  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1084  | 
lemma integral_normal_moment_abs_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1085  | 
"integral\<^sup>L lborel (\<lambda>x. normal_density \<mu> \<sigma> x * \<bar>x - \<mu>\<bar>^(2 * k + 1)) = 2 ^ k * \<sigma> ^ (2 * k + 1) * fact k * sqrt (2 / pi)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1086  | 
using normal_moment_abs_odd by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1087  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1088  | 
lemma integral_normal_moment_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1089  | 
"integral\<^sup>L lborel (\<lambda>x. normal_density \<mu> \<sigma> x * (x - \<mu>)^(2 * k + 1)) = 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1090  | 
using normal_moment_odd by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1091  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1092  | 
end  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1093  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1094  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1095  | 
context  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1096  | 
fixes \<sigma> :: real  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1097  | 
assumes \<sigma>_pos[arith]: "0 < \<sigma>"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1098  | 
begin  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1099  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1100  | 
lemma normal_moment_nz_1: "has_bochner_integral lborel (\<lambda>x. normal_density \<mu> \<sigma> x * x) \<mu>"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1101  | 
proof -  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1102  | 
note normal_moment_even[OF \<sigma>_pos, of \<mu> 0]  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1103  | 
note normal_moment_odd[OF \<sigma>_pos, of \<mu> 0] has_bochner_integral_mult_left[of \<mu>, OF this]  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1104  | 
note has_bochner_integral_add[OF this]  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1105  | 
then show ?thesis  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1106  | 
by (simp add: power2_eq_square field_simps)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1107  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1108  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1109  | 
lemma integral_normal_moment_nz_1:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1110  | 
"integral\<^sup>L lborel (\<lambda>x. normal_density \<mu> \<sigma> x * x) = \<mu>"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1111  | 
using normal_moment_nz_1 by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1112  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1113  | 
lemma integrable_normal_moment_nz_1: "integrable lborel (\<lambda>x. normal_density \<mu> \<sigma> x * x)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1114  | 
using normal_moment_nz_1 by rule  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1115  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1116  | 
lemma integrable_normal_moment: "integrable lborel (\<lambda>x. normal_density \<mu> \<sigma> x * (x - \<mu>)^k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1117  | 
proof cases  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1118  | 
assume "even k" then show ?thesis  | 
| 
58709
 
efdc6c533bd3
prefer generic elimination rules for even/odd over specialized unfold rules for nat
 
haftmann 
parents: 
57514 
diff
changeset
 | 
1119  | 
using integrable.intros[OF normal_moment_even] by (auto elim: evenE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1120  | 
next  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1121  | 
assume "odd k" then show ?thesis  | 
| 
58709
 
efdc6c533bd3
prefer generic elimination rules for even/odd over specialized unfold rules for nat
 
haftmann 
parents: 
57514 
diff
changeset
 | 
1122  | 
using integrable.intros[OF normal_moment_odd] by (auto elim: oddE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1123  | 
qed  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1124  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1125  | 
lemma integrable_normal_moment_abs: "integrable lborel (\<lambda>x. normal_density \<mu> \<sigma> x * \<bar>x - \<mu>\<bar>^k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1126  | 
proof cases  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1127  | 
assume "even k" then show ?thesis  | 
| 
58709
 
efdc6c533bd3
prefer generic elimination rules for even/odd over specialized unfold rules for nat
 
haftmann 
parents: 
57514 
diff
changeset
 | 
1128  | 
using integrable.intros[OF normal_moment_even] by (auto simp add: power_even_abs elim: evenE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1129  | 
next  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1130  | 
assume "odd k" then show ?thesis  | 
| 
58709
 
efdc6c533bd3
prefer generic elimination rules for even/odd over specialized unfold rules for nat
 
haftmann 
parents: 
57514 
diff
changeset
 | 
1131  | 
using integrable.intros[OF normal_moment_abs_odd] by (auto elim: oddE)  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1132  | 
qed  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1133  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1134  | 
lemma integrable_normal_density[simp, intro]: "integrable lborel (normal_density \<mu> \<sigma>)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1135  | 
using integrable_normal_moment[of \<mu> 0] by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1136  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1137  | 
lemma integral_normal_density[simp]: "(\<integral>x. normal_density \<mu> \<sigma> x \<partial>lborel) = 1"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1138  | 
using integral_normal_moment_even[of \<sigma> \<mu> 0] by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1139  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1140  | 
lemma prob_space_normal_density:  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1141  | 
"prob_space (density lborel (normal_density \<mu> \<sigma>))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1142  | 
proof qed (simp add: emeasure_density nn_integral_eq_integral normal_density_nonneg)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1143  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1144  | 
end  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1145  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1146  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1147  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1148  | 
context  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1149  | 
fixes k :: nat  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1150  | 
begin  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1151  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1152  | 
lemma std_normal_moment_even:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1153  | 
"has_bochner_integral lborel (\<lambda>x. std_normal_density x * x ^ (2 * k)) (fact (2 * k) / (2^k * fact k))"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1154  | 
using normal_moment_even[of 1 0 k] by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1155  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1156  | 
lemma std_normal_moment_abs_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1157  | 
"has_bochner_integral lborel (\<lambda>x. std_normal_density x * \<bar>x\<bar>^(2 * k + 1)) (sqrt (2/pi) * 2^k * fact k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1158  | 
using normal_moment_abs_odd[of 1 0 k] by (simp add: ac_simps)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1159  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1160  | 
lemma std_normal_moment_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1161  | 
"has_bochner_integral lborel (\<lambda>x. std_normal_density x * x^(2 * k + 1)) 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1162  | 
using normal_moment_odd[of 1 0 k] by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1163  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1164  | 
lemma integral_std_normal_moment_even:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1165  | 
"integral\<^sup>L lborel (\<lambda>x. std_normal_density x * x^(2*k)) = fact (2 * k) / (2^k * fact k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1166  | 
using std_normal_moment_even by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1167  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1168  | 
lemma integral_std_normal_moment_abs_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1169  | 
"integral\<^sup>L lborel (\<lambda>x. std_normal_density x * \<bar>x\<bar>^(2 * k + 1)) = sqrt (2 / pi) * 2^k * fact k"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1170  | 
using std_normal_moment_abs_odd by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1171  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1172  | 
lemma integral_std_normal_moment_odd:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1173  | 
"integral\<^sup>L lborel (\<lambda>x. std_normal_density x * x^(2 * k + 1)) = 0"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1174  | 
using std_normal_moment_odd by (rule has_bochner_integral_integral_eq)  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1175  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1176  | 
lemma integrable_std_normal_moment_abs: "integrable lborel (\<lambda>x. std_normal_density x * \<bar>x\<bar>^k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1177  | 
using integrable_normal_moment_abs[of 1 0 k] by simp  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1178  | 
|
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1179  | 
lemma integrable_std_normal_moment: "integrable lborel (\<lambda>x. std_normal_density x * x^k)"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1180  | 
using integrable_normal_moment[of 1 0 k] by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1181  | 
|
| 50419 | 1182  | 
end  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1183  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1184  | 
lemma (in prob_space) normal_density_affine:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1185  | 
assumes X: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1186  | 
assumes [simp, arith]: "0 < \<sigma>" "\<alpha> \<noteq> 0"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1187  | 
shows "distributed M lborel (\<lambda>x. \<beta> + \<alpha> * X x) (normal_density (\<beta> + \<alpha> * \<mu>) (\<bar>\<alpha>\<bar> * \<sigma>))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1188  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1189  | 
have eq: "\<And>x. \<bar>\<alpha>\<bar> * normal_density (\<beta> + \<alpha> * \<mu>) (\<bar>\<alpha>\<bar> * \<sigma>) (x * \<alpha> + \<beta>) =  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1190  | 
normal_density \<mu> \<sigma> x"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1191  | 
by (simp add: normal_density_def real_sqrt_mult field_simps)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1192  | 
(simp add: power2_eq_square field_simps)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1193  | 
show ?thesis  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1194  | 
by (rule distributed_affineI[OF _ \<open>\<alpha> \<noteq> 0\<close>, where t=\<beta>])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1195  | 
(simp_all add: eq X ennreal_mult'[symmetric])  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1196  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1197  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1198  | 
lemma (in prob_space) normal_standard_normal_convert:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1199  | 
assumes pos_var[simp, arith]: "0 < \<sigma>"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1200  | 
shows "distributed M lborel X (normal_density \<mu> \<sigma>) = distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) std_normal_density"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1201  | 
proof auto  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1202  | 
assume "distributed M lborel X (\<lambda>x. ennreal (normal_density \<mu> \<sigma> x))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1203  | 
then have "distributed M lborel (\<lambda>x. -\<mu> / \<sigma> + (1/\<sigma>) * X x) (\<lambda>x. ennreal (normal_density (-\<mu> / \<sigma> + (1/\<sigma>)* \<mu>) (\<bar>1/\<sigma>\<bar> * \<sigma>) x))"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1204  | 
by(rule normal_density_affine) auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1205  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1206  | 
then show "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ennreal (std_normal_density x))"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1207  | 
by (simp add: diff_divide_distrib[symmetric] field_simps)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1208  | 
next  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1209  | 
assume *: "distributed M lborel (\<lambda>x. (X x - \<mu>) / \<sigma>) (\<lambda>x. ennreal (std_normal_density x))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1210  | 
have "distributed M lborel (\<lambda>x. \<mu> + \<sigma> * ((X x - \<mu>) / \<sigma>)) (\<lambda>x. ennreal (normal_density \<mu> \<bar>\<sigma>\<bar> x))"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1211  | 
using normal_density_affine[OF *, of \<sigma> \<mu>] by simp  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1212  | 
then show "distributed M lborel X (\<lambda>x. ennreal (normal_density \<mu> \<sigma> x))" by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1213  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1214  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1215  | 
lemma conv_normal_density_zero_mean:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1216  | 
assumes [simp, arith]: "0 < \<sigma>" "0 < \<tau>"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1217  | 
shows "(\<lambda>x. \<integral>\<^sup>+y. ennreal (normal_density 0 \<sigma> (x - y) * normal_density 0 \<tau> y) \<partial>lborel) =  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1218  | 
normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2))" (is "?LHS = ?RHS")  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1219  | 
proof -  | 
| 63040 | 1220  | 
define \<sigma>' \<tau>' where "\<sigma>' = \<sigma>\<^sup>2" and "\<tau>' = \<tau>\<^sup>2"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1221  | 
then have [simp, arith]: "0 < \<sigma>'" "0 < \<tau>'"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1222  | 
by simp_all  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1223  | 
let ?\<sigma> = "sqrt ((\<sigma>' * \<tau>') / (\<sigma>' + \<tau>'))"  | 
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1224  | 
have sqrt: "(sqrt (2 * pi * (\<sigma>' + \<tau>')) * sqrt (2 * pi * (\<sigma>' * \<tau>') / (\<sigma>' + \<tau>'))) =  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1225  | 
(sqrt (2 * pi * \<sigma>') * sqrt (2 * pi * \<tau>'))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1226  | 
by (subst power_eq_iff_eq_base[symmetric, where n=2])  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1227  | 
(simp_all add: real_sqrt_mult[symmetric] power2_eq_square)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1228  | 
have "?LHS =  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1229  | 
(\<lambda>x. \<integral>\<^sup>+y. ennreal((normal_density 0 (sqrt (\<sigma>' + \<tau>')) x) * normal_density (\<tau>' * x / (\<sigma>' + \<tau>')) ?\<sigma> y) \<partial>lborel)"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1230  | 
apply (intro ext nn_integral_cong)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1231  | 
apply (simp add: normal_density_def \<sigma>'_def[symmetric] \<tau>'_def[symmetric] sqrt mult_exp_exp)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1232  | 
apply (simp add: divide_simps power2_eq_square)  | 
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70365 
diff
changeset
 | 
1233  | 
apply (simp add: algebra_simps)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1234  | 
done  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1235  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1236  | 
also have "... =  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1237  | 
(\<lambda>x. (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x) * \<integral>\<^sup>+y. ennreal( normal_density (\<tau>\<^sup>2* x / (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) ?\<sigma> y) \<partial>lborel)"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1238  | 
by (subst nn_integral_cmult[symmetric])  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1239  | 
(auto simp: \<sigma>'_def \<tau>'_def normal_density_def ennreal_mult'[symmetric])  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1240  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1241  | 
also have "... = (\<lambda>x. (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1242  | 
by (subst nn_integral_eq_integral) (auto simp: normal_density_nonneg)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1243  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1244  | 
finally show ?thesis by fast  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1245  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1246  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1247  | 
lemma conv_std_normal_density:  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1248  | 
"(\<lambda>x. \<integral>\<^sup>+y. ennreal (std_normal_density (x - y) * std_normal_density y) \<partial>lborel) =  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1249  | 
(normal_density 0 (sqrt 2))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1250  | 
by (subst conv_normal_density_zero_mean) simp_all  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1251  | 
|
| 64267 | 1252  | 
lemma (in prob_space) add_indep_normal:  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1253  | 
assumes indep: "indep_var borel X borel Y"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1254  | 
assumes pos_var[arith]: "0 < \<sigma>" "0 < \<tau>"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1255  | 
assumes normalX[simp]: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1256  | 
assumes normalY[simp]: "distributed M lborel Y (normal_density \<nu> \<tau>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1257  | 
shows "distributed M lborel (\<lambda>x. X x + Y x) (normal_density (\<mu> + \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1258  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1259  | 
have ind[simp]: "indep_var borel (\<lambda>x. - \<mu> + X x) borel (\<lambda>x. - \<nu> + Y x)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1260  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1261  | 
have "indep_var borel ( (\<lambda>x. -\<mu> + x) o X) borel ((\<lambda>x. - \<nu> + x) o Y)"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1262  | 
by (auto intro!: indep_var_compose assms)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1263  | 
then show ?thesis by (simp add: o_def)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1264  | 
qed  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1265  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1266  | 
have "distributed M lborel (\<lambda>x. -\<mu> + 1 * X x) (normal_density (- \<mu> + 1 * \<mu>) (\<bar>1\<bar> * \<sigma>))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1267  | 
by(rule normal_density_affine[OF normalX pos_var(1), of 1 "-\<mu>"]) simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1268  | 
then have 1[simp]: "distributed M lborel (\<lambda>x. - \<mu> + X x) (normal_density 0 \<sigma>)" by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1269  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1270  | 
have "distributed M lborel (\<lambda>x. -\<nu> + 1 * Y x) (normal_density (- \<nu> + 1 * \<nu>) (\<bar>1\<bar> * \<tau>))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1271  | 
by(rule normal_density_affine[OF normalY pos_var(2), of 1 "-\<nu>"]) simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1272  | 
then have 2[simp]: "distributed M lborel (\<lambda>x. - \<nu> + Y x) (normal_density 0 \<tau>)" by simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1273  | 
|
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1274  | 
have *: "distributed M lborel (\<lambda>x. (- \<mu> + X x) + (- \<nu> + Y x)) (\<lambda>x. ennreal (normal_density 0 (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1275  | 
using distributed_convolution[OF ind 1 2] conv_normal_density_zero_mean[OF pos_var]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1276  | 
by (simp add: ennreal_mult'[symmetric] normal_density_nonneg)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1277  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1278  | 
have "distributed M lborel (\<lambda>x. \<mu> + \<nu> + 1 * (- \<mu> + X x + (- \<nu> + Y x)))  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1279  | 
(\<lambda>x. ennreal (normal_density (\<mu> + \<nu> + 1 * 0) (\<bar>1\<bar> * sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)) x))"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1280  | 
by (rule normal_density_affine[OF *, of 1 "\<mu> + \<nu>"]) (auto simp: add_pos_pos)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1281  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1282  | 
then show ?thesis by auto  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1283  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1284  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1285  | 
lemma (in prob_space) diff_indep_normal:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1286  | 
assumes indep[simp]: "indep_var borel X borel Y"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1287  | 
assumes [simp, arith]: "0 < \<sigma>" "0 < \<tau>"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1288  | 
assumes normalX[simp]: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1289  | 
assumes normalY[simp]: "distributed M lborel Y (normal_density \<nu> \<tau>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1290  | 
shows "distributed M lborel (\<lambda>x. X x - Y x) (normal_density (\<mu> - \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1291  | 
proof -  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1292  | 
have "distributed M lborel (\<lambda>x. 0 + - 1 * Y x) (\<lambda>x. ennreal (normal_density (0 + - 1 * \<nu>) (\<bar>- 1\<bar> * \<tau>) x))"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1293  | 
by(rule normal_density_affine, auto)  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1294  | 
then have [simp]:"distributed M lborel (\<lambda>x. - Y x) (\<lambda>x. ennreal (normal_density (- \<nu>) \<tau> x))" by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1295  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1296  | 
have "distributed M lborel (\<lambda>x. X x + (- Y x)) (normal_density (\<mu> + - \<nu>) (sqrt (\<sigma>\<^sup>2 + \<tau>\<^sup>2)))"  | 
| 64267 | 1297  | 
apply (rule add_indep_normal)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1298  | 
apply (rule indep_var_compose[unfolded comp_def, of borel _ borel _ "\<lambda>x. x" _ "\<lambda>x. - x"])  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1299  | 
apply simp_all  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1300  | 
done  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1301  | 
then show ?thesis by simp  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1302  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1303  | 
|
| 64267 | 1304  | 
lemma (in prob_space) sum_indep_normal:  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1305  | 
  assumes "finite I" "I \<noteq> {}" "indep_vars (\<lambda>i. borel) X I"
 | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1306  | 
assumes "\<And>i. i \<in> I \<Longrightarrow> 0 < \<sigma> i"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1307  | 
assumes normal: "\<And>i. i \<in> I \<Longrightarrow> distributed M lborel (X i) (normal_density (\<mu> i) (\<sigma> i))"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1308  | 
shows "distributed M lborel (\<lambda>x. \<Sum>i\<in>I. X i x) (normal_density (\<Sum>i\<in>I. \<mu> i) (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2)))"  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1309  | 
using assms  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1310  | 
proof (induct I rule: finite_ne_induct)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1311  | 
case (singleton i) then show ?case by (simp add : abs_of_pos)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1312  | 
next  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1313  | 
case (insert i I)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1314  | 
then have 1: "distributed M lborel (\<lambda>x. (X i x) + (\<Sum>i\<in>I. X i x))  | 
| 64267 | 1315  | 
(normal_density (\<mu> i + sum \<mu> I) (sqrt ((\<sigma> i)\<^sup>2 + (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2))\<^sup>2)))"  | 
1316  | 
apply (intro add_indep_normal indep_vars_sum insert real_sqrt_gt_zero)  | 
|
1317  | 
apply (auto intro: indep_vars_subset intro!: sum_pos)  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1318  | 
apply fastforce  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1319  | 
done  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1320  | 
have 2: "(\<lambda>x. (X i x)+ (\<Sum>i\<in>I. X i x)) = (\<lambda>x. (\<Sum>j\<in>insert i I. X j x))"  | 
| 64267 | 1321  | 
"\<mu> i + sum \<mu> I = sum \<mu> (insert i I)"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1322  | 
using insert by auto  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1323  | 
|
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1324  | 
have 3: "(sqrt ((\<sigma> i)\<^sup>2 + (sqrt (\<Sum>i\<in>I. (\<sigma> i)\<^sup>2))\<^sup>2)) = (sqrt (\<Sum>i\<in>insert i I. (\<sigma> i)\<^sup>2))"  | 
| 64267 | 1325  | 
using insert by (simp add: sum_nonneg)  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1326  | 
|
| 
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61359 
diff
changeset
 | 
1327  | 
show ?case using 1 2 3 by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1328  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1329  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1330  | 
lemma (in prob_space) standard_normal_distributed_expectation:  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1331  | 
assumes D: "distributed M lborel X std_normal_density"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1332  | 
shows "expectation X = 0"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1333  | 
using integral_std_normal_moment_odd[of 0]  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1334  | 
distributed_integral[OF D, of "\<lambda>x. x", symmetric]  | 
| 
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1335  | 
by (auto simp: )  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1336  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1337  | 
lemma (in prob_space) normal_distributed_expectation:  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1338  | 
assumes \<sigma>[arith]: "0 < \<sigma>"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1339  | 
assumes D: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1340  | 
shows "expectation X = \<mu>"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1341  | 
using integral_normal_moment_nz_1[OF \<sigma>, of \<mu>] distributed_integral[OF D, of "\<lambda>x. x", symmetric]  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1342  | 
by (auto simp: field_simps)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1343  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1344  | 
lemma (in prob_space) normal_distributed_variance:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1345  | 
fixes a b :: real  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1346  | 
assumes [simp, arith]: "0 < \<sigma>"  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1347  | 
assumes D: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1348  | 
shows "variance X = \<sigma>\<^sup>2"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1349  | 
using integral_normal_moment_even[of \<sigma> \<mu> 1]  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1350  | 
by (subst distributed_integral[OF D, symmetric])  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1351  | 
(simp_all add: eval_nat_numeral normal_distributed_expectation[OF assms])  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1352  | 
|
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1353  | 
lemma (in prob_space) standard_normal_distributed_variance:  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1354  | 
"distributed M lborel X std_normal_density \<Longrightarrow> variance X = 1"  | 
| 
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1355  | 
using normal_distributed_variance[of 1 X 0] by simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1356  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1357  | 
lemma (in information_space) entropy_normal_density:  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1358  | 
assumes [arith]: "0 < \<sigma>"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1359  | 
assumes D: "distributed M lborel X (normal_density \<mu> \<sigma>)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1360  | 
shows "entropy b lborel X = log b (2 * pi * exp 1 * \<sigma>\<^sup>2) / 2"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1361  | 
proof -  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1362  | 
have "entropy b lborel X = - (\<integral> x. normal_density \<mu> \<sigma> x * log b (normal_density \<mu> \<sigma> x) \<partial>lborel)"  | 
| 
62975
 
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
 
hoelzl 
parents: 
62390 
diff
changeset
 | 
1363  | 
using D by (rule entropy_distr) simp  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1364  | 
also have "\<dots> = - (\<integral> x. normal_density \<mu> \<sigma> x * (- ln (2 * pi * \<sigma>\<^sup>2) - (x - \<mu>)\<^sup>2 / \<sigma>\<^sup>2) / (2 * ln b) \<partial>lborel)"  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63540 
diff
changeset
 | 
1365  | 
by (intro arg_cong[where f="uminus"] Bochner_Integration.integral_cong)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1366  | 
(auto simp: normal_density_def field_simps ln_mult log_def ln_div ln_sqrt)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1367  | 
also have "\<dots> = - (\<integral>x. - (normal_density \<mu> \<sigma> x * (ln (2 * pi * \<sigma>\<^sup>2)) + (normal_density \<mu> \<sigma> x * (x - \<mu>)\<^sup>2) / \<sigma>\<^sup>2) / (2 * ln b) \<partial>lborel)"  | 
| 
70817
 
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
 
haftmann 
parents: 
70365 
diff
changeset
 | 
1368  | 
by (intro arg_cong[where f="uminus"] Bochner_Integration.integral_cong) (auto simp: field_split_simps field_simps)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1369  | 
also have "\<dots> = (\<integral>x. normal_density \<mu> \<sigma> x * (ln (2 * pi * \<sigma>\<^sup>2)) + (normal_density \<mu> \<sigma> x * (x - \<mu>)\<^sup>2) / \<sigma>\<^sup>2 \<partial>lborel) / (2 * ln b)"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1370  | 
by (simp del: minus_add_distrib)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1371  | 
also have "\<dots> = (ln (2 * pi * \<sigma>\<^sup>2) + 1) / (2 * ln b)"  | 
| 
57254
 
d3d91422f408
lemmas about the moments of the normal distribution
 
hoelzl 
parents: 
57252 
diff
changeset
 | 
1372  | 
using integral_normal_moment_even[of \<sigma> \<mu> 1] by (simp add: integrable_normal_moment fact_numeral)  | 
| 
57252
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1373  | 
also have "\<dots> = log b (2 * pi * exp 1 * \<sigma>\<^sup>2) / 2"  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1374  | 
by (simp add: log_def field_simps ln_mult)  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1375  | 
finally show ?thesis .  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1376  | 
qed  | 
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1377  | 
|
| 
 
19b7ace1c5da
properties of normal distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
1378  | 
end  |