author | wenzelm |
Fri, 24 Sep 2021 22:23:26 +0200 | |
changeset 74362 | 0135a0c77b64 |
parent 70817 | dd675800469d |
permissions | -rw-r--r-- |
63627 | 1 |
(* Title: HOL/Analysis/Radon_Nikodym.thy |
42067 | 2 |
Author: Johannes Hölzl, TU München |
3 |
*) |
|
4 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
5 |
section \<open>Radon-Nikod{\'y}m Derivative\<close> |
42067 | 6 |
|
38656 | 7 |
theory Radon_Nikodym |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
8 |
imports Bochner_Integration |
38656 | 9 |
begin |
10 |
||
70136 | 11 |
definition\<^marker>\<open>tag important\<close> diff_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure" |
63329 | 12 |
where |
13 |
"diff_measure M N = measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)" |
|
47694 | 14 |
|
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:
61169
diff
changeset
|
15 |
lemma |
47694 | 16 |
shows space_diff_measure[simp]: "space (diff_measure M N) = space M" |
17 |
and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M" |
|
18 |
by (auto simp: diff_measure_def) |
|
19 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
20 |
lemma emeasure_diff_measure: |
47694 | 21 |
assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N" |
22 |
assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M" |
|
23 |
shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A") |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
24 |
unfolding diff_measure_def |
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
25 |
proof (rule emeasure_measure_of_sigma) |
47694 | 26 |
show "sigma_algebra (space M) (sets M)" .. |
27 |
show "positive (sets M) ?\<mu>" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
28 |
using pos by (simp add: positive_def) |
47694 | 29 |
show "countably_additive (sets M) ?\<mu>" |
30 |
proof (rule countably_additiveI) |
|
31 |
fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> sets M" and "disjoint_family A" |
|
32 |
then have suminf: |
|
33 |
"(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)" |
|
34 |
"(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)" |
|
35 |
by (simp_all add: suminf_emeasure sets_eq) |
|
36 |
with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) = |
|
37 |
(\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
38 |
using fin pos[of "A _"] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
39 |
by (intro ennreal_suminf_minus) |
63329 | 40 |
(auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure) |
47694 | 41 |
then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) = |
42 |
emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) " |
|
43 |
by (simp add: suminf) |
|
44 |
qed |
|
45 |
qed fact |
|
46 |
||
64911 | 47 |
text \<open>An equivalent characterization of sigma-finite spaces is the existence of integrable |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
48 |
positive functions (or, still equivalently, the existence of a probability measure which is in |
64911 | 49 |
the same measure class as the original measure).\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
50 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
51 |
proposition (in sigma_finite_measure) obtain_positive_integrable_function: |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
52 |
obtains f::"'a \<Rightarrow> real" where |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
53 |
"f \<in> borel_measurable M" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
54 |
"\<And>x. f x > 0" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
55 |
"\<And>x. f x \<le> 1" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
56 |
"integrable M f" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
57 |
proof - |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
58 |
obtain A :: "nat \<Rightarrow> 'a set" where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
59 |
using sigma_finite by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
60 |
then have [measurable]:"A n \<in> sets M" for n by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
61 |
define g where "g = (\<lambda>x. (\<Sum>n. (1/2)^(Suc n) * indicator (A n) x / (1+ measure M (A n))))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
62 |
have [measurable]: "g \<in> borel_measurable M" unfolding g_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
63 |
have *: "summable (\<lambda>n. (1/2)^(Suc n) * indicator (A n) x / (1+ measure M (A n)))" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
64 |
apply (rule summable_comparison_test'[of "\<lambda>n. (1/2)^(Suc n)" 0]) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
65 |
using power_half_series summable_def by (auto simp add: indicator_def divide_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
66 |
have "g x \<le> (\<Sum>n. (1/2)^(Suc n))" for x unfolding g_def |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
67 |
apply (rule suminf_le) using * power_half_series summable_def by (auto simp add: indicator_def divide_simps) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
68 |
then have g_le_1: "g x \<le> 1" for x using power_half_series sums_unique by fastforce |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
69 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
70 |
have g_pos: "g x > 0" if "x \<in> space M" for x |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
71 |
unfolding g_def proof (subst suminf_pos_iff[OF *[of x]], auto) |
64911 | 72 |
obtain i where "x \<in> A i" using \<open>(\<Union>i. A i) = space M\<close> \<open>x \<in> space M\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
73 |
then have "0 < (1 / 2) ^ Suc i * indicator (A i) x / (1 + Sigma_Algebra.measure M (A i))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
74 |
unfolding indicator_def apply (auto simp add: divide_simps) using measure_nonneg[of M "A i"] |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
75 |
by (auto, meson add_nonneg_nonneg linorder_not_le mult_nonneg_nonneg zero_le_numeral zero_le_one zero_le_power) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
76 |
then show "\<exists>i. 0 < (1 / 2) ^ i * indicator (A i) x / (2 + 2 * Sigma_Algebra.measure M (A i))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
77 |
by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
78 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
79 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
80 |
have "integrable M g" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
81 |
unfolding g_def proof (rule integrable_suminf) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
82 |
fix n |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
83 |
show "integrable M (\<lambda>x. (1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n)))" |
64911 | 84 |
using \<open>emeasure M (A n) \<noteq> \<infinity>\<close> |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
85 |
by (auto intro!: integrable_mult_right integrable_divide_zero integrable_real_indicator simp add: top.not_eq_extremum) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
86 |
next |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
87 |
show "AE x in M. summable (\<lambda>n. norm ((1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n))))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
88 |
using * by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
89 |
show "summable (\<lambda>n. (\<integral>x. norm ((1 / 2) ^ Suc n * indicator (A n) x / (1 + Sigma_Algebra.measure M (A n))) \<partial>M))" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
90 |
apply (rule summable_comparison_test'[of "\<lambda>n. (1/2)^(Suc n)" 0], auto) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
91 |
using power_half_series summable_def apply auto[1] |
70817
dd675800469d
dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents:
70136
diff
changeset
|
92 |
apply (auto simp add: field_split_simps) using measure_nonneg[of M] not_less by fastforce |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
93 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
94 |
|
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
95 |
define f where "f = (\<lambda>x. if x \<in> space M then g x else 1)" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
96 |
have "f x > 0" for x unfolding f_def using g_pos by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
97 |
moreover have "f x \<le> 1" for x unfolding f_def using g_le_1 by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
98 |
moreover have [measurable]: "f \<in> borel_measurable M" unfolding f_def by auto |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
99 |
moreover have "integrable M f" |
64911 | 100 |
apply (subst integrable_cong[of _ _ _ g]) unfolding f_def using \<open>integrable M g\<close> by auto |
64283
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
101 |
ultimately show "(\<And>f. f \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 < f x) \<Longrightarrow> (\<And>x. f x \<le> 1) \<Longrightarrow> integrable M f \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
102 |
by (meson that) |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
103 |
qed |
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents:
63627
diff
changeset
|
104 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
105 |
lemma (in sigma_finite_measure) Ex_finite_integrable_function: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
106 |
"\<exists>h\<in>borel_measurable M. integral\<^sup>N M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>)" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
107 |
proof - |
38656 | 108 |
obtain A :: "nat \<Rightarrow> 'a set" where |
50003 | 109 |
range[measurable]: "range A \<subseteq> sets M" and |
38656 | 110 |
space: "(\<Union>i. A i) = space M" and |
47694 | 111 |
measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and |
38656 | 112 |
disjoint: "disjoint_family A" |
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62083
diff
changeset
|
113 |
using sigma_finite_disjoint by blast |
47694 | 114 |
let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)" |
63329 | 115 |
have [measurable]: "\<And>i. A i \<in> sets M" |
116 |
using range by fastforce+ |
|
38656 | 117 |
have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)" |
118 |
proof |
|
47694 | 119 |
fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
120 |
using measure[of i] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
121 |
by (auto intro!: dense simp: ennreal_inverse_positive ennreal_mult_eq_top_iff power_eq_top_ennreal) |
38656 | 122 |
qed |
123 |
from choice[OF this] obtain n where n: "\<And>i. 0 < n i" |
|
47694 | 124 |
"\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
125 |
{ fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this |
46731 | 126 |
let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x" |
38656 | 127 |
show ?thesis |
128 |
proof (safe intro!: bexI[of _ ?h] del: notI) |
|
63329 | 129 |
have "integral\<^sup>N M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos |
56996 | 130 |
by (simp add: nn_integral_suminf nn_integral_cmult_indicator) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
131 |
also have "\<dots> \<le> (\<Sum>i. ennreal ((1/2)^Suc i))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
132 |
proof (intro suminf_le allI) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
133 |
fix N |
47694 | 134 |
have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
135 |
using n[of N] by (intro mult_right_mono) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
136 |
also have "\<dots> = (1/2)^Suc N * (inverse (emeasure M (A N)) * emeasure M (A N))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
137 |
using measure[of N] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
138 |
by (simp add: ennreal_inverse_power divide_ennreal_def ennreal_inverse_mult |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
139 |
power_eq_top_ennreal less_top[symmetric] mult_ac |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
140 |
del: power_Suc) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
141 |
also have "\<dots> \<le> inverse (ennreal 2) ^ Suc N" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
142 |
using measure[of N] |
63329 | 143 |
by (cases "emeasure M (A N)"; cases "emeasure M (A N) = 0") |
144 |
(auto simp: inverse_ennreal ennreal_mult[symmetric] divide_ennreal_def simp del: power_Suc) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
145 |
also have "\<dots> = ennreal (inverse 2 ^ Suc N)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
146 |
by (subst ennreal_power[symmetric], simp) (simp add: inverse_ennreal) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
147 |
finally show "n N * emeasure M (A N) \<le> ennreal ((1/2)^Suc N)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
148 |
by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
149 |
qed auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
150 |
also have "\<dots> < top" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
151 |
unfolding less_top[symmetric] |
63329 | 152 |
by (rule ennreal_suminf_neq_top) |
153 |
(auto simp: summable_geometric summable_Suc_iff simp del: power_Suc) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
154 |
finally show "integral\<^sup>N M ?h \<noteq> \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
155 |
by (auto simp: top_unique) |
38656 | 156 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
157 |
{ fix x assume "x \<in> space M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
158 |
then obtain i where "x \<in> A i" using space[symmetric] by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
159 |
with disjoint n have "?h x = n i" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
160 |
by (auto intro!: suminf_cmult_indicator intro: less_imp_le) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
161 |
then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by (auto simp: less_top[symmetric]) } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
162 |
note pos = this |
50003 | 163 |
qed measurable |
38656 | 164 |
qed |
165 |
||
69683 | 166 |
subsection "Absolutely continuous" |
40871 | 167 |
|
70136 | 168 |
definition\<^marker>\<open>tag important\<close> absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where |
47694 | 169 |
"absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N" |
170 |
||
171 |
lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M" |
|
172 |
unfolding absolutely_continuous_def by (auto simp: null_sets_count_space) |
|
38656 | 173 |
|
47694 | 174 |
lemma absolutely_continuousI_density: |
175 |
"f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)" |
|
176 |
by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in) |
|
177 |
||
178 |
lemma absolutely_continuousI_point_measure_finite: |
|
179 |
"(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)" |
|
180 |
unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff) |
|
181 |
||
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
182 |
lemma absolutely_continuousD: |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
183 |
"absolutely_continuous M N \<Longrightarrow> A \<in> sets M \<Longrightarrow> emeasure M A = 0 \<Longrightarrow> emeasure N A = 0" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
184 |
by (auto simp: absolutely_continuous_def null_sets_def) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
185 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
186 |
lemma absolutely_continuous_AE: |
47694 | 187 |
assumes sets_eq: "sets M' = sets M" |
188 |
and "absolutely_continuous M M'" "AE x in M. P x" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
189 |
shows "AE x in M'. P x" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
190 |
proof - |
61808 | 191 |
from \<open>AE x in M. P x\<close> obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N" |
47694 | 192 |
unfolding eventually_ae_filter by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
193 |
show "AE x in M'. P x" |
47694 | 194 |
proof (rule AE_I') |
195 |
show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp |
|
61808 | 196 |
from \<open>absolutely_continuous M M'\<close> show "N \<in> null_sets M'" |
47694 | 197 |
using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto |
40859 | 198 |
qed |
199 |
qed |
|
200 |
||
69683 | 201 |
subsection "Existence of the Radon-Nikodym derivative" |
40871 | 202 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
203 |
proposition |
69173
38beaaebe736
tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
64911
diff
changeset
|
204 |
(in finite_measure) Radon_Nikodym_finite_measure: |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
205 |
assumes "finite_measure N" and sets_eq[simp]: "sets N = sets M" |
47694 | 206 |
assumes "absolutely_continuous M N" |
63329 | 207 |
shows "\<exists>f \<in> borel_measurable M. density M f = N" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
208 |
proof - |
47694 | 209 |
interpret N: finite_measure N by fact |
63329 | 210 |
define G where "G = {g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A}" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
211 |
have [measurable_dest]: "f \<in> G \<Longrightarrow> f \<in> borel_measurable M" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
212 |
and G_D: "\<And>A. f \<in> G \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) \<le> N A" for f |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
213 |
by (auto simp: G_def) |
50003 | 214 |
note this[measurable_dest] |
38656 | 215 |
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto |
216 |
hence "G \<noteq> {}" by auto |
|
63329 | 217 |
{ fix f g assume f[measurable]: "f \<in> G" and g[measurable]: "g \<in> G" |
38656 | 218 |
have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def |
219 |
proof safe |
|
220 |
let ?A = "{x \<in> space M. f x \<le> g x}" |
|
221 |
have "?A \<in> sets M" using f g unfolding G_def by auto |
|
63329 | 222 |
fix A assume [measurable]: "A \<in> sets M" |
38656 | 223 |
have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A" |
61808 | 224 |
using sets.sets_into_space[OF \<open>A \<in> sets M\<close>] by auto |
38656 | 225 |
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x = |
226 |
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x" |
|
227 |
by (auto simp: indicator_def max_def) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
228 |
hence "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
229 |
(\<integral>\<^sup>+x. g x * indicator (?A \<inter> A) x \<partial>M) + |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
230 |
(\<integral>\<^sup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)" |
56996 | 231 |
by (auto cong: nn_integral_cong intro!: nn_integral_add) |
47694 | 232 |
also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)" |
63329 | 233 |
using f g unfolding G_def by (auto intro!: add_mono) |
47694 | 234 |
also have "\<dots> = N A" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
235 |
using union by (subst plus_emeasure) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
236 |
finally show "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" . |
63329 | 237 |
qed auto } |
38656 | 238 |
note max_in_G = this |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
239 |
{ fix f assume "incseq f" and f: "\<And>i. f i \<in> G" |
50003 | 240 |
then have [measurable]: "\<And>i. f i \<in> borel_measurable M" by (auto simp: G_def) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
241 |
have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def |
38656 | 242 |
proof safe |
50003 | 243 |
show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M" by measurable |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
244 |
next |
38656 | 245 |
fix A assume "A \<in> sets M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
246 |
have "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) = |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
247 |
(\<integral>\<^sup>+x. (SUP i. f i x * indicator A x) \<partial>M)" |
56996 | 248 |
by (intro nn_integral_cong) (simp split: split_indicator) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
249 |
also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i x * indicator A x \<partial>M))" |
61808 | 250 |
using \<open>incseq f\<close> f \<open>A \<in> sets M\<close> |
56996 | 251 |
by (intro nn_integral_monotone_convergence_SUP) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
252 |
(auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
253 |
finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
254 |
using f \<open>A \<in> sets M\<close> by (auto intro!: SUP_least simp: G_D) |
38656 | 255 |
qed } |
256 |
note SUP_in_G = this |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
257 |
let ?y = "SUP g \<in> G. integral\<^sup>N M g" |
47694 | 258 |
have y_le: "?y \<le> N (space M)" unfolding G_def |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
259 |
proof (safe intro!: SUP_least) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
260 |
fix g assume "\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A" |
56996 | 261 |
from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \<le> N (space M)" |
262 |
by (simp cong: nn_integral_cong) |
|
38656 | 263 |
qed |
74362 | 264 |
from ennreal_SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] |
265 |
obtain ys :: "nat \<Rightarrow> ennreal" |
|
266 |
where ys: "range ys \<subseteq> integral\<^sup>N M ` G \<and> Sup (integral\<^sup>N M ` G) = Sup (range ys)" |
|
267 |
by auto |
|
56996 | 268 |
then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n" |
38656 | 269 |
proof safe |
56996 | 270 |
fix n assume "range ys \<subseteq> integral\<^sup>N M ` G" |
271 |
hence "ys n \<in> integral\<^sup>N M ` G" by auto |
|
272 |
thus "\<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n" by auto |
|
38656 | 273 |
qed |
56996 | 274 |
from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^sup>N M (gs n) = ys n" by auto |
275 |
hence y_eq: "?y = (SUP i. integral\<^sup>N M (gs i))" using ys by auto |
|
46731 | 276 |
let ?g = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})" |
63040 | 277 |
define f where [abs_def]: "f x = (SUP i. ?g i x)" for x |
46731 | 278 |
let ?F = "\<lambda>A x. f x * indicator A x" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
279 |
have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto |
38656 | 280 |
{ fix i have "?g i \<in> G" |
281 |
proof (induct i) |
|
282 |
case 0 thus ?case by simp fact |
|
283 |
next |
|
284 |
case (Suc i) |
|
61808 | 285 |
with Suc gs_not_empty \<open>gs (Suc i) \<in> G\<close> show ?case |
38656 | 286 |
by (auto simp add: atMost_Suc intro!: max_in_G) |
287 |
qed } |
|
288 |
note g_in_G = this |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
289 |
have "incseq ?g" using gs_not_empty |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
290 |
by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
291 |
|
50003 | 292 |
from SUP_in_G[OF this g_in_G] have [measurable]: "f \<in> G" unfolding f_def . |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
293 |
then have [measurable]: "f \<in> borel_measurable M" unfolding G_def by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
294 |
|
56996 | 295 |
have "integral\<^sup>N M f = (SUP i. integral\<^sup>N M (?g i))" unfolding f_def |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
296 |
using g_in_G \<open>incseq ?g\<close> by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def) |
38656 | 297 |
also have "\<dots> = ?y" |
298 |
proof (rule antisym) |
|
56996 | 299 |
show "(SUP i. integral\<^sup>N M (?g i)) \<le> ?y" |
56166 | 300 |
using g_in_G by (auto intro: SUP_mono) |
56996 | 301 |
show "?y \<le> (SUP i. integral\<^sup>N M (?g i))" unfolding y_eq |
302 |
by (auto intro!: SUP_mono nn_integral_mono Max_ge) |
|
38656 | 303 |
qed |
56996 | 304 |
finally have int_f_eq_y: "integral\<^sup>N M f = ?y" . |
47694 | 305 |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
306 |
have upper_bound: "\<forall>A\<in>sets M. N A \<le> density M f A" |
38656 | 307 |
proof (rule ccontr) |
308 |
assume "\<not> ?thesis" |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
309 |
then obtain A where A[measurable]: "A \<in> sets M" and f_less_N: "density M f A < N A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
310 |
by (auto simp: not_le) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
311 |
then have pos_A: "0 < M A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
312 |
using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, OF A] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
313 |
by (auto simp: zero_less_iff_neq_zero) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
314 |
|
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
315 |
define b where "b = (N A - density M f A) / M A / 2" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
316 |
with f_less_N pos_A have "0 < b" "b \<noteq> top" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
317 |
by (auto intro!: diff_gr0_ennreal simp: zero_less_iff_neq_zero diff_eq_0_iff_ennreal ennreal_divide_eq_top_iff) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
318 |
|
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
319 |
let ?f = "\<lambda>x. f x + b" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
320 |
have "nn_integral M f \<noteq> top" |
64911 | 321 |
using \<open>f \<in> G\<close>[THEN G_D, of "space M"] by (auto simp: top_unique cong: nn_integral_cong) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
322 |
with \<open>b \<noteq> top\<close> interpret Mf: finite_measure "density M ?f" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
323 |
by (intro finite_measureI) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
324 |
(auto simp: field_simps mult_indicator_subset ennreal_mult_eq_top_iff |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
325 |
emeasure_density nn_integral_cmult_indicator nn_integral_add |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
326 |
cong: nn_integral_cong) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
327 |
|
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
328 |
from unsigned_Hahn_decomposition[of "density M ?f" N A] |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
329 |
obtain Y where [measurable]: "Y \<in> sets M" and [simp]: "Y \<subseteq> A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
330 |
and Y1: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> density M ?f C \<le> N C" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
331 |
and Y2: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> A \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> N C \<le> density M ?f C" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
332 |
by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
333 |
|
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
334 |
let ?f' = "\<lambda>x. f x + b * indicator Y x" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
335 |
have "M Y \<noteq> 0" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
336 |
proof |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
337 |
assume "M Y = 0" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
338 |
then have "N Y = 0" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
339 |
using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, of Y] by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
340 |
then have "N A = N (A - Y)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
341 |
by (subst emeasure_Diff) auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
342 |
also have "\<dots> \<le> density M ?f (A - Y)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
343 |
by (rule Y2) auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
344 |
also have "\<dots> \<le> density M ?f A - density M ?f Y" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
345 |
by (subst emeasure_Diff) auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
346 |
also have "\<dots> \<le> density M ?f A - 0" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
347 |
by (intro ennreal_minus_mono) auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
348 |
also have "density M ?f A = b * M A + density M f A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
349 |
by (simp add: emeasure_density field_simps mult_indicator_subset nn_integral_add nn_integral_cmult_indicator) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
350 |
also have "\<dots> < N A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
351 |
using f_less_N pos_A |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
352 |
by (cases "density M f A"; cases "M A"; cases "N A") |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
353 |
(auto simp: b_def ennreal_less_iff ennreal_minus divide_ennreal ennreal_numeral[symmetric] |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
354 |
ennreal_plus[symmetric] ennreal_mult[symmetric] field_simps |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
355 |
simp del: ennreal_numeral ennreal_plus) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
356 |
finally show False |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
357 |
by simp |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
358 |
qed |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
359 |
then have "nn_integral M f < nn_integral M ?f'" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
360 |
using \<open>0 < b\<close> \<open>nn_integral M f \<noteq> top\<close> |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
361 |
by (simp add: nn_integral_add nn_integral_cmult_indicator ennreal_zero_less_mult_iff zero_less_iff_neq_zero) |
38656 | 362 |
moreover |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
363 |
have "?f' \<in> G" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
364 |
unfolding G_def |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
365 |
proof safe |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
366 |
fix X assume [measurable]: "X \<in> sets M" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
367 |
have "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) = density M f (X - Y) + density M ?f (X \<inter> Y)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
368 |
by (auto simp add: emeasure_density nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_cong) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
369 |
also have "\<dots> \<le> N (X - Y) + N (X \<inter> Y)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
370 |
using G_D[OF \<open>f \<in> G\<close>] by (intro add_mono Y1) (auto simp: emeasure_density) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
371 |
also have "\<dots> = N X" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
372 |
by (subst plus_emeasure) (auto intro!: arg_cong2[where f=emeasure]) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
373 |
finally show "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) \<le> N X" . |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
374 |
qed simp |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
375 |
then have "nn_integral M ?f' \<le> ?y" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
376 |
by (rule SUP_upper) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
377 |
ultimately show False |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
378 |
by (simp add: int_f_eq_y) |
38656 | 379 |
qed |
380 |
show ?thesis |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
381 |
proof (intro bexI[of _ f] measure_eqI conjI antisym) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
382 |
fix A assume "A \<in> sets (density M f)" then show "emeasure (density M f) A \<le> emeasure N A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
383 |
by (auto simp: emeasure_density intro!: G_D[OF \<open>f \<in> G\<close>]) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
384 |
next |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
385 |
fix A assume A: "A \<in> sets (density M f)" then show "emeasure N A \<le> emeasure (density M f) A" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
386 |
using upper_bound by auto |
47694 | 387 |
qed auto |
38656 | 388 |
qed |
389 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
390 |
lemma (in finite_measure) split_space_into_finite_sets_and_rest: |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
391 |
assumes ac: "absolutely_continuous M N" and sets_eq[simp]: "sets N = sets M" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
392 |
shows "\<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> (\<forall>i. N (B i) \<noteq> \<infinity>) \<and> |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
393 |
(\<forall>A\<in>sets M. A \<inter> (\<Union>i. B i) = {} \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>))" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
394 |
proof - |
47694 | 395 |
let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}" |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
396 |
let ?a = "SUP Q\<in>?Q. emeasure M Q" |
47694 | 397 |
have "{} \<in> ?Q" by auto |
38656 | 398 |
then have Q_not_empty: "?Q \<noteq> {}" by blast |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
399 |
have "?a \<le> emeasure M (space M)" using sets.sets_into_space |
47694 | 400 |
by (auto intro!: SUP_least emeasure_mono) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
401 |
then have "?a \<noteq> \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
402 |
using finite_emeasure_space |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
403 |
by (auto simp: less_top[symmetric] top_unique simp del: SUP_eq_top_iff Sup_eq_top_iff) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
404 |
from ennreal_SUP_countable_SUP [OF Q_not_empty, of "emeasure M"] |
47694 | 405 |
obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" |
38656 | 406 |
by auto |
47694 | 407 |
then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto |
408 |
from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q" |
|
38656 | 409 |
by auto |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69173
diff
changeset
|
410 |
then have a_Lim: "?a = (SUP i. emeasure M (Q' i))" using a by simp |
46731 | 411 |
let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i" |
47694 | 412 |
have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)" |
413 |
proof (rule SUP_emeasure_incseq[of ?O]) |
|
414 |
show "range ?O \<subseteq> sets M" using Q' by auto |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
415 |
show "incseq ?O" by (fastforce intro!: incseq_SucI) |
38656 | 416 |
qed |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
417 |
have Q'_sets[measurable]: "\<And>i. Q' i \<in> sets M" using Q' by auto |
47694 | 418 |
have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto |
38656 | 419 |
then have O_in_G: "\<And>i. ?O i \<in> ?Q" |
420 |
proof (safe del: notI) |
|
47694 | 421 |
fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto |
422 |
then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))" |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
423 |
by (simp add: emeasure_subadditive_finite) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
424 |
also have "\<dots> < \<infinity>" using Q' by (simp add: less_top) |
47694 | 425 |
finally show "N (?O i) \<noteq> \<infinity>" by simp |
38656 | 426 |
qed auto |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
427 |
have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce |
47694 | 428 |
have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric] |
38656 | 429 |
proof (rule antisym) |
47694 | 430 |
show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim |
431 |
using Q' by (auto intro!: SUP_mono emeasure_mono) |
|
62343
24106dc44def
prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents:
62083
diff
changeset
|
432 |
show "(SUP i. emeasure M (?O i)) \<le> ?a" |
38656 | 433 |
proof (safe intro!: Sup_mono, unfold bex_simps) |
434 |
fix i |
|
52141
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
haftmann
parents:
51329
diff
changeset
|
435 |
have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto |
47694 | 436 |
then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and> |
52141
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
haftmann
parents:
51329
diff
changeset
|
437 |
emeasure M (\<Union>(Q' ` {..i})) \<le> emeasure M x" |
38656 | 438 |
using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) |
439 |
qed |
|
440 |
qed |
|
46731 | 441 |
let ?O_0 = "(\<Union>i. ?O i)" |
38656 | 442 |
have "?O_0 \<in> sets M" using Q' by auto |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
443 |
have "disjointed Q' i \<in> sets M" for i |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
444 |
using sets.range_disjointed_sets[of Q' M] using Q'_sets by (auto simp: subset_eq) |
38656 | 445 |
note Q_sets = this |
40859 | 446 |
show ?thesis |
447 |
proof (intro bexI exI conjI ballI impI allI) |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
448 |
show "disjoint_family (disjointed Q')" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
449 |
by (rule disjoint_family_disjointed) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
450 |
show "range (disjointed Q') \<subseteq> sets M" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
451 |
using Q'_sets by (intro sets.range_disjointed_sets) auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
452 |
{ fix A assume A: "A \<in> sets M" "A \<inter> (\<Union>i. disjointed Q' i) = {}" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
453 |
then have A1: "A \<inter> (\<Union>i. Q' i) = {}" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
454 |
unfolding UN_disjointed_eq by auto |
47694 | 455 |
show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>" |
40859 | 456 |
proof (rule disjCI, simp) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
457 |
assume *: "emeasure M A = 0 \<or> N A \<noteq> top" |
47694 | 458 |
show "emeasure M A = 0 \<and> N A = 0" |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
459 |
proof (cases "emeasure M A = 0") |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
460 |
case True |
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
461 |
with ac A have "N A = 0" |
40859 | 462 |
unfolding absolutely_continuous_def by auto |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
463 |
with True show ?thesis by simp |
40859 | 464 |
next |
53374
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
wenzelm
parents:
53015
diff
changeset
|
465 |
case False |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
466 |
with * have "N A \<noteq> \<infinity>" by auto |
47694 | 467 |
with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
468 |
using Q' A1 by (auto intro!: plus_emeasure simp: set_eq_iff) |
47694 | 469 |
also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))" |
470 |
proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified]) |
|
40859 | 471 |
show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M" |
61808 | 472 |
using \<open>N A \<noteq> \<infinity>\<close> O_sets A by auto |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
473 |
qed (fastforce intro!: incseq_SucI) |
40859 | 474 |
also have "\<dots> \<le> ?a" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
475 |
proof (safe intro!: SUP_least) |
40859 | 476 |
fix i have "?O i \<union> A \<in> ?Q" |
477 |
proof (safe del: notI) |
|
478 |
show "?O i \<union> A \<in> sets M" using O_sets A by auto |
|
47694 | 479 |
from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
480 |
using emeasure_subadditive[of "?O i" N A] A O_sets by auto |
47694 | 481 |
with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
482 |
using \<open>N A \<noteq> \<infinity>\<close> by (auto simp: top_unique) |
40859 | 483 |
qed |
47694 | 484 |
then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper) |
40859 | 485 |
qed |
47694 | 486 |
finally have "emeasure M A = 0" |
487 |
unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure) |
|
61808 | 488 |
with \<open>emeasure M A \<noteq> 0\<close> show ?thesis by auto |
40859 | 489 |
qed |
490 |
qed } |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
491 |
{ fix i |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
492 |
have "N (disjointed Q' i) \<le> N (Q' i)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
493 |
by (auto intro!: emeasure_mono simp: disjointed_def) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
494 |
then show "N (disjointed Q' i) \<noteq> \<infinity>" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
495 |
using Q'(2)[of i] by (auto simp: top_unique) } |
40859 | 496 |
qed |
497 |
qed |
|
498 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
499 |
proposition (in finite_measure) Radon_Nikodym_finite_measure_infinite: |
47694 | 500 |
assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M" |
63329 | 501 |
shows "\<exists>f\<in>borel_measurable M. density M f = N" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
502 |
proof - |
40859 | 503 |
from split_space_into_finite_sets_and_rest[OF assms] |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
504 |
obtain Q :: "nat \<Rightarrow> 'a set" |
40859 | 505 |
where Q: "disjoint_family Q" "range Q \<subseteq> sets M" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
506 |
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>" |
47694 | 507 |
and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force |
40859 | 508 |
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto |
47694 | 509 |
let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))" |
63329 | 510 |
have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). density (?M i) f = ?N i" |
47694 | 511 |
proof (intro allI finite_measure.Radon_Nikodym_finite_measure) |
38656 | 512 |
fix i |
47694 | 513 |
from Q show "finite_measure (?M i)" |
56996 | 514 |
by (auto intro!: finite_measureI cong: nn_integral_cong |
47694 | 515 |
simp add: emeasure_density subset_eq sets_eq) |
516 |
from Q have "emeasure (?N i) (space N) = emeasure N (Q i)" |
|
56996 | 517 |
by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: nn_integral_cong) |
47694 | 518 |
with Q_fin show "finite_measure (?N i)" |
519 |
by (auto intro!: finite_measureI) |
|
520 |
show "sets (?N i) = sets (?M i)" by (simp add: sets_eq) |
|
50003 | 521 |
have [measurable]: "\<And>A. A \<in> sets M \<Longrightarrow> A \<in> sets N" by (simp add: sets_eq) |
47694 | 522 |
show "absolutely_continuous (?M i) (?N i)" |
61808 | 523 |
using \<open>absolutely_continuous M N\<close> \<open>Q i \<in> sets M\<close> |
47694 | 524 |
by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq |
525 |
intro!: absolutely_continuous_AE[OF sets_eq]) |
|
38656 | 526 |
qed |
47694 | 527 |
from choice[OF this[unfolded Bex_def]] |
528 |
obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x" |
|
529 |
and f_density: "\<And>i. density (?M i) (f i) = ?N i" |
|
54776
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
immler
parents:
53374
diff
changeset
|
530 |
by force |
47694 | 531 |
{ fix A i assume A: "A \<in> sets M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
532 |
with Q borel have "(\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A" |
56996 | 533 |
by (auto simp add: emeasure_density nn_integral_density subset_eq |
534 |
intro!: nn_integral_cong split: split_indicator) |
|
47694 | 535 |
also have "\<dots> = emeasure N (Q i \<inter> A)" |
536 |
using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
537 |
finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. } |
47694 | 538 |
note integral_eq = this |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
539 |
let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator (space M - (\<Union>i. Q i)) x" |
38656 | 540 |
show ?thesis |
541 |
proof (safe intro!: bexI[of _ ?f]) |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
542 |
show "?f \<in> borel_measurable M" using borel Q_sets |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
543 |
by (auto intro!: measurable_If) |
47694 | 544 |
show "density M ?f = N" |
545 |
proof (rule measure_eqI) |
|
546 |
fix A assume "A \<in> sets (density M ?f)" |
|
547 |
then have "A \<in> sets M" by simp |
|
548 |
have Qi: "\<And>i. Q i \<in> sets M" using Q by auto |
|
549 |
have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M" |
|
550 |
"\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x" |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
551 |
using borel Qi \<open>A \<in> sets M\<close> by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
552 |
have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator ((space M - (\<Union>i. Q i)) \<inter> A) x \<partial>M)" |
56996 | 553 |
using borel by (intro nn_integral_cong) (auto simp: indicator_def) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
554 |
also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
555 |
using borel Qi \<open>A \<in> sets M\<close> |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
556 |
by (subst nn_integral_add) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
557 |
(auto simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
558 |
also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)" |
61808 | 559 |
by (subst integral_eq[OF \<open>A \<in> sets M\<close>], subst nn_integral_suminf) auto |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
560 |
finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)" . |
47694 | 561 |
moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)" |
61808 | 562 |
using Q Q_sets \<open>A \<in> sets M\<close> |
47694 | 563 |
by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
564 |
moreover |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
565 |
have "(space M - (\<Union>x. Q x)) \<inter> A \<inter> (\<Union>x. Q x) = {}" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
566 |
by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
567 |
then have "\<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A) = N ((space M - (\<Union>i. Q i)) \<inter> A)" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
568 |
using in_Q0[of "(space M - (\<Union>i. Q i)) \<inter> A"] \<open>A \<in> sets M\<close> Q by (auto simp: ennreal_top_mult) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
569 |
moreover have "(space M - (\<Union>i. Q i)) \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
570 |
using Q_sets \<open>A \<in> sets M\<close> by auto |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
571 |
moreover have "((\<Union>i. Q i) \<inter> A) \<union> ((space M - (\<Union>i. Q i)) \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> ((space M - (\<Union>i. Q i)) \<inter> A) = {}" |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
572 |
using \<open>A \<in> sets M\<close> sets.sets_into_space by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
573 |
ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
574 |
using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "(space M - (\<Union>i. Q i)) \<inter> A"] by (simp add: sets_eq) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
575 |
with \<open>A \<in> sets M\<close> borel Q show "emeasure (density M ?f) A = N A" |
50003 | 576 |
by (auto simp: subset_eq emeasure_density) |
47694 | 577 |
qed (simp add: sets_eq) |
38656 | 578 |
qed |
579 |
qed |
|
580 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
581 |
theorem (in sigma_finite_measure) Radon_Nikodym: |
47694 | 582 |
assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M" |
63329 | 583 |
shows "\<exists>f \<in> borel_measurable M. density M f = N" |
38656 | 584 |
proof - |
585 |
from Ex_finite_integrable_function |
|
56996 | 586 |
obtain h where finite: "integral\<^sup>N M h \<noteq> \<infinity>" and |
38656 | 587 |
borel: "h \<in> borel_measurable M" and |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
588 |
nn: "\<And>x. 0 \<le> h x" and |
38656 | 589 |
pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
590 |
"\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
591 |
let ?T = "\<lambda>A. (\<integral>\<^sup>+x. h x * indicator A x \<partial>M)" |
47694 | 592 |
let ?MT = "density M h" |
593 |
from borel finite nn interpret T: finite_measure ?MT |
|
56996 | 594 |
by (auto intro!: finite_measureI cong: nn_integral_cong simp: emeasure_density) |
47694 | 595 |
have "absolutely_continuous ?MT N" "sets N = sets ?MT" |
596 |
proof (unfold absolutely_continuous_def, safe) |
|
597 |
fix A assume "A \<in> null_sets ?MT" |
|
598 |
with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0" |
|
599 |
by (auto simp add: null_sets_density_iff) |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
600 |
with pos sets.sets_into_space have "AE x in M. x \<notin> A" |
61810 | 601 |
by (elim eventually_mono) (auto simp: not_le[symmetric]) |
47694 | 602 |
then have "A \<in> null_sets M" |
61808 | 603 |
using \<open>A \<in> sets M\<close> by (simp add: AE_iff_null_sets) |
47694 | 604 |
with ac show "A \<in> null_sets N" |
605 |
by (auto simp: absolutely_continuous_def) |
|
606 |
qed (auto simp add: sets_eq) |
|
607 |
from T.Radon_Nikodym_finite_measure_infinite[OF this] |
|
63329 | 608 |
obtain f where f_borel: "f \<in> borel_measurable M" "density ?MT f = N" by auto |
47694 | 609 |
with nn borel show ?thesis |
610 |
by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq) |
|
38656 | 611 |
qed |
612 |
||
69683 | 613 |
subsection \<open>Uniqueness of densities\<close> |
40859 | 614 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
615 |
lemma finite_density_unique: |
40859 | 616 |
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
47694 | 617 |
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x" |
56996 | 618 |
and fin: "integral\<^sup>N M f \<noteq> \<infinity>" |
49785 | 619 |
shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
620 |
proof (intro iffI ballI) |
47694 | 621 |
fix A assume eq: "AE x in M. f x = g x" |
49785 | 622 |
with borel show "density M f = density M g" |
623 |
by (auto intro: density_cong) |
|
40859 | 624 |
next |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
625 |
let ?P = "\<lambda>f A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M" |
49785 | 626 |
assume "density M f = density M g" |
627 |
with borel have eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" |
|
628 |
by (simp add: emeasure_density[symmetric]) |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
629 |
from this[THEN bspec, OF sets.top] fin |
56996 | 630 |
have g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" by (simp cong: nn_integral_cong) |
40859 | 631 |
{ fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
47694 | 632 |
and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x" |
56996 | 633 |
and g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" |
40859 | 634 |
let ?N = "{x\<in>space M. g x < f x}" |
635 |
have N: "?N \<in> sets M" using borel by simp |
|
56996 | 636 |
have "?P g ?N \<le> integral\<^sup>N M g" using pos |
637 |
by (intro nn_integral_mono_AE) (auto split: split_indicator) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
638 |
then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by (auto simp: top_unique) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
639 |
have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)" |
56996 | 640 |
by (auto intro!: nn_integral_cong simp: indicator_def) |
40859 | 641 |
also have "\<dots> = ?P f ?N - ?P g ?N" |
56996 | 642 |
proof (rule nn_integral_diff) |
40859 | 643 |
show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M" |
644 |
using borel N by auto |
|
47694 | 645 |
show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
646 |
using pos by (auto split: split_indicator) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
647 |
qed fact |
40859 | 648 |
also have "\<dots> = 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
649 |
unfolding eq[THEN bspec, OF N] using Pg_fin by auto |
47694 | 650 |
finally have "AE x in M. f x \<le> g x" |
56996 | 651 |
using pos borel nn_integral_PInf_AE[OF borel(2) g_fin] |
652 |
by (subst (asm) nn_integral_0_iff_AE) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
653 |
(auto split: split_indicator simp: not_less ennreal_minus_eq_0) } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
654 |
from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq |
47694 | 655 |
show "AE x in M. f x = g x" by auto |
40859 | 656 |
qed |
657 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
658 |
lemma (in finite_measure) density_unique_finite_measure: |
40859 | 659 |
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M" |
47694 | 660 |
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
661 |
assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. f' x * indicator A x \<partial>M)" |
40859 | 662 |
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A") |
47694 | 663 |
shows "AE x in M. f x = f' x" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
664 |
proof - |
47694 | 665 |
let ?D = "\<lambda>f. density M f" |
666 |
let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A" |
|
46731 | 667 |
let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x" |
47694 | 668 |
|
669 |
have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M" |
|
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:
61169
diff
changeset
|
670 |
using borel by (auto intro!: absolutely_continuousI_density) |
47694 | 671 |
from split_space_into_finite_sets_and_rest[OF this] |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
672 |
obtain Q :: "nat \<Rightarrow> 'a set" |
40859 | 673 |
where Q: "disjoint_family Q" "range Q \<subseteq> sets M" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
674 |
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>" |
47694 | 675 |
and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
676 |
with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>" |
47694 | 677 |
and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq) |
678 |
||
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
679 |
from Q have Q_sets[measurable]: "\<And>i. Q i \<in> sets M" by auto |
47694 | 680 |
let ?D = "{x\<in>space M. f x \<noteq> f' x}" |
681 |
have "?D \<in> sets M" using borel by auto |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
682 |
have *: "\<And>i x A. \<And>y::ennreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x" |
40859 | 683 |
unfolding indicator_def by auto |
47694 | 684 |
have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos |
40859 | 685 |
by (intro finite_density_unique[THEN iffD1] allI) |
50003 | 686 |
(auto intro!: f measure_eqI simp: emeasure_density * subset_eq) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
687 |
moreover have "AE x in M. ?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x" |
40859 | 688 |
proof (rule AE_I') |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
689 |
{ fix f :: "'a \<Rightarrow> ennreal" assume borel: "f \<in> borel_measurable M" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
690 |
and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
691 |
let ?A = "\<lambda>i. (space M - (\<Union>i. Q i)) \<inter> {x \<in> space M. f x < (i::nat)}" |
47694 | 692 |
have "(\<Union>i. ?A i) \<in> null_sets M" |
40859 | 693 |
proof (rule null_sets_UN) |
43923 | 694 |
fix i ::nat have "?A i \<in> sets M" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
695 |
using borel by auto |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
696 |
have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ennreal) * indicator (?A i) x \<partial>M)" |
61808 | 697 |
unfolding eq[OF \<open>?A i \<in> sets M\<close>] |
56996 | 698 |
by (auto intro!: nn_integral_mono simp: indicator_def) |
47694 | 699 |
also have "\<dots> = i * emeasure M (?A i)" |
61808 | 700 |
using \<open>?A i \<in> sets M\<close> by (auto intro!: nn_integral_cmult_indicator) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
701 |
also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by (auto simp: ennreal_mult_less_top of_nat_less_top) |
47694 | 702 |
finally have "?N (?A i) \<noteq> \<infinity>" by simp |
61808 | 703 |
then show "?A i \<in> null_sets M" using in_Q0[OF \<open>?A i \<in> sets M\<close>] \<open>?A i \<in> sets M\<close> by auto |
40859 | 704 |
qed |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
705 |
also have "(\<Union>i. ?A i) = (space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
706 |
by (auto simp: ennreal_Ex_less_of_nat less_top[symmetric]) |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
707 |
finally have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp } |
40859 | 708 |
from this[OF borel(1) refl] this[OF borel(2) f] |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
709 |
have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
710 |
then show "((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets M" by (rule null_sets.Un) |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
711 |
show "{x \<in> space M. ?f (space M - (\<Union>i. Q i)) x \<noteq> ?f' (space M - (\<Union>i. Q i)) x} \<subseteq> |
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
712 |
((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def) |
40859 | 713 |
qed |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
714 |
moreover have "AE x in M. (?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow> |
40859 | 715 |
?f (space M) x = ?f' (space M) x" |
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
716 |
by (auto simp: indicator_def) |
47694 | 717 |
ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x" |
718 |
unfolding AE_all_countable[symmetric] |
|
63330
8d591640c3bd
Probability: introduce Hahn decomposition; use it to clean up Radon_Nikodym
hoelzl
parents:
63329
diff
changeset
|
719 |
by eventually_elim (auto split: if_split_asm simp: indicator_def) |
47694 | 720 |
then show "AE x in M. f x = f' x" by auto |
40859 | 721 |
qed |
722 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
723 |
proposition (in sigma_finite_measure) density_unique: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
724 |
assumes f: "f \<in> borel_measurable M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
725 |
assumes f': "f' \<in> borel_measurable M" |
47694 | 726 |
assumes density_eq: "density M f = density M f'" |
727 |
shows "AE x in M. f x = f' x" |
|
40859 | 728 |
proof - |
729 |
obtain h where h_borel: "h \<in> borel_measurable M" |
|
56996 | 730 |
and fin: "integral\<^sup>N M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x" |
40859 | 731 |
using Ex_finite_integrable_function by auto |
47694 | 732 |
then have h_nn: "AE x in M. 0 \<le> h x" by auto |
733 |
let ?H = "density M h" |
|
734 |
interpret h: finite_measure ?H |
|
735 |
using fin h_borel pos |
|
56996 | 736 |
by (intro finite_measureI) (simp cong: nn_integral_cong emeasure_density add: fin) |
47694 | 737 |
let ?fM = "density M f" |
738 |
let ?f'M = "density M f'" |
|
40859 | 739 |
{ fix A assume "A \<in> sets M" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
740 |
then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A" |
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
741 |
using pos(1) sets.sets_into_space by (force simp: indicator_def) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
742 |
then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M" |
61808 | 743 |
using h_borel \<open>A \<in> sets M\<close> h_nn by (subst nn_integral_0_iff) auto } |
40859 | 744 |
note h_null_sets = this |
745 |
{ fix A assume "A \<in> sets M" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
746 |
have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)" |
61808 | 747 |
using \<open>A \<in> sets M\<close> h_borel h_nn f f' |
56996 | 748 |
by (intro nn_integral_density[symmetric]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
749 |
also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)" |
47694 | 750 |
by (simp_all add: density_eq) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
751 |
also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h x * indicator A x) \<partial>M)" |
61808 | 752 |
using \<open>A \<in> sets M\<close> h_borel h_nn f f' |
56996 | 753 |
by (intro nn_integral_density) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
754 |
finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
755 |
by (simp add: ac_simps) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
756 |
then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)" |
61808 | 757 |
using \<open>A \<in> sets M\<close> h_borel h_nn f f' |
56996 | 758 |
by (subst (asm) (1 2) nn_integral_density[symmetric]) auto } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
759 |
then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
760 |
by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
761 |
with AE_space[of M] pos show "AE x in M. f x = f' x" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
762 |
unfolding AE_density[OF h_borel] by auto |
40859 | 763 |
qed |
764 |
||
47694 | 765 |
lemma (in sigma_finite_measure) density_unique_iff: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
766 |
assumes f: "f \<in> borel_measurable M" and f': "f' \<in> borel_measurable M" |
47694 | 767 |
shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)" |
768 |
using density_unique[OF assms] density_cong[OF f f'] by auto |
|
769 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
770 |
lemma sigma_finite_density_unique: |
49785 | 771 |
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
772 |
and fin: "sigma_finite_measure (density M f)" |
|
773 |
shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)" |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
774 |
proof |
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:
61169
diff
changeset
|
775 |
assume "AE x in M. f x = g x" with borel show "density M f = density M g" |
49785 | 776 |
by (auto intro: density_cong) |
777 |
next |
|
778 |
assume eq: "density M f = density M g" |
|
61605 | 779 |
interpret f: sigma_finite_measure "density M f" by fact |
74362 | 780 |
from f.sigma_finite_incseq obtain A where cover: "range A \<subseteq> sets (density M f)" |
781 |
"\<Union> (range A) = space (density M f)" |
|
782 |
"\<And>i. emeasure (density M f) (A i) \<noteq> \<infinity>" |
|
783 |
"incseq A" |
|
784 |
by auto |
|
49785 | 785 |
have "AE x in M. \<forall>i. x \<in> A i \<longrightarrow> f x = g x" |
786 |
unfolding AE_all_countable |
|
787 |
proof |
|
788 |
fix i |
|
789 |
have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))" |
|
790 |
unfolding eq .. |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
791 |
moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>" |
49785 | 792 |
using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq) |
793 |
ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
794 |
using borel cover(1) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
795 |
by (intro finite_density_unique[THEN iffD1]) (auto simp: density_density_eq subset_eq) |
49785 | 796 |
then show "AE x in M. x \<in> A i \<longrightarrow> f x = g x" |
797 |
by auto |
|
798 |
qed |
|
799 |
with AE_space show "AE x in M. f x = g x" |
|
800 |
apply eventually_elim |
|
801 |
using cover(2)[symmetric] |
|
802 |
apply auto |
|
803 |
done |
|
804 |
qed |
|
805 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
806 |
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite': |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
807 |
assumes f: "f \<in> borel_measurable M" |
47694 | 808 |
shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)" |
809 |
(is "sigma_finite_measure ?N \<longleftrightarrow> _") |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
810 |
proof |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
811 |
assume "sigma_finite_measure ?N" |
47694 | 812 |
then interpret N: sigma_finite_measure ?N . |
813 |
from N.Ex_finite_integrable_function obtain h where |
|
56996 | 814 |
h: "h \<in> borel_measurable M" "integral\<^sup>N ?N h \<noteq> \<infinity>" and |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
815 |
fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
816 |
by auto |
47694 | 817 |
have "AE x in M. f x * h x \<noteq> \<infinity>" |
40859 | 818 |
proof (rule AE_I') |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
819 |
have "integral\<^sup>N ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
820 |
using f h by (auto intro!: nn_integral_density) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
821 |
then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>" |
40859 | 822 |
using h(2) by simp |
47694 | 823 |
then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
824 |
using f h(1) by (auto intro!: nn_integral_PInf[unfolded infinity_ennreal_def] borel_measurable_vimage) |
40859 | 825 |
qed auto |
47694 | 826 |
then show "AE x in M. f x \<noteq> \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
827 |
using fin by (auto elim!: AE_Ball_mp simp: less_top ennreal_mult_less_top) |
40859 | 828 |
next |
47694 | 829 |
assume AE: "AE x in M. f x \<noteq> \<infinity>" |
74362 | 830 |
from sigma_finite obtain Q :: "nat \<Rightarrow> 'a set" |
831 |
where Q: "range Q \<subseteq> sets M" "\<Union> (range Q) = space M" "\<And>i. emeasure M (Q i) \<noteq> \<infinity>" |
|
832 |
by auto |
|
63040 | 833 |
define A where "A i = |
834 |
f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M" for i |
|
40859 | 835 |
{ fix i j have "A i \<inter> Q j \<in> sets M" |
836 |
unfolding A_def using f Q |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
837 |
apply (rule_tac sets.Int) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
838 |
by (cases i) (auto intro: measurable_sets[OF f(1)]) } |
40859 | 839 |
note A_in_sets = this |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
840 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
841 |
show "sigma_finite_measure ?N" |
61169 | 842 |
proof (standard, intro exI conjI ballI) |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
843 |
show "countable (range (\<lambda>(i, j). A i \<inter> Q j))" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
844 |
by auto |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
845 |
show "range (\<lambda>(i, j). A i \<inter> Q j) \<subseteq> sets (density M f)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
846 |
using A_in_sets by auto |
40859 | 847 |
next |
69745 | 848 |
have "\<Union>(range (\<lambda>(i, j). A i \<inter> Q j)) = (\<Union>i j. A i \<inter> Q j)" |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
849 |
by auto |
40859 | 850 |
also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto |
851 |
also have "(\<Union>i. A i) = space M" |
|
852 |
proof safe |
|
853 |
fix x assume x: "x \<in> space M" |
|
854 |
show "x \<in> (\<Union>i. A i)" |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
855 |
proof (cases "f x" rule: ennreal_cases) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
856 |
case top with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0]) |
40859 | 857 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
858 |
case (real r) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
859 |
with ennreal_Ex_less_of_nat[of "f x"] obtain n :: nat where "f x < n" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
860 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
861 |
also have "n < (Suc n :: ennreal)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
862 |
by simp |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
863 |
finally show ?thesis |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
864 |
using x real by (auto simp: A_def ennreal_of_nat_eq_real_of_nat intro!: exI[of _ "Suc n"]) |
40859 | 865 |
qed |
866 |
qed (auto simp: A_def) |
|
69745 | 867 |
finally show "\<Union>(range (\<lambda>(i, j). A i \<inter> Q j)) = space ?N" by simp |
40859 | 868 |
next |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
869 |
fix X assume "X \<in> range (\<lambda>(i, j). A i \<inter> Q j)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
870 |
then obtain i j where [simp]:"X = A i \<inter> Q j" by auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
871 |
have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>" |
40859 | 872 |
proof (cases i) |
873 |
case 0 |
|
47694 | 874 |
have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0" |
61808 | 875 |
using AE by (auto simp: A_def \<open>i = 0\<close>) |
56996 | 876 |
from nn_integral_cong_AE[OF this] show ?thesis by simp |
40859 | 877 |
next |
878 |
case (Suc n) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
879 |
then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le> |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
880 |
(\<integral>\<^sup>+x. (Suc n :: ennreal) * indicator (Q j) x \<partial>M)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
881 |
by (auto intro!: nn_integral_mono simp: indicator_def A_def ennreal_of_nat_eq_real_of_nat) |
47694 | 882 |
also have "\<dots> = Suc n * emeasure M (Q j)" |
56996 | 883 |
using Q by (auto intro!: nn_integral_cmult_indicator) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
884 |
also have "\<dots> < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
885 |
using Q by (auto simp: ennreal_mult_less_top less_top of_nat_less_top) |
40859 | 886 |
finally show ?thesis by simp |
887 |
qed |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
888 |
then show "emeasure ?N X \<noteq> \<infinity>" |
47694 | 889 |
using A_in_sets Q f by (auto simp: emeasure_density) |
40859 | 890 |
qed |
891 |
qed |
|
892 |
||
49778
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
hoelzl
parents:
47694
diff
changeset
|
893 |
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: |
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
hoelzl
parents:
47694
diff
changeset
|
894 |
"f \<in> borel_measurable M \<Longrightarrow> sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
895 |
by (subst sigma_finite_iff_density_finite') |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
896 |
(auto simp: max_def intro!: measurable_If) |
49778
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
hoelzl
parents:
47694
diff
changeset
|
897 |
|
69683 | 898 |
subsection \<open>Radon-Nikodym derivative\<close> |
38656 | 899 |
|
70136 | 900 |
definition\<^marker>\<open>tag important\<close> RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
901 |
"RN_deriv M N = |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
902 |
(if \<exists>f. f \<in> borel_measurable M \<and> density M f = N |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
903 |
then SOME f. f \<in> borel_measurable M \<and> density M f = N |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
904 |
else (\<lambda>_. 0))" |
38656 | 905 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
906 |
lemma RN_derivI: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
907 |
assumes "f \<in> borel_measurable M" "density M f = N" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
908 |
shows "density M (RN_deriv M N) = N" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
909 |
proof - |
63540 | 910 |
have *: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
911 |
using assms by auto |
63540 | 912 |
then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
913 |
by (rule someI2_ex) auto |
63540 | 914 |
with * show ?thesis |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
915 |
by (auto simp: RN_deriv_def) |
40859 | 916 |
qed |
917 |
||
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
918 |
lemma borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M" |
38656 | 919 |
proof - |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
920 |
{ assume ex: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
921 |
have 1: "(SOME f. f \<in> borel_measurable M \<and> density M f = N) \<in> borel_measurable M" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
922 |
using ex by (rule someI2_ex) auto } |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
923 |
from this show ?thesis |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
924 |
by (auto simp: RN_deriv_def) |
38656 | 925 |
qed |
926 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
927 |
lemma density_RN_deriv_density: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
928 |
assumes f: "f \<in> borel_measurable M" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
929 |
shows "density M (RN_deriv M (density M f)) = density M f" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
930 |
by (rule RN_derivI[OF f]) simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
931 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
932 |
lemma (in sigma_finite_measure) density_RN_deriv: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
933 |
"absolutely_continuous M N \<Longrightarrow> sets N = sets M \<Longrightarrow> density M (RN_deriv M N) = N" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
934 |
by (metis RN_derivI Radon_Nikodym) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
935 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
936 |
lemma (in sigma_finite_measure) RN_deriv_nn_integral: |
47694 | 937 |
assumes N: "absolutely_continuous M N" "sets N = sets M" |
40859 | 938 |
and f: "f \<in> borel_measurable M" |
56996 | 939 |
shows "integral\<^sup>N N f = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
940 |
proof - |
56996 | 941 |
have "integral\<^sup>N N f = integral\<^sup>N (density M (RN_deriv M N)) f" |
47694 | 942 |
using N by (simp add: density_RN_deriv) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
943 |
also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
944 |
using f by (simp add: nn_integral_density) |
47694 | 945 |
finally show ?thesis by simp |
40859 | 946 |
qed |
947 |
||
47694 | 948 |
lemma (in sigma_finite_measure) RN_deriv_unique: |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
949 |
assumes f: "f \<in> borel_measurable M" |
47694 | 950 |
and eq: "density M f = N" |
951 |
shows "AE x in M. f x = RN_deriv M N x" |
|
49785 | 952 |
unfolding eq[symmetric] |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
953 |
by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
954 |
density_RN_deriv_density[symmetric]) |
49785 | 955 |
|
956 |
lemma RN_deriv_unique_sigma_finite: |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
957 |
assumes f: "f \<in> borel_measurable M" |
49785 | 958 |
and eq: "density M f = N" and fin: "sigma_finite_measure N" |
959 |
shows "AE x in M. f x = RN_deriv M N x" |
|
960 |
using fin unfolding eq[symmetric] |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
961 |
by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
962 |
density_RN_deriv_density[symmetric]) |
47694 | 963 |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
964 |
lemma (in sigma_finite_measure) RN_deriv_distr: |
47694 | 965 |
fixes T :: "'a \<Rightarrow> 'b" |
966 |
assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M" |
|
967 |
and inv: "\<forall>x\<in>space M. T' (T x) = x" |
|
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
968 |
and ac[simp]: "absolutely_continuous (distr M M' T) (distr N M' T)" |
47694 | 969 |
and N: "sets N = sets M" |
970 |
shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x" |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
971 |
proof (rule RN_deriv_unique) |
47694 | 972 |
have [simp]: "sets N = sets M" by fact |
973 |
note sets_eq_imp_space_eq[OF N, simp] |
|
974 |
have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def) |
|
975 |
{ fix A assume "A \<in> sets M" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
976 |
with inv T T' sets.sets_into_space[OF this] |
47694 | 977 |
have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A" |
978 |
by (auto simp: measurable_def) } |
|
979 |
note eq = this[simp] |
|
980 |
{ fix A assume "A \<in> sets M" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
981 |
with inv T T' sets.sets_into_space[OF this] |
47694 | 982 |
have "(T' \<circ> T) -` A \<inter> space M = A" |
983 |
by (auto simp: measurable_def) } |
|
984 |
note eq2 = this[simp] |
|
985 |
let ?M' = "distr M M' T" and ?N' = "distr N M' T" |
|
986 |
interpret M': sigma_finite_measure ?M' |
|
41832 | 987 |
proof |
74362 | 988 |
from sigma_finite_countable obtain F |
989 |
where F: "countable F \<and> F \<subseteq> sets M \<and> \<Union> F = space M \<and> (\<forall>a\<in>F. emeasure M a \<noteq> \<infinity>)" .. |
|
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
990 |
show "\<exists>A. countable A \<and> A \<subseteq> sets (distr M M' T) \<and> \<Union>A = space (distr M M' T) \<and> (\<forall>a\<in>A. emeasure (distr M M' T) a \<noteq> \<infinity>)" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
991 |
proof (intro exI conjI ballI) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
992 |
show *: "(\<lambda>A. T' -` A \<inter> space ?M') ` F \<subseteq> sets ?M'" |
47694 | 993 |
using F T' by (auto simp: measurable_def) |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
994 |
show "\<Union>((\<lambda>A. T' -` A \<inter> space ?M')`F) = space ?M'" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
995 |
using F T'[THEN measurable_space] by (auto simp: set_eq_iff) |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
996 |
next |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
997 |
fix X assume "X \<in> (\<lambda>A. T' -` A \<inter> space ?M')`F" |
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
998 |
then obtain A where [simp]: "X = T' -` A \<inter> space ?M'" and "A \<in> F" by auto |
61808 | 999 |
have "X \<in> sets M'" using F T' \<open>A\<in>F\<close> by auto |
41832 | 1000 |
moreover |
61808 | 1001 |
have Fi: "A \<in> sets M" using F \<open>A\<in>F\<close> by auto |
57447
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
hoelzl
parents:
56996
diff
changeset
|
1002 |
ultimately show "emeasure ?M' X \<noteq> \<infinity>" |
61808 | 1003 |
using F T T' \<open>A\<in>F\<close> by (simp add: emeasure_distr) |
74362 | 1004 |
qed (use F in auto) |
41832 | 1005 |
qed |
47694 | 1006 |
have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M" |
50021
d96a3f468203
add support for function application to measurability prover
hoelzl
parents:
50003
diff
changeset
|
1007 |
using T ac by measurable |
47694 | 1008 |
then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M" |
41832 | 1009 |
by (simp add: comp_def) |
47694 | 1010 |
|
1011 |
have "N = distr N M (T' \<circ> T)" |
|
1012 |
by (subst measure_of_of_measure[of N, symmetric]) |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50021
diff
changeset
|
1013 |
(auto simp add: distr_def sets.sigma_sets_eq intro!: measure_of_eq sets.space_closed) |
47694 | 1014 |
also have "\<dots> = distr (distr N M' T) M T'" |
1015 |
using T T' by (simp add: distr_distr) |
|
1016 |
also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'" |
|
1017 |
using ac by (simp add: M'.density_RN_deriv) |
|
1018 |
also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1019 |
by (simp add: distr_density_distr[OF T T', OF inv]) |
47694 | 1020 |
finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N" |
1021 |
by (simp add: comp_def) |
|
41832 | 1022 |
qed |
1023 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
1024 |
lemma (in sigma_finite_measure) RN_deriv_finite: |
47694 | 1025 |
assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" |
1026 |
shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>" |
|
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
1027 |
proof - |
47694 | 1028 |
interpret N: sigma_finite_measure N by fact |
1029 |
from N show ?thesis |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1030 |
using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] density_RN_deriv[OF ac] |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1031 |
by simp |
40859 | 1032 |
qed |
1033 |
||
69739 | 1034 |
lemma (in sigma_finite_measure) |
47694 | 1035 |
assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M" |
40859 | 1036 |
and f: "f \<in> borel_measurable M" |
47694 | 1037 |
shows RN_deriv_integrable: "integrable N f \<longleftrightarrow> |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1038 |
integrable M (\<lambda>x. enn2real (RN_deriv M N x) * f x)" (is ?integrable) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1039 |
and RN_deriv_integral: "integral\<^sup>L N f = (\<integral>x. enn2real (RN_deriv M N x) * f x \<partial>M)" (is ?integral) |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
1040 |
proof - |
47694 | 1041 |
note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp] |
1042 |
interpret N: sigma_finite_measure N by fact |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1043 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1044 |
have eq: "density M (RN_deriv M N) = density M (\<lambda>x. enn2real (RN_deriv M N x))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1045 |
proof (rule density_cong) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1046 |
from RN_deriv_finite[OF assms(1,2,3)] |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1047 |
show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1048 |
by eventually_elim (auto simp: less_top) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1049 |
qed (insert ac, auto) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1050 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1051 |
show ?integrable |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1052 |
apply (subst density_RN_deriv[OF ac, symmetric]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1053 |
unfolding eq |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1054 |
apply (intro integrable_real_density f AE_I2 enn2real_nonneg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1055 |
apply (insert ac, auto) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1056 |
done |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1057 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1058 |
show ?integral |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1059 |
apply (subst density_RN_deriv[OF ac, symmetric]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1060 |
unfolding eq |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1061 |
apply (intro integral_real_density f AE_I2 enn2real_nonneg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1062 |
apply (insert ac, auto) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56537
diff
changeset
|
1063 |
done |
40859 | 1064 |
qed |
1065 |
||
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
1066 |
proposition (in sigma_finite_measure) real_RN_deriv: |
47694 | 1067 |
assumes "finite_measure N" |
1068 |
assumes ac: "absolutely_continuous M N" "sets N = sets M" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1069 |
obtains D where "D \<in> borel_measurable M" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1070 |
and "AE x in M. RN_deriv M N x = ennreal (D x)" |
47694 | 1071 |
and "AE x in N. 0 < D x" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1072 |
and "\<And>x. 0 \<le> D x" |
69730
0c3dcb3a17f6
tagged 5 theories
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents:
69683
diff
changeset
|
1073 |
proof |
47694 | 1074 |
interpret N: finite_measure N by fact |
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:
61169
diff
changeset
|
1075 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1076 |
note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac] |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1077 |
|
47694 | 1078 |
let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1079 |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1080 |
show "(\<lambda>x. enn2real (RN_deriv M N x)) \<in> borel_measurable M" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1081 |
using RN by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1082 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
1083 |
have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1084 |
using RN(1) by (subst RN(2)[symmetric]) (auto simp: emeasure_density) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
1085 |
also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)" |
56996 | 1086 |
by (intro nn_integral_cong) (auto simp: indicator_def) |
47694 | 1087 |
also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)" |
56996 | 1088 |
using RN by (intro nn_integral_cmult_indicator) auto |
47694 | 1089 |
finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" . |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1090 |
moreover |
47694 | 1091 |
have "emeasure M (?RN \<infinity>) = 0" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1092 |
proof (rule ccontr) |
47694 | 1093 |
assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1094 |
then have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1095 |
by (auto simp: zero_less_iff_neq_zero) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1096 |
with eq have "N (?RN \<infinity>) = \<infinity>" by (simp add: ennreal_mult_eq_top_iff) |
47694 | 1097 |
with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1098 |
qed |
47694 | 1099 |
ultimately have "AE x in M. RN_deriv M N x < \<infinity>" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1100 |
using RN by (intro AE_iff_measurable[THEN iffD2]) (auto simp: less_top[symmetric]) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1101 |
then show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1102 |
by auto |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1103 |
then have eq: "AE x in N. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))" |
47694 | 1104 |
using ac absolutely_continuous_AE by auto |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1105 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1106 |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
1107 |
have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)" |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1108 |
by (subst RN(2)[symmetric]) (auto simp: emeasure_density) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
1109 |
also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)" |
56996 | 1110 |
by (intro nn_integral_cong) (auto simp: indicator_def) |
47694 | 1111 |
finally have "AE x in N. RN_deriv M N x \<noteq> 0" |
1112 |
using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq) |
|
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1113 |
with eq show "AE x in N. 0 < enn2real (RN_deriv M N x)" |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1114 |
by (auto simp: enn2real_positive_iff less_top[symmetric] zero_less_iff_neq_zero) |
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1115 |
qed (rule enn2real_nonneg) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1116 |
|
38656 | 1117 |
lemma (in sigma_finite_measure) RN_deriv_singleton: |
47694 | 1118 |
assumes ac: "absolutely_continuous M N" "sets N = sets M" |
1119 |
and x: "{x} \<in> sets M" |
|
1120 |
shows "N {x} = RN_deriv M N x * emeasure M {x}" |
|
38656 | 1121 |
proof - |
61808 | 1122 |
from \<open>{x} \<in> sets M\<close> |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
52141
diff
changeset
|
1123 |
have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)" |
56996 | 1124 |
by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong) |
62975
1d066f6ab25d
Probability: move emeasure and nn_integral from ereal to ennreal
hoelzl
parents:
62623
diff
changeset
|
1125 |
with x density_RN_deriv[OF ac] show ?thesis |
62083 | 1126 |
by (auto simp: max_def) |
38656 | 1127 |
qed |
1128 |
||
1129 |
end |