author | bulwahn |
Mon, 23 Jan 2012 14:00:52 +0100 | |
changeset 46311 | 56fae81902ce |
parent 45777 | c36637603821 |
child 46731 | 5302e932d1e5 |
permissions | -rw-r--r-- |
42067 | 1 |
(* Title: HOL/Probability/Radon_Nikodym.thy |
2 |
Author: Johannes Hölzl, TU München |
|
3 |
*) |
|
4 |
||
5 |
header {*Radon-Nikod{\'y}m derivative*} |
|
6 |
||
38656 | 7 |
theory Radon_Nikodym |
8 |
imports Lebesgue_Integration |
|
9 |
begin |
|
10 |
||
11 |
lemma (in sigma_finite_measure) Ex_finite_integrable_function: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
12 |
shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)" |
38656 | 13 |
proof - |
14 |
obtain A :: "nat \<Rightarrow> 'a set" where |
|
15 |
range: "range A \<subseteq> sets M" and |
|
16 |
space: "(\<Union>i. A i) = space M" and |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
17 |
measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" and |
38656 | 18 |
disjoint: "disjoint_family A" |
19 |
using disjoint_sigma_finite by auto |
|
20 |
let "?B i" = "2^Suc i * \<mu> (A i)" |
|
21 |
have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)" |
|
22 |
proof |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
23 |
fix i have Ai: "A i \<in> sets M" using range by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
24 |
from measure positive_measure[OF this] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
25 |
show "\<exists>x. 0 < x \<and> x < inverse (?B i)" |
43920 | 26 |
by (auto intro!: ereal_dense simp: ereal_0_gt_inverse) |
38656 | 27 |
qed |
28 |
from choice[OF this] obtain n where n: "\<And>i. 0 < n i" |
|
29 |
"\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
30 |
{ fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
31 |
let "?h x" = "\<Sum>i. n i * indicator (A i) x" |
38656 | 32 |
show ?thesis |
33 |
proof (safe intro!: bexI[of _ ?h] del: notI) |
|
39092 | 34 |
have "\<And>i. A i \<in> sets M" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
35 |
using range by fastforce+ |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
36 |
then have "integral\<^isup>P M ?h = (\<Sum>i. n i * \<mu> (A i))" using pos |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
37 |
by (simp add: positive_integral_suminf positive_integral_cmult_indicator) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
38 |
also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
39 |
proof (rule suminf_le_pos) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
40 |
fix N |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
41 |
have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
42 |
using positive_measure[OF `A N \<in> sets M`] n[of N] |
43920 | 43 |
by (intro ereal_mult_right_mono) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
44 |
also have "\<dots> \<le> (1 / 2) ^ Suc N" |
38656 | 45 |
using measure[of N] n[of N] |
43920 | 46 |
by (cases rule: ereal2_cases[of "n N" "\<mu> (A N)"]) |
47 |
(simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
48 |
finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
49 |
show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp |
38656 | 50 |
qed |
43920 | 51 |
finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto |
38656 | 52 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
53 |
{ 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
|
54 |
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
|
55 |
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
|
56 |
by (auto intro!: suminf_cmult_indicator intro: less_imp_le) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
57 |
then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
58 |
note pos = this |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
59 |
fix x show "0 \<le> ?h x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
60 |
proof cases |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
61 |
assume "x \<in> space M" then show "0 \<le> ?h x" using pos by (auto intro: less_imp_le) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
62 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
63 |
assume "x \<notin> space M" then have "\<And>i. x \<notin> A i" using space by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
64 |
then show "0 \<le> ?h x" by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
65 |
qed |
38656 | 66 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
67 |
show "?h \<in> borel_measurable M" using range n |
43920 | 68 |
by (auto intro!: borel_measurable_psuminf borel_measurable_ereal_times ereal_0_le_mult intro: less_imp_le) |
38656 | 69 |
qed |
70 |
qed |
|
71 |
||
40871 | 72 |
subsection "Absolutely continuous" |
73 |
||
38656 | 74 |
definition (in measure_space) |
43920 | 75 |
"absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: ereal))" |
38656 | 76 |
|
41832 | 77 |
lemma (in measure_space) absolutely_continuous_AE: |
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
|
78 |
assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M" |
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
|
79 |
and "absolutely_continuous (measure M')" "AE x. P x" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
80 |
shows "AE x in M'. P x" |
40859 | 81 |
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
|
82 |
interpret \<nu>: measure_space M' by fact |
40859 | 83 |
from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N" |
84 |
unfolding almost_everywhere_def by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
85 |
show "AE x in M'. P x" |
40859 | 86 |
proof (rule \<nu>.AE_I') |
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
|
87 |
show "{x\<in>space M'. \<not> P x} \<subseteq> N" by simp fact |
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
|
88 |
from `absolutely_continuous (measure M')` show "N \<in> \<nu>.null_sets" |
40859 | 89 |
using N unfolding absolutely_continuous_def by auto |
90 |
qed |
|
91 |
qed |
|
92 |
||
39097 | 93 |
lemma (in finite_measure_space) absolutely_continuousI: |
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
|
94 |
assumes "finite_measure_space (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure_space ?\<nu>") |
39097 | 95 |
assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0" |
96 |
shows "absolutely_continuous \<nu>" |
|
97 |
proof (unfold absolutely_continuous_def sets_eq_Pow, safe) |
|
98 |
fix N assume "\<mu> N = 0" "N \<subseteq> space M" |
|
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
|
99 |
interpret v: finite_measure_space ?\<nu> by fact |
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
|
100 |
have "\<nu> N = measure ?\<nu> (\<Union>x\<in>N. {x})" by simp |
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
|
101 |
also have "\<dots> = (\<Sum>x\<in>N. measure ?\<nu> {x})" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
102 |
proof (rule v.measure_setsum[symmetric]) |
39097 | 103 |
show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset) |
104 |
show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto |
|
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
|
105 |
fix x assume "x \<in> N" thus "{x} \<in> sets ?\<nu>" using `N \<subseteq> space M` sets_eq_Pow by auto |
39097 | 106 |
qed |
107 |
also have "\<dots> = 0" |
|
108 |
proof (safe intro!: setsum_0') |
|
109 |
fix x assume "x \<in> N" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
110 |
hence "\<mu> {x} \<le> \<mu> N" "0 \<le> \<mu> {x}" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
111 |
using sets_eq_Pow `N \<subseteq> space M` positive_measure[of "{x}"] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
112 |
by (auto intro!: measure_mono) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
113 |
then have "\<mu> {x} = 0" using `\<mu> N = 0` by simp |
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
|
114 |
thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto |
39097 | 115 |
qed |
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
|
116 |
finally show "\<nu> N = 0" by simp |
39097 | 117 |
qed |
118 |
||
40871 | 119 |
lemma (in measure_space) density_is_absolutely_continuous: |
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
|
120 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
40871 | 121 |
shows "absolutely_continuous \<nu>" |
122 |
using assms unfolding absolutely_continuous_def |
|
123 |
by (simp add: positive_integral_null_set) |
|
124 |
||
125 |
subsection "Existence of the Radon-Nikodym derivative" |
|
126 |
||
38656 | 127 |
lemma (in finite_measure) Radon_Nikodym_aux_epsilon: |
128 |
fixes e :: real assumes "0 < e" |
|
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
|
129 |
assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
130 |
shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
131 |
\<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
132 |
(\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)" |
38656 | 133 |
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
|
134 |
interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
135 |
let "?d A" = "\<mu>' A - M'.\<mu>' A" |
38656 | 136 |
let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B) |
137 |
then {} |
|
138 |
else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)" |
|
139 |
def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}" |
|
140 |
have A_simps[simp]: |
|
141 |
"A 0 = {}" |
|
142 |
"\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all |
|
143 |
{ fix A assume "A \<in> sets M" |
|
144 |
have "?A A \<in> sets M" |
|
145 |
by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) } |
|
146 |
note A'_in_sets = this |
|
147 |
{ fix n have "A n \<in> sets M" |
|
148 |
proof (induct n) |
|
149 |
case (Suc n) thus "A (Suc n) \<in> sets M" |
|
150 |
using A'_in_sets[of "A n"] by (auto split: split_if_asm) |
|
151 |
qed (simp add: A_def) } |
|
152 |
note A_in_sets = this |
|
153 |
hence "range A \<subseteq> sets M" by auto |
|
154 |
{ fix n B |
|
155 |
assume Ex: "\<exists>B. B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> -e" |
|
156 |
hence False: "\<not> (\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B)" by (auto simp: not_less) |
|
157 |
have "?d (A (Suc n)) \<le> ?d (A n) - e" unfolding A_simps if_not_P[OF False] |
|
158 |
proof (rule someI2_ex[OF Ex]) |
|
159 |
fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e" |
|
160 |
hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto |
|
161 |
hence "?d (A n \<union> B) = ?d (A n) + ?d B" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
162 |
using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by simp |
38656 | 163 |
also have "\<dots> \<le> ?d (A n) - e" using dB by simp |
164 |
finally show "?d (A n \<union> B) \<le> ?d (A n) - e" . |
|
165 |
qed } |
|
166 |
note dA_epsilon = this |
|
167 |
{ fix n have "?d (A (Suc n)) \<le> ?d (A n)" |
|
168 |
proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e") |
|
169 |
case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp |
|
170 |
next |
|
171 |
case False |
|
172 |
hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le) |
|
173 |
thus ?thesis by simp |
|
174 |
qed } |
|
175 |
note dA_mono = this |
|
176 |
show ?thesis |
|
177 |
proof (cases "\<exists>n. \<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B") |
|
178 |
case True then obtain n where B: "\<And>B. \<lbrakk> B \<in> sets M; B \<subseteq> space M - A n\<rbrakk> \<Longrightarrow> -e < ?d B" by blast |
|
179 |
show ?thesis |
|
180 |
proof (safe intro!: bexI[of _ "space M - A n"]) |
|
181 |
fix B assume "B \<in> sets M" "B \<subseteq> space M - A n" |
|
182 |
from B[OF this] show "-e < ?d B" . |
|
183 |
next |
|
184 |
show "space M - A n \<in> sets M" by (rule compl_sets) fact |
|
185 |
next |
|
186 |
show "?d (space M) \<le> ?d (space M - A n)" |
|
187 |
proof (induct n) |
|
188 |
fix n assume "?d (space M) \<le> ?d (space M - A n)" |
|
189 |
also have "\<dots> \<le> ?d (space M - A (Suc n))" |
|
190 |
using A_in_sets sets_into_space dA_mono[of n] |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
191 |
by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff) |
38656 | 192 |
finally show "?d (space M) \<le> ?d (space M - A (Suc n))" . |
193 |
qed simp |
|
194 |
qed |
|
195 |
next |
|
196 |
case False hence B: "\<And>n. \<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e" |
|
197 |
by (auto simp add: not_less) |
|
198 |
{ fix n have "?d (A n) \<le> - real n * e" |
|
199 |
proof (induct n) |
|
200 |
case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
201 |
next |
43920 | 202 |
case 0 with M'.empty_measure show ?case by (simp add: zero_ereal_def) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
203 |
qed } note dA_less = this |
38656 | 204 |
have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq |
205 |
proof (rule incseq_SucI) |
|
206 |
fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto |
|
207 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
208 |
have A: "incseq A" by (auto intro!: incseq_SucI) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
209 |
from finite_continuity_from_below[OF _ A] `range A \<subseteq> sets M` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
210 |
M'.finite_continuity_from_below[OF _ A] |
38656 | 211 |
have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)" |
44568
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
43923
diff
changeset
|
212 |
by (auto intro!: tendsto_diff) |
38656 | 213 |
obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto |
214 |
moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less] |
|
215 |
have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps) |
|
216 |
ultimately show ?thesis by auto |
|
217 |
qed |
|
218 |
qed |
|
219 |
||
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
220 |
lemma (in finite_measure) restricted_measure_subset: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
221 |
assumes S: "S \<in> sets M" and X: "X \<subseteq> S" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
222 |
shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
223 |
proof cases |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
224 |
note r = restricted_finite_measure[OF S] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
225 |
{ assume "X \<in> sets M" with S X show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
226 |
unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def by auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
227 |
{ assume "X \<notin> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
228 |
moreover then have "S \<inter> X \<notin> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
229 |
using X by (simp add: Int_absorb1) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
230 |
ultimately show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
231 |
unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def using S by auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
232 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
233 |
|
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
234 |
lemma (in finite_measure) restricted_measure: |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
235 |
assumes X: "S \<in> sets M" "X \<in> sets (restricted_space S)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
236 |
shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
237 |
proof - |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
238 |
from X have "S \<in> sets M" "X \<subseteq> S" by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
239 |
from restricted_measure_subset[OF this] show ?thesis . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
240 |
qed |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
241 |
|
38656 | 242 |
lemma (in finite_measure) Radon_Nikodym_aux: |
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
|
243 |
assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?M'") |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
244 |
shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
245 |
\<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
246 |
(\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)" |
38656 | 247 |
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
|
248 |
interpret M': finite_measure ?M' where |
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
|
249 |
"space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
250 |
let "?d A" = "\<mu>' A - M'.\<mu>' A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
251 |
let "?P A B n" = "A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)" |
39092 | 252 |
let "?r S" = "restricted_space S" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
253 |
{ fix S n assume S: "S \<in> sets M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
254 |
note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
255 |
then have "finite_measure (?r S)" "0 < 1 / real (Suc n)" |
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
|
256 |
"finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
257 |
from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
258 |
have "?P X S n" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
259 |
proof (intro conjI ballI impI) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
260 |
show "X \<in> sets M" "X \<subseteq> S" using X(1) `S \<in> sets M` by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
261 |
have "S \<in> op \<inter> S ` sets M" using `S \<in> sets M` by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
262 |
then show "?d S \<le> ?d X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
263 |
using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
264 |
fix C assume "C \<in> sets M" "C \<subseteq> X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
265 |
then have "C \<in> sets (restricted_space S)" "C \<subseteq> X" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
266 |
using `S \<in> sets M` `X \<subseteq> S` by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
267 |
with X(2) show "- 1 / real (Suc n) < ?d C" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
268 |
using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto |
38656 | 269 |
qed |
270 |
hence "\<exists>A. ?P A S n" by auto } |
|
271 |
note Ex_P = this |
|
272 |
def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)" |
|
273 |
have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def) |
|
274 |
have A_0[simp]: "A 0 = space M" unfolding A_def by simp |
|
275 |
{ fix i have "A i \<in> sets M" unfolding A_def |
|
276 |
proof (induct i) |
|
277 |
case (Suc i) |
|
278 |
from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc |
|
279 |
by (rule someI2_ex) simp |
|
280 |
qed simp } |
|
281 |
note A_in_sets = this |
|
282 |
{ fix n have "?P (A (Suc n)) (A n) n" |
|
283 |
using Ex_P[OF A_in_sets] unfolding A_Suc |
|
284 |
by (rule someI2_ex) simp } |
|
285 |
note P_A = this |
|
286 |
have "range A \<subseteq> sets M" using A_in_sets by auto |
|
287 |
have A_mono: "\<And>i. A (Suc i) \<subseteq> A i" using P_A by simp |
|
288 |
have mono_dA: "mono (\<lambda>i. ?d (A i))" using P_A by (simp add: mono_iff_le_Suc) |
|
289 |
have epsilon: "\<And>C i. \<lbrakk> C \<in> sets M; C \<subseteq> A (Suc i) \<rbrakk> \<Longrightarrow> - 1 / real (Suc i) < ?d C" |
|
290 |
using P_A by auto |
|
291 |
show ?thesis |
|
292 |
proof (safe intro!: bexI[of _ "\<Inter>i. A i"]) |
|
293 |
show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
294 |
have A: "decseq A" using A_mono by (auto intro!: decseq_SucI) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
295 |
from |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
296 |
finite_continuity_from_above[OF `range A \<subseteq> sets M` A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
297 |
M'.finite_continuity_from_above[OF `range A \<subseteq> sets M` A] |
44568
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
huffman
parents:
43923
diff
changeset
|
298 |
have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (intro tendsto_diff) |
38656 | 299 |
thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _] |
300 |
by (rule_tac LIMSEQ_le_const) (auto intro!: exI) |
|
301 |
next |
|
302 |
fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)" |
|
303 |
show "0 \<le> ?d B" |
|
304 |
proof (rule ccontr) |
|
305 |
assume "\<not> 0 \<le> ?d B" |
|
306 |
hence "0 < - ?d B" by auto |
|
307 |
from ex_inverse_of_nat_Suc_less[OF this] |
|
308 |
obtain n where *: "?d B < - 1 / real (Suc n)" |
|
309 |
by (auto simp: real_eq_of_nat inverse_eq_divide field_simps) |
|
310 |
have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc) |
|
311 |
from epsilon[OF B(1) this] * |
|
312 |
show False by auto |
|
313 |
qed |
|
314 |
qed |
|
315 |
qed |
|
316 |
||
317 |
lemma (in finite_measure) Radon_Nikodym_finite_measure: |
|
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
|
318 |
assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'") |
38656 | 319 |
assumes "absolutely_continuous \<nu>" |
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
|
320 |
shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
38656 | 321 |
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
|
322 |
interpret M': finite_measure ?M' |
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
|
323 |
where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" |
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
|
324 |
using assms(1) by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
325 |
def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A)}" |
38656 | 326 |
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto |
327 |
hence "G \<noteq> {}" by auto |
|
328 |
{ fix f g assume f: "f \<in> G" and g: "g \<in> G" |
|
329 |
have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def |
|
330 |
proof safe |
|
331 |
show "?max \<in> borel_measurable M" using f g unfolding G_def by auto |
|
332 |
let ?A = "{x \<in> space M. f x \<le> g x}" |
|
333 |
have "?A \<in> sets M" using f g unfolding G_def by auto |
|
334 |
fix A assume "A \<in> sets M" |
|
335 |
hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto |
|
336 |
have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A" |
|
337 |
using sets_into_space[OF `A \<in> sets M`] by auto |
|
338 |
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x = |
|
339 |
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x" |
|
340 |
by (auto simp: indicator_def max_def) |
|
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
|
341 |
hence "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) = |
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
|
342 |
(\<integral>\<^isup>+x. g x * indicator (?A \<inter> A) x \<partial>M) + |
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
|
343 |
(\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)" |
38656 | 344 |
using f g sets unfolding G_def |
345 |
by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator) |
|
346 |
also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)" |
|
347 |
using f g sets unfolding G_def by (auto intro!: add_mono) |
|
348 |
also have "\<dots> = \<nu> A" |
|
349 |
using M'.measure_additive[OF sets] union by auto |
|
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
|
350 |
finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" . |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
351 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
352 |
fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max) |
38656 | 353 |
qed } |
354 |
note max_in_G = this |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
355 |
{ fix f assume "incseq f" and f: "\<And>i. f i \<in> G" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
356 |
have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def |
38656 | 357 |
proof safe |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
358 |
show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
359 |
using f by (auto simp: G_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
360 |
{ fix x show "0 \<le> (SUP i. f i x)" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
361 |
using f by (auto simp: G_def intro: SUP_upper2) } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
362 |
next |
38656 | 363 |
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
|
364 |
have "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) = |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
365 |
(\<integral>\<^isup>+x. (SUP i. f i x * indicator A x) \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
366 |
by (intro positive_integral_cong) (simp split: split_indicator) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
367 |
also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
368 |
using `incseq f` f `A \<in> sets M` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
369 |
by (intro positive_integral_monotone_convergence_SUP) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
370 |
(auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
371 |
finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> \<nu> A" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
372 |
using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def) |
38656 | 373 |
qed } |
374 |
note SUP_in_G = this |
|
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
|
375 |
let ?y = "SUP g : G. integral\<^isup>P M g" |
38656 | 376 |
have "?y \<le> \<nu> (space M)" unfolding G_def |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
377 |
proof (safe intro!: SUP_least) |
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
|
378 |
fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A" |
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
|
379 |
from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)" |
38656 | 380 |
by (simp cong: positive_integral_cong) |
381 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
382 |
from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^isup>P M"] guess ys .. note ys = this |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
383 |
then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" |
38656 | 384 |
proof safe |
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
|
385 |
fix n assume "range ys \<subseteq> integral\<^isup>P M ` G" |
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
|
386 |
hence "ys n \<in> integral\<^isup>P M ` G" by auto |
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
|
387 |
thus "\<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" by auto |
38656 | 388 |
qed |
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
|
389 |
from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^isup>P M (gs n) = ys n" by auto |
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
|
390 |
hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto |
38656 | 391 |
let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
392 |
def f \<equiv> "\<lambda>x. SUP i. ?g i x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
393 |
let "?F A x" = "f x * indicator A x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
394 |
have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto |
38656 | 395 |
{ fix i have "?g i \<in> G" |
396 |
proof (induct i) |
|
397 |
case 0 thus ?case by simp fact |
|
398 |
next |
|
399 |
case (Suc i) |
|
400 |
with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case |
|
401 |
by (auto simp add: atMost_Suc intro!: max_in_G) |
|
402 |
qed } |
|
403 |
note g_in_G = this |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
404 |
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
|
405 |
by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
406 |
from SUP_in_G[OF this g_in_G] have "f \<in> G" unfolding f_def . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
407 |
then have [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
408 |
have "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding f_def |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
409 |
using g_in_G `incseq ?g` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
410 |
by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def) |
38656 | 411 |
also have "\<dots> = ?y" |
412 |
proof (rule antisym) |
|
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
|
413 |
show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y" |
44918 | 414 |
using g_in_G |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
415 |
by (auto intro!: exI Sup_mono simp: SUP_def) |
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
|
416 |
show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq |
38656 | 417 |
by (auto intro!: SUP_mono positive_integral_mono Max_ge) |
418 |
qed |
|
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
|
419 |
finally have int_f_eq_y: "integral\<^isup>P M f = ?y" . |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
420 |
have "\<And>x. 0 \<le> f x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
421 |
unfolding f_def using `\<And>i. gs i \<in> G` |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
422 |
by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
423 |
let "?t A" = "\<nu> A - (\<integral>\<^isup>+x. ?F A x \<partial>M)" |
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
|
424 |
let ?M = "M\<lparr> measure := ?t\<rparr>" |
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
|
425 |
interpret M: sigma_algebra ?M |
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
|
426 |
by (intro sigma_algebra_cong) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
427 |
have f_le_\<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> \<nu> A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
428 |
using `f \<in> G` unfolding G_def by auto |
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
|
429 |
have fmM: "finite_measure ?M" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
430 |
proof |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
431 |
show "measure_space ?M" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
432 |
proof (default, simp_all add: countably_additive_def positive_def, safe del: notI) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
433 |
fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
434 |
have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. (\<Sum>n. ?F (A n) x) \<partial>M)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
435 |
using `range A \<subseteq> sets M` `\<And>x. 0 \<le> f x` |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
436 |
by (intro positive_integral_suminf[symmetric]) auto |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
437 |
also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
438 |
using `\<And>x. 0 \<le> f x` |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
439 |
by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`]) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
440 |
finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" . |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
441 |
moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
442 |
using M'.measure_countably_additive A by (simp add: comp_def) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
443 |
moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<infinity>" using M'.finite_measure A by (simp add: countable_UN) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
444 |
moreover { |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
445 |
have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
446 |
using A `f \<in> G` unfolding G_def by (auto simp: countable_UN) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
447 |
also have "\<nu> (\<Union>i. A i) < \<infinity>" using v_fin by simp |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
448 |
finally have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<noteq> \<infinity>" by simp } |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
449 |
moreover have "\<And>i. (\<integral>\<^isup>+x. ?F (A i) x \<partial>M) \<le> \<nu> (A i)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
450 |
using A by (intro f_le_\<nu>) auto |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
451 |
ultimately |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
452 |
show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
453 |
by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
454 |
next |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
455 |
fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
456 |
using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
457 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
458 |
next |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
459 |
show "measure ?M (space ?M) \<noteq> \<infinity>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
460 |
using positive_integral_positive[of "?F (space M)"] |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
461 |
by (cases rule: ereal2_cases[of "\<nu> (space M)" "\<integral>\<^isup>+ x. ?F (space M) x \<partial>M"]) auto |
38656 | 462 |
qed |
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
|
463 |
then interpret M: finite_measure ?M |
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
|
464 |
where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t" |
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
|
465 |
by (simp_all add: fmM) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
466 |
have ac: "absolutely_continuous ?t" unfolding absolutely_continuous_def |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
467 |
proof |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
468 |
fix N assume N: "N \<in> null_sets" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
469 |
with `absolutely_continuous \<nu>` have "\<nu> N = 0" unfolding absolutely_continuous_def by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
470 |
moreover with N have "(\<integral>\<^isup>+ x. ?F N x \<partial>M) \<le> \<nu> N" using `f \<in> G` by (auto simp: G_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
471 |
ultimately show "\<nu> N - (\<integral>\<^isup>+ x. ?F N x \<partial>M) = 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
472 |
using positive_integral_positive by (auto intro!: antisym) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
473 |
qed |
38656 | 474 |
have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0" |
475 |
proof (rule ccontr) |
|
476 |
assume "\<not> ?thesis" |
|
477 |
then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A" |
|
478 |
by (auto simp: not_le) |
|
479 |
note pos |
|
480 |
also have "?t A \<le> ?t (space M)" |
|
481 |
using M.measure_mono[of A "space M"] A sets_into_space by simp |
|
482 |
finally have pos_t: "0 < ?t (space M)" by simp |
|
483 |
moreover |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
484 |
then have "\<mu> (space M) \<noteq> 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
485 |
using ac unfolding absolutely_continuous_def by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
486 |
then have pos_M: "0 < \<mu> (space M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
487 |
using positive_measure[OF top] by (simp add: le_less) |
38656 | 488 |
moreover |
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
|
489 |
have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)" |
38656 | 490 |
using `f \<in> G` unfolding G_def by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
491 |
hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>" |
38656 | 492 |
using M'.finite_measure_of_space by auto |
493 |
moreover |
|
494 |
def b \<equiv> "?t (space M) / \<mu> (space M) / 2" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
495 |
ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
496 |
using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"] |
43920 | 497 |
by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
498 |
(simp_all add: field_simps) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
499 |
then have b: "b \<noteq> 0" "0 \<le> b" "0 < b" "b \<noteq> \<infinity>" by auto |
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
|
500 |
let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>" |
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
|
501 |
interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
502 |
have Mb: "finite_measure ?Mb" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
503 |
proof |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
504 |
show "measure_space ?Mb" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
505 |
proof |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
506 |
show "positive ?Mb (measure ?Mb)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
507 |
using `0 \<le> b` by (auto simp: positive_def) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
508 |
show "countably_additive ?Mb (measure ?Mb)" |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
509 |
using `0 \<le> b` measure_countably_additive |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
510 |
by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq) |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
511 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
512 |
show "measure ?Mb (space ?Mb) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
513 |
using b by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
514 |
qed |
38656 | 515 |
from M.Radon_Nikodym_aux[OF this] |
516 |
obtain A0 where "A0 \<in> sets M" and |
|
517 |
space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
518 |
*: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
519 |
unfolding M.\<mu>'_def finite_measure.\<mu>'_def[OF Mb] by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
520 |
{ fix B assume B: "B \<in> sets M" "B \<subseteq> A0" |
38656 | 521 |
with *[OF this] have "b * \<mu> B \<le> ?t B" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
522 |
using M'.finite_measure b finite_measure M.positive_measure[OF B(1)] |
43920 | 523 |
by (cases rule: ereal2_cases[of "?t B" "b * \<mu> B"]) auto } |
38656 | 524 |
note bM_le_t = this |
525 |
let "?f0 x" = "f x + b * indicator A0 x" |
|
526 |
{ fix A assume A: "A \<in> sets M" |
|
527 |
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto |
|
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
|
528 |
have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) = |
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
|
529 |
(\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
530 |
by (auto intro!: positive_integral_cong split: split_indicator) |
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
|
531 |
hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) = |
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
|
532 |
(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
533 |
using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
534 |
by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) } |
38656 | 535 |
note f0_eq = this |
536 |
{ fix A assume A: "A \<in> sets M" |
|
537 |
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto |
|
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
|
538 |
have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A" |
38656 | 539 |
using `f \<in> G` A unfolding G_def by auto |
540 |
note f0_eq[OF A] |
|
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
|
541 |
also have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0) \<le> |
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
|
542 |
(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t (A \<inter> A0)" |
38656 | 543 |
using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M` |
544 |
by (auto intro!: add_left_mono) |
|
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
|
545 |
also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t A" |
38656 | 546 |
using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`] |
547 |
by (auto intro!: add_left_mono) |
|
548 |
also have "\<dots> \<le> \<nu> A" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
549 |
using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`] positive_integral_positive[of "?F A"] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
550 |
by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") auto |
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
|
551 |
finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
552 |
hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def |
43920 | 553 |
by (auto intro!: borel_measurable_indicator borel_measurable_ereal_add |
554 |
borel_measurable_ereal_times ereal_add_nonneg_nonneg) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
555 |
have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
556 |
"b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>" |
38656 | 557 |
using `A0 \<in> sets M` b |
558 |
finite_measure[of A0] M.finite_measure[of A0] |
|
559 |
finite_measure_of_space M.finite_measure_of_space |
|
560 |
by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
561 |
have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>" |
43920 | 562 |
using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff |
38656 | 563 |
by (auto cong: positive_integral_cong) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
564 |
have "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def |
38656 | 565 |
using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
566 |
using positive_integral_positive[of "?F (space M)"] |
43920 | 567 |
by (cases rule: ereal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
568 |
(auto simp: field_simps mult_less_cancel_left) |
38656 | 569 |
also have "\<dots> \<le> ?t A0 - b * \<mu> A0" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
570 |
using space_less_A0 b |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
571 |
using |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
572 |
`A0 \<in> sets M`[THEN M.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
573 |
top[THEN M.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
574 |
apply safe |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
575 |
apply simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
576 |
using |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
577 |
`A0 \<in> sets M`[THEN real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
578 |
`A0 \<in> sets M`[THEN M'.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
579 |
top[THEN real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
580 |
top[THEN M'.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
581 |
by (cases b) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
582 |
finally have 1: "b * \<mu> A0 < ?t A0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
583 |
using |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
584 |
`A0 \<in> sets M`[THEN M.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
585 |
apply safe |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
586 |
apply simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
587 |
using |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
588 |
`A0 \<in> sets M`[THEN real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
589 |
`A0 \<in> sets M`[THEN M'.real_measure] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
590 |
by (cases b) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
591 |
have "0 < ?t A0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
592 |
using b `A0 \<in> sets M` by (auto intro!: le_less_trans[OF _ 1]) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
593 |
then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def |
38656 | 594 |
using `A0 \<in> sets M` by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
595 |
then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto |
43920 | 596 |
hence "0 < b * \<mu> A0" using b by (auto simp: ereal_zero_less_0_iff) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
597 |
with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
598 |
using `f \<in> G` |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
599 |
by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive) |
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
|
600 |
also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space |
38656 | 601 |
by (simp cong: positive_integral_cong) |
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
|
602 |
finally have "?y < integral\<^isup>P M ?f0" by simp |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
603 |
moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper) |
38656 | 604 |
ultimately show False by auto |
605 |
qed |
|
606 |
show ?thesis |
|
607 |
proof (safe intro!: bexI[of _ f]) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
608 |
fix A assume A: "A\<in>sets M" |
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
|
609 |
show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
38656 | 610 |
proof (rule antisym) |
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
|
611 |
show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A" |
38656 | 612 |
using `f \<in> G` `A \<in> sets M` unfolding G_def by auto |
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
|
613 |
show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
38656 | 614 |
using upper_bound[THEN bspec, OF `A \<in> sets M`] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
615 |
using M'.real_measure[OF A] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
616 |
by (cases "integral\<^isup>P M (?F A)") auto |
38656 | 617 |
qed |
618 |
qed simp |
|
619 |
qed |
|
620 |
||
40859 | 621 |
lemma (in finite_measure) split_space_into_finite_sets_and_rest: |
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
|
622 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") |
40859 | 623 |
assumes ac: "absolutely_continuous \<nu>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
624 |
shows "\<exists>A0\<in>sets M. \<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> A0 = space M - (\<Union>i. B i) \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
625 |
(\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<infinity>)) \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
626 |
(\<forall>i. \<nu> (B i) \<noteq> \<infinity>)" |
38656 | 627 |
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
|
628 |
interpret v: measure_space ?N |
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
|
629 |
where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>" |
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
|
630 |
by fact auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
631 |
let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<infinity>}" |
38656 | 632 |
let ?a = "SUP Q:?Q. \<mu> Q" |
633 |
have "{} \<in> ?Q" using v.empty_measure by auto |
|
634 |
then have Q_not_empty: "?Q \<noteq> {}" by blast |
|
635 |
have "?a \<le> \<mu> (space M)" using sets_into_space |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
636 |
by (auto intro!: SUP_least measure_mono top) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
637 |
then have "?a \<noteq> \<infinity>" using finite_measure_of_space |
38656 | 638 |
by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
639 |
from SUPR_countable_SUPR[OF Q_not_empty, of \<mu>] |
38656 | 640 |
obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" |
641 |
by auto |
|
642 |
then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto |
|
643 |
from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q" |
|
644 |
by auto |
|
645 |
then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp |
|
646 |
let "?O n" = "\<Union>i\<le>n. Q' i" |
|
647 |
have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)" |
|
648 |
proof (rule continuity_from_below[of ?O]) |
|
649 |
show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
650 |
show "incseq ?O" by (fastforce intro!: incseq_SucI) |
38656 | 651 |
qed |
652 |
have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto |
|
653 |
have O_sets: "\<And>i. ?O i \<in> sets M" |
|
654 |
using Q' by (auto intro!: finite_UN Un) |
|
655 |
then have O_in_G: "\<And>i. ?O i \<in> ?Q" |
|
656 |
proof (safe del: notI) |
|
657 |
fix i have "Q' ` {..i} \<subseteq> sets M" |
|
658 |
using Q' by (auto intro: finite_UN) |
|
659 |
with v.measure_finitely_subadditive[of "{.. i}" Q'] |
|
660 |
have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
661 |
also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
662 |
finally show "\<nu> (?O i) \<noteq> \<infinity>" by simp |
38656 | 663 |
qed auto |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
664 |
have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce |
38656 | 665 |
have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric] |
666 |
proof (rule antisym) |
|
667 |
show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim |
|
668 |
using Q' by (auto intro!: SUP_mono measure_mono finite_UN) |
|
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
669 |
show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUP_def |
38656 | 670 |
proof (safe intro!: Sup_mono, unfold bex_simps) |
671 |
fix i |
|
672 |
have *: "(\<Union>Q' ` {..i}) = ?O i" by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
673 |
then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<infinity>) \<and> |
38656 | 674 |
\<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x" |
675 |
using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) |
|
676 |
qed |
|
677 |
qed |
|
678 |
let "?O_0" = "(\<Union>i. ?O i)" |
|
679 |
have "?O_0 \<in> sets M" using Q' by auto |
|
40859 | 680 |
def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n" |
38656 | 681 |
{ fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) } |
682 |
note Q_sets = this |
|
40859 | 683 |
show ?thesis |
684 |
proof (intro bexI exI conjI ballI impI allI) |
|
685 |
show "disjoint_family Q" |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
686 |
by (fastforce simp: disjoint_family_on_def Q_def |
40859 | 687 |
split: nat.split_asm) |
688 |
show "range Q \<subseteq> sets M" |
|
689 |
using Q_sets by auto |
|
690 |
{ fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
691 |
show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>" |
40859 | 692 |
proof (rule disjCI, simp) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
693 |
assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<infinity>" |
40859 | 694 |
show "\<mu> A = 0 \<and> \<nu> A = 0" |
695 |
proof cases |
|
696 |
assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0" |
|
697 |
unfolding absolutely_continuous_def by auto |
|
698 |
ultimately show ?thesis by simp |
|
699 |
next |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
700 |
assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<infinity>" using positive_measure[OF A(1)] by auto |
40859 | 701 |
with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)" |
702 |
using Q' by (auto intro!: measure_additive countable_UN) |
|
703 |
also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))" |
|
704 |
proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified]) |
|
705 |
show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
706 |
using `\<nu> A \<noteq> \<infinity>` O_sets A by auto |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
707 |
qed (fastforce intro!: incseq_SucI) |
40859 | 708 |
also have "\<dots> \<le> ?a" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
709 |
proof (safe intro!: SUP_least) |
40859 | 710 |
fix i have "?O i \<union> A \<in> ?Q" |
711 |
proof (safe del: notI) |
|
712 |
show "?O i \<union> A \<in> sets M" using O_sets A by auto |
|
713 |
from O_in_G[of i] |
|
714 |
moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A" |
|
715 |
using v.measure_subadditive[of "?O i" A] A O_sets by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
716 |
ultimately show "\<nu> (?O i \<union> A) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
717 |
using `\<nu> A \<noteq> \<infinity>` by auto |
40859 | 718 |
qed |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44918
diff
changeset
|
719 |
then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule SUP_upper) |
40859 | 720 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
721 |
finally have "\<mu> A = 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
722 |
unfolding a_eq using real_measure[OF `?O_0 \<in> sets M`] real_measure[OF A(1)] by auto |
40859 | 723 |
with `\<mu> A \<noteq> 0` show ?thesis by auto |
724 |
qed |
|
725 |
qed } |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
726 |
{ fix i show "\<nu> (Q i) \<noteq> \<infinity>" |
40859 | 727 |
proof (cases i) |
728 |
case 0 then show ?thesis |
|
729 |
unfolding Q_def using Q'[of 0] by simp |
|
730 |
next |
|
731 |
case (Suc n) |
|
732 |
then show ?thesis unfolding Q_def |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
733 |
using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
734 |
using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
735 |
using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono] |
43920 | 736 |
by (cases rule: ereal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto |
40859 | 737 |
qed } |
738 |
show "space M - ?O_0 \<in> sets M" using Q'_sets by auto |
|
739 |
{ fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)" |
|
740 |
proof (induct j) |
|
741 |
case 0 then show ?case by (simp add: Q_def) |
|
742 |
next |
|
743 |
case (Suc j) |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
744 |
have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastforce |
40859 | 745 |
have "{..j} \<union> {..Suc j} = {..Suc j}" by auto |
746 |
then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)" |
|
747 |
by (simp add: UN_Un[symmetric] Q_def del: UN_Un) |
|
748 |
then show ?case using Suc by (auto simp add: eq atMost_Suc) |
|
749 |
qed } |
|
750 |
then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp |
|
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44568
diff
changeset
|
751 |
then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastforce |
40859 | 752 |
qed |
753 |
qed |
|
754 |
||
755 |
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite: |
|
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
|
756 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") |
40859 | 757 |
assumes "absolutely_continuous \<nu>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
758 |
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))" |
40859 | 759 |
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
|
760 |
interpret v: measure_space ?N |
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
|
761 |
where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>" |
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
|
762 |
by fact auto |
40859 | 763 |
from split_space_into_finite_sets_and_rest[OF assms] |
764 |
obtain Q0 and Q :: "nat \<Rightarrow> 'a set" |
|
765 |
where Q: "disjoint_family Q" "range Q \<subseteq> sets M" |
|
766 |
and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
767 |
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
768 |
and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<infinity>" by force |
40859 | 769 |
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
770 |
have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. |
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
|
771 |
\<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))" |
38656 | 772 |
proof |
773 |
fix i |
|
43920 | 774 |
have indicator_eq: "\<And>f x A. (f x :: ereal) * indicator (Q i \<inter> A) x * indicator (Q i) x |
38656 | 775 |
= (f x * indicator (Q i) x) * indicator A x" |
776 |
unfolding indicator_def by auto |
|
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
|
777 |
have fm: "finite_measure (restricted_space (Q i))" |
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
|
778 |
(is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]]) |
38656 | 779 |
then interpret R: finite_measure ?R . |
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
|
780 |
have fmv: "finite_measure (restricted_space (Q i) \<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?Q") |
38656 | 781 |
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
|
782 |
show "measure_space ?Q" |
38656 | 783 |
using v.restricted_measure_space Q_sets[of i] by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
784 |
show "measure ?Q (space ?Q) \<noteq> \<infinity>" using Q_fin by simp |
38656 | 785 |
qed |
786 |
have "R.absolutely_continuous \<nu>" |
|
787 |
using `absolutely_continuous \<nu>` `Q i \<in> sets M` |
|
788 |
by (auto simp: R.absolutely_continuous_def absolutely_continuous_def) |
|
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
|
789 |
from R.Radon_Nikodym_finite_measure[OF fmv this] |
38656 | 790 |
obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M" |
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
|
791 |
and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \<partial>M)" |
38656 | 792 |
unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`] |
793 |
positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
794 |
then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. |
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
|
795 |
\<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
796 |
by (auto intro!: exI[of _ "\<lambda>x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
797 |
split: split_indicator split_if_asm simp: max_def) |
38656 | 798 |
qed |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
799 |
from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x" |
38656 | 800 |
and f: "\<And>A i. A \<in> sets M \<Longrightarrow> |
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
|
801 |
\<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" |
38656 | 802 |
by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
803 |
let "?f x" = "(\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x" |
38656 | 804 |
show ?thesis |
805 |
proof (safe intro!: bexI[of _ ?f]) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
806 |
show "?f \<in> borel_measurable M" using Q0 borel Q_sets |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
807 |
by (auto intro!: measurable_If) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
808 |
show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def) |
38656 | 809 |
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
|
810 |
have Qi: "\<And>i. Q i \<in> sets M" using Q by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
811 |
have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
812 |
"\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x" |
43920 | 813 |
using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
814 |
have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
815 |
using borel by (intro positive_integral_cong) (auto simp: indicator_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
816 |
also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
817 |
using borel Qi Q0(1) `A \<in> sets M` |
43920 | 818 |
by (subst positive_integral_add) (auto simp del: ereal_infty_mult |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
819 |
simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
820 |
also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
821 |
by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
822 |
finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)" . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
823 |
moreover have "(\<Sum>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)" |
40859 | 824 |
using Q Q_sets `A \<in> sets M` |
825 |
by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified]) |
|
826 |
(auto simp: disjoint_family_on_def) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
827 |
moreover have "\<infinity> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)" |
40859 | 828 |
proof - |
829 |
have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast |
|
830 |
from in_Q0[OF this] show ?thesis by auto |
|
38656 | 831 |
qed |
40859 | 832 |
moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M" |
833 |
using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN) |
|
834 |
moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}" |
|
835 |
using `A \<in> sets M` sets_into_space Q0 by auto |
|
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
|
836 |
ultimately show "\<nu> A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)" |
40859 | 837 |
using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"] |
838 |
by simp |
|
38656 | 839 |
qed |
840 |
qed |
|
841 |
||
842 |
lemma (in sigma_finite_measure) Radon_Nikodym: |
|
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
|
843 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
844 |
assumes ac: "absolutely_continuous \<nu>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
845 |
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))" |
38656 | 846 |
proof - |
847 |
from Ex_finite_integrable_function |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
848 |
obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and |
38656 | 849 |
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
|
850 |
nn: "\<And>x. 0 \<le> h x" and |
38656 | 851 |
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
|
852 |
"\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto |
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
|
853 |
let "?T A" = "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M)" |
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
|
854 |
let ?MT = "M\<lparr> measure := ?T \<rparr>" |
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
|
855 |
interpret T: finite_measure ?MT |
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
|
856 |
where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
857 |
using borel finite nn |
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
858 |
by (auto intro!: measure_space_density finite_measureI cong: positive_integral_cong) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
859 |
have "T.absolutely_continuous \<nu>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
860 |
proof (unfold T.absolutely_continuous_def, safe) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
861 |
fix N assume "N \<in> sets M" "(\<integral>\<^isup>+x. h x * indicator N x \<partial>M) = 0" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
862 |
with borel ac pos have "AE x. x \<notin> N" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
863 |
by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
864 |
then have "N \<in> null_sets" using `N \<in> sets M` sets_into_space |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
865 |
by (subst (asm) AE_iff_measurable[OF `N \<in> sets M`]) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
866 |
then show "\<nu> N = 0" using ac by (auto simp: absolutely_continuous_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
867 |
qed |
38656 | 868 |
from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
869 |
obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and |
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
|
870 |
fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>?MT)" |
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
|
871 |
by (auto simp: measurable_def) |
38656 | 872 |
show ?thesis |
873 |
proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"]) |
|
874 |
show "(\<lambda>x. h x * f x) \<in> borel_measurable M" |
|
43920 | 875 |
using borel f_borel by (auto intro: borel_measurable_ereal_times) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
876 |
show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto |
38656 | 877 |
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
|
878 |
then show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
879 |
unfolding fT[OF `A \<in> sets M`] mult_assoc using nn borel f_borel |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
880 |
by (intro positive_integral_translated_density) auto |
38656 | 881 |
qed |
882 |
qed |
|
883 |
||
40859 | 884 |
section "Uniqueness of densities" |
885 |
||
886 |
lemma (in measure_space) finite_density_unique: |
|
887 |
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
888 |
assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
889 |
and fin: "integral\<^isup>P M f \<noteq> \<infinity>" |
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
|
890 |
shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M)) |
40859 | 891 |
\<longleftrightarrow> (AE x. f x = g x)" |
892 |
(is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _") |
|
893 |
proof (intro iffI ballI) |
|
894 |
fix A assume eq: "AE x. f x = g x" |
|
41705 | 895 |
then show "?P f A = ?P g A" |
896 |
by (auto intro: positive_integral_cong_AE) |
|
40859 | 897 |
next |
898 |
assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" |
|
899 |
from this[THEN bspec, OF top] fin |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
900 |
have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong) |
40859 | 901 |
{ fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
902 |
and pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
903 |
and g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" |
40859 | 904 |
let ?N = "{x\<in>space M. g x < f x}" |
905 |
have N: "?N \<in> sets M" using borel by simp |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
906 |
have "?P g ?N \<le> integral\<^isup>P M g" using pos |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
907 |
by (intro positive_integral_mono_AE) (auto split: split_indicator) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
908 |
then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by auto |
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
|
909 |
have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)" |
40859 | 910 |
by (auto intro!: positive_integral_cong simp: indicator_def) |
911 |
also have "\<dots> = ?P f ?N - ?P g ?N" |
|
912 |
proof (rule positive_integral_diff) |
|
913 |
show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M" |
|
914 |
using borel N by auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
915 |
show "AE x. g x * indicator ?N x \<le> f x * indicator ?N x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
916 |
"AE x. 0 \<le> g x * indicator ?N x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
917 |
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
|
918 |
qed fact |
40859 | 919 |
also have "\<dots> = 0" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
920 |
unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
921 |
finally have "AE x. f x \<le> g x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
922 |
using pos borel positive_integral_PInf_AE[OF borel(2) g_fin] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
923 |
by (subst (asm) positive_integral_0_iff_AE) |
43920 | 924 |
(auto split: split_indicator simp: not_less ereal_minus_le_iff) } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
925 |
from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
926 |
show "AE x. f x = g x" by auto |
40859 | 927 |
qed |
928 |
||
929 |
lemma (in finite_measure) density_unique_finite_measure: |
|
930 |
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
931 |
assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> f' x" |
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
|
932 |
assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)" |
40859 | 933 |
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A") |
934 |
shows "AE x. f x = f' x" |
|
935 |
proof - |
|
936 |
let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A" |
|
937 |
let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x" |
|
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
|
938 |
interpret M: measure_space "M\<lparr> measure := ?\<nu>\<rparr>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
939 |
using borel(1) pos(1) by (rule measure_space_density) simp |
40859 | 940 |
have ac: "absolutely_continuous ?\<nu>" |
941 |
using f by (rule density_is_absolutely_continuous) |
|
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
|
942 |
from split_space_into_finite_sets_and_rest[OF `measure_space (M\<lparr> measure := ?\<nu>\<rparr>)` ac] |
40859 | 943 |
obtain Q0 and Q :: "nat \<Rightarrow> 'a set" |
944 |
where Q: "disjoint_family Q" "range Q \<subseteq> sets M" |
|
945 |
and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
946 |
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
947 |
and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<infinity>" by force |
40859 | 948 |
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto |
949 |
let ?N = "{x\<in>space M. f x \<noteq> f' x}" |
|
950 |
have "?N \<in> sets M" using borel by auto |
|
43920 | 951 |
have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x" |
40859 | 952 |
unfolding indicator_def by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
953 |
have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos |
40859 | 954 |
by (intro finite_density_unique[THEN iffD1] allI) |
43920 | 955 |
(auto intro!: borel_measurable_ereal_times f Int simp: *) |
41705 | 956 |
moreover have "AE x. ?f Q0 x = ?f' Q0 x" |
40859 | 957 |
proof (rule AE_I') |
43920 | 958 |
{ fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M" |
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
|
959 |
and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
43923 | 960 |
let "?A i" = "Q0 \<inter> {x \<in> space M. f x < (i::nat)}" |
40859 | 961 |
have "(\<Union>i. ?A i) \<in> null_sets" |
962 |
proof (rule null_sets_UN) |
|
43923 | 963 |
fix i ::nat have "?A i \<in> sets M" |
40859 | 964 |
using borel Q0(1) by auto |
43923 | 965 |
have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)" |
40859 | 966 |
unfolding eq[OF `?A i \<in> sets M`] |
967 |
by (auto intro!: positive_integral_mono simp: indicator_def) |
|
43923 | 968 |
also have "\<dots> = i * \<mu> (?A i)" |
40859 | 969 |
using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
970 |
also have "\<dots> < \<infinity>" |
40859 | 971 |
using `?A i \<in> sets M`[THEN finite_measure] by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
972 |
finally have "?\<nu> (?A i) \<noteq> \<infinity>" by simp |
40859 | 973 |
then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto |
974 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
975 |
also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
976 |
by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
977 |
finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" by simp } |
40859 | 978 |
from this[OF borel(1) refl] this[OF borel(2) f] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
979 |
have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets" by simp_all |
42866 | 980 |
then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets" by (rule nullsets.Un) |
40859 | 981 |
show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq> |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
982 |
(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def) |
40859 | 983 |
qed |
41705 | 984 |
moreover have "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow> |
40859 | 985 |
?f (space M) x = ?f' (space M) x" |
986 |
by (auto simp: indicator_def Q0) |
|
41705 | 987 |
ultimately have "AE x. ?f (space M) x = ?f' (space M) x" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
988 |
by (auto simp: AE_all_countable[symmetric]) |
41705 | 989 |
then show "AE x. f x = f' x" by auto |
40859 | 990 |
qed |
991 |
||
992 |
lemma (in sigma_finite_measure) density_unique: |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
993 |
assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
994 |
assumes f': "f' \<in> borel_measurable M" "AE x. 0 \<le> f' x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
995 |
assumes eq: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)" |
40859 | 996 |
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A") |
997 |
shows "AE x. f x = f' x" |
|
998 |
proof - |
|
999 |
obtain h where h_borel: "h \<in> borel_measurable M" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1000 |
and fin: "integral\<^isup>P 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 | 1001 |
using Ex_finite_integrable_function by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1002 |
then have h_nn: "AE x. 0 \<le> h x" by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1003 |
let ?H = "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1004 |
have H: "measure_space ?H" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1005 |
using h_borel h_nn by (rule measure_space_density) simp |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1006 |
then interpret h: measure_space ?H . |
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
|
1007 |
interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>" |
45777
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
hoelzl
parents:
45769
diff
changeset
|
1008 |
by (intro H finite_measureI) (simp cong: positive_integral_cong add: fin) |
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
|
1009 |
let ?fM = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)\<rparr>" |
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
|
1010 |
interpret f: measure_space ?fM |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1011 |
using f by (rule measure_space_density) simp |
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
|
1012 |
let ?f'M = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)\<rparr>" |
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
|
1013 |
interpret f': measure_space ?f'M |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1014 |
using f' by (rule measure_space_density) simp |
40859 | 1015 |
{ 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
|
1016 |
then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1017 |
using pos(1) sets_into_space by (force simp: indicator_def) |
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
|
1018 |
then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1019 |
using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto } |
40859 | 1020 |
note h_null_sets = this |
1021 |
{ 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
|
1022 |
have "(\<integral>\<^isup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1023 |
using `A \<in> sets M` h_borel h_nn f f' |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1024 |
by (intro positive_integral_translated_density[symmetric]) auto |
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
|
1025 |
also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1026 |
by (rule f'.positive_integral_cong_measure) (simp_all add: eq) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1027 |
also have "\<dots> = (\<integral>\<^isup>+x. f' x * (h x * indicator A x) \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1028 |
using `A \<in> sets M` h_borel h_nn f f' |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1029 |
by (intro positive_integral_translated_density) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1030 |
finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1031 |
by (simp add: ac_simps) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1032 |
then have "(\<integral>\<^isup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^isup>+x. (f' x * indicator A x) \<partial>?H)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1033 |
using `A \<in> sets M` h_borel h_nn f f' |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1034 |
by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1035 |
then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1036 |
by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1037 |
simp_all |
40859 | 1038 |
then show "AE x. f x = f' x" |
1039 |
unfolding h.almost_everywhere_def almost_everywhere_def |
|
1040 |
by (auto simp add: h_null_sets) |
|
1041 |
qed |
|
1042 |
||
1043 |
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: |
|
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
|
1044 |
assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1045 |
and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" |
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
|
1046 |
and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+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
|
1047 |
shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<infinity>)" |
40859 | 1048 |
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
|
1049 |
assume "sigma_finite_measure ?N" |
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
|
1050 |
then interpret \<nu>: sigma_finite_measure ?N |
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
|
1051 |
where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>" |
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
|
1052 |
and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def) |
40859 | 1053 |
from \<nu>.Ex_finite_integrable_function obtain h where |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1054 |
h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<infinity>" and |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1055 |
h_nn: "\<And>x. 0 \<le> h x" and |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1056 |
fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1057 |
have "AE x. f x * h x \<noteq> \<infinity>" |
40859 | 1058 |
proof (rule AE_I') |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1059 |
have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn |
41705 | 1060 |
by (subst \<nu>.positive_integral_cong_measure[symmetric, |
1061 |
of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"]) |
|
1062 |
(auto intro!: positive_integral_translated_density simp: eq) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1063 |
then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<infinity>" |
40859 | 1064 |
using h(2) by simp |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1065 |
then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1066 |
using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage) |
40859 | 1067 |
qed auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1068 |
then show "AE x. f x \<noteq> \<infinity>" |
41705 | 1069 |
using fin by (auto elim!: AE_Ball_mp) |
40859 | 1070 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1071 |
assume AE: "AE x. f x \<noteq> \<infinity>" |
40859 | 1072 |
from sigma_finite guess Q .. note Q = this |
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
|
1073 |
interpret \<nu>: measure_space ?N |
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
|
1074 |
where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>" |
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
|
1075 |
and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def) |
43923 | 1076 |
def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M" |
40859 | 1077 |
{ fix i j have "A i \<inter> Q j \<in> sets M" |
1078 |
unfolding A_def using f Q |
|
1079 |
apply (rule_tac Int) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1080 |
by (cases i) (auto intro: measurable_sets[OF f(1)]) } |
40859 | 1081 |
note A_in_sets = this |
1082 |
let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j" |
|
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
|
1083 |
show "sigma_finite_measure ?N" |
40859 | 1084 |
proof (default, intro exI conjI subsetI allI) |
1085 |
fix x assume "x \<in> range ?A" |
|
1086 |
then obtain n where n: "x = ?A n" by auto |
|
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
|
1087 |
then show "x \<in> sets ?N" using A_in_sets by (cases "prod_decode n") auto |
40859 | 1088 |
next |
1089 |
have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)" |
|
1090 |
proof safe |
|
1091 |
fix x i j assume "x \<in> A i" "x \<in> Q j" |
|
1092 |
then show "x \<in> (\<Union>i. case prod_decode i of (i, j) \<Rightarrow> A i \<inter> Q j)" |
|
1093 |
by (intro UN_I[of "prod_encode (i,j)"]) auto |
|
1094 |
qed auto |
|
1095 |
also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto |
|
1096 |
also have "(\<Union>i. A i) = space M" |
|
1097 |
proof safe |
|
1098 |
fix x assume x: "x \<in> space M" |
|
1099 |
show "x \<in> (\<Union>i. A i)" |
|
1100 |
proof (cases "f x") |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1101 |
case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0]) |
40859 | 1102 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1103 |
case (real r) |
43923 | 1104 |
with less_PInf_Ex_of_nat[of "f x"] obtain n :: nat where "f x < n" by (auto simp: real_eq_of_nat) |
45769
2d5b1af2426a
real is better supported than real_of_nat, use it in the nat => ereal coercion
hoelzl
parents:
44941
diff
changeset
|
1105 |
then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"] simp: real_eq_of_nat) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1106 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1107 |
case MInf with x show ?thesis |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1108 |
unfolding A_def by (auto intro!: exI[of _ "Suc 0"]) |
40859 | 1109 |
qed |
1110 |
qed (auto simp: A_def) |
|
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
|
1111 |
finally show "(\<Union>i. ?A i) = space ?N" by simp |
40859 | 1112 |
next |
1113 |
fix n obtain i j where |
|
1114 |
[simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1115 |
have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>" |
40859 | 1116 |
proof (cases i) |
1117 |
case 0 |
|
1118 |
have "AE x. f x * indicator (A i \<inter> Q j) x = 0" |
|
41705 | 1119 |
using AE by (auto simp: A_def `i = 0`) |
1120 |
from positive_integral_cong_AE[OF this] show ?thesis by simp |
|
40859 | 1121 |
next |
1122 |
case (Suc n) |
|
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
|
1123 |
then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le> |
43923 | 1124 |
(\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)" |
45769
2d5b1af2426a
real is better supported than real_of_nat, use it in the nat => ereal coercion
hoelzl
parents:
44941
diff
changeset
|
1125 |
by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat) |
43923 | 1126 |
also have "\<dots> = Suc n * \<mu> (Q j)" |
40859 | 1127 |
using Q by (auto intro!: positive_integral_cmult_indicator) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1128 |
also have "\<dots> < \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1129 |
using Q by (auto simp: real_eq_of_nat[symmetric]) |
40859 | 1130 |
finally show ?thesis by simp |
1131 |
qed |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1132 |
then show "measure ?N (?A n) \<noteq> \<infinity>" |
40859 | 1133 |
using A_in_sets Q eq by auto |
1134 |
qed |
|
1135 |
qed |
|
1136 |
||
40871 | 1137 |
section "Radon-Nikodym derivative" |
38656 | 1138 |
|
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
|
1139 |
definition |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1140 |
"RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> |
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
|
1141 |
(\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))" |
38656 | 1142 |
|
40859 | 1143 |
lemma (in sigma_finite_measure) RN_deriv_cong: |
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
|
1144 |
assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M" |
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
|
1145 |
and \<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A" |
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
|
1146 |
shows "RN_deriv M' \<nu>' x = RN_deriv M \<nu> x" |
40859 | 1147 |
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
|
1148 |
interpret \<mu>': sigma_finite_measure M' |
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
|
1149 |
using cong by (rule sigma_finite_measure_cong) |
40859 | 1150 |
show ?thesis |
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
|
1151 |
unfolding RN_deriv_def |
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
|
1152 |
by (simp add: cong \<nu> positive_integral_cong_measure[OF cong] measurable_def) |
40859 | 1153 |
qed |
1154 |
||
38656 | 1155 |
lemma (in sigma_finite_measure) RN_deriv: |
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
|
1156 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" |
38656 | 1157 |
assumes "absolutely_continuous \<nu>" |
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
|
1158 |
shows "RN_deriv M \<nu> \<in> borel_measurable M" (is ?borel) |
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
|
1159 |
and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)" |
38656 | 1160 |
(is "\<And>A. _ \<Longrightarrow> ?int A") |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1161 |
and "0 \<le> RN_deriv M \<nu> x" |
38656 | 1162 |
proof - |
1163 |
note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1164 |
then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1165 |
from Ex show "0 \<le> RN_deriv M \<nu> x" unfolding RN_deriv_def |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1166 |
by (rule someI2_ex) simp |
38656 | 1167 |
fix A assume "A \<in> sets M" |
1168 |
from Ex show "?int A" unfolding RN_deriv_def |
|
1169 |
by (rule someI2_ex) (simp add: `A \<in> sets M`) |
|
1170 |
qed |
|
1171 |
||
40859 | 1172 |
lemma (in sigma_finite_measure) RN_deriv_positive_integral: |
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
|
1173 |
assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" |
40859 | 1174 |
and f: "f \<in> borel_measurable M" |
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
|
1175 |
shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)" |
40859 | 1176 |
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
|
1177 |
interpret \<nu>: measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1178 |
note RN = RN_deriv[OF \<nu>] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1179 |
have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1180 |
unfolding positive_integral_max_0 .. |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1181 |
also have "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>) = |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1182 |
(\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1183 |
by (intro \<nu>.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2)) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1184 |
also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * max 0 (f x) \<partial>M)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1185 |
by (intro positive_integral_translated_density) (auto simp add: RN f) |
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
|
1186 |
also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1187 |
using RN_deriv(3)[OF \<nu>] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1188 |
by (auto intro!: positive_integral_cong_pos split: split_if_asm |
43920 | 1189 |
simp: max_def ereal_mult_le_0_iff) |
40859 | 1190 |
finally show ?thesis . |
1191 |
qed |
|
1192 |
||
1193 |
lemma (in sigma_finite_measure) RN_deriv_unique: |
|
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
|
1194 |
assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1195 |
and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" |
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
|
1196 |
and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" |
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
|
1197 |
shows "AE x. f x = RN_deriv M \<nu> x" |
40859 | 1198 |
proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1199 |
show "AE x. 0 \<le> RN_deriv M \<nu> x" using RN_deriv[OF \<nu>] by auto |
40859 | 1200 |
fix A assume A: "A \<in> sets M" |
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
|
1201 |
show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)" |
40859 | 1202 |
unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] .. |
1203 |
qed |
|
1204 |
||
41832 | 1205 |
lemma (in sigma_finite_measure) RN_deriv_vimage: |
1206 |
assumes T: "T \<in> measure_preserving M M'" |
|
1207 |
and T': "T' \<in> measure_preserving (M'\<lparr> measure := \<nu>' \<rparr>) (M\<lparr> measure := \<nu> \<rparr>)" |
|
1208 |
and inv: "\<And>x. x \<in> space M \<Longrightarrow> T' (T x) = x" |
|
1209 |
and \<nu>': "measure_space (M'\<lparr>measure := \<nu>'\<rparr>)" "measure_space.absolutely_continuous M' \<nu>'" |
|
1210 |
shows "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x" |
|
1211 |
proof (rule RN_deriv_unique) |
|
1212 |
interpret \<nu>': measure_space "M'\<lparr>measure := \<nu>'\<rparr>" by fact |
|
1213 |
show "measure_space (M\<lparr> measure := \<nu>\<rparr>)" |
|
1214 |
by (rule \<nu>'.measure_space_vimage[OF _ T'], simp) default |
|
1215 |
interpret M': measure_space M' |
|
1216 |
proof (rule measure_space_vimage) |
|
1217 |
have "sigma_algebra (M'\<lparr> measure := \<nu>'\<rparr>)" by default |
|
1218 |
then show "sigma_algebra M'" by simp |
|
1219 |
qed fact |
|
1220 |
show "absolutely_continuous \<nu>" unfolding absolutely_continuous_def |
|
1221 |
proof safe |
|
1222 |
fix N assume N: "N \<in> sets M" and N_0: "\<mu> N = 0" |
|
1223 |
then have N': "T' -` N \<inter> space M' \<in> sets M'" |
|
1224 |
using T' by (auto simp: measurable_def measure_preserving_def) |
|
1225 |
have "T -` (T' -` N \<inter> space M') \<inter> space M = N" |
|
1226 |
using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def) |
|
1227 |
then have "measure M' (T' -` N \<inter> space M') = 0" |
|
1228 |
using measure_preservingD[OF T N'] N_0 by auto |
|
1229 |
with \<nu>'(2) N' show "\<nu> N = 0" using measure_preservingD[OF T', of N] N |
|
1230 |
unfolding M'.absolutely_continuous_def measurable_def by auto |
|
1231 |
qed |
|
1232 |
interpret M': sigma_finite_measure M' |
|
1233 |
proof |
|
1234 |
from sigma_finite guess F .. note F = this |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1235 |
show "\<exists>A::nat \<Rightarrow> 'c set. range A \<subseteq> sets M' \<and> (\<Union>i. A i) = space M' \<and> (\<forall>i. M'.\<mu> (A i) \<noteq> \<infinity>)" |
41832 | 1236 |
proof (intro exI conjI allI) |
1237 |
show *: "range (\<lambda>i. T' -` F i \<inter> space M') \<subseteq> sets M'" |
|
1238 |
using F T' by (auto simp: measurable_def measure_preserving_def) |
|
1239 |
show "(\<Union>i. T' -` F i \<inter> space M') = space M'" |
|
1240 |
using F T' by (force simp: measurable_def measure_preserving_def) |
|
1241 |
fix i |
|
1242 |
have "T' -` F i \<inter> space M' \<in> sets M'" using * by auto |
|
1243 |
note measure_preservingD[OF T this, symmetric] |
|
1244 |
moreover |
|
1245 |
have Fi: "F i \<in> sets M" using F by auto |
|
1246 |
then have "T -` (T' -` F i \<inter> space M') \<inter> space M = F i" |
|
1247 |
using T inv sets_into_space[OF Fi] |
|
1248 |
by (auto simp: measurable_def measure_preserving_def) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1249 |
ultimately show "measure M' (T' -` F i \<inter> space M') \<noteq> \<infinity>" |
41832 | 1250 |
using F by simp |
1251 |
qed |
|
1252 |
qed |
|
1253 |
have "(RN_deriv M' \<nu>') \<circ> T \<in> borel_measurable M" |
|
1254 |
by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+ |
|
1255 |
then show "(\<lambda>x. RN_deriv M' \<nu>' (T x)) \<in> borel_measurable M" |
|
1256 |
by (simp add: comp_def) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1257 |
show "AE x. 0 \<le> RN_deriv M' \<nu>' (T x)" using M'.RN_deriv(3)[OF \<nu>'] by auto |
41832 | 1258 |
fix A let ?A = "T' -` A \<inter> space M'" |
1259 |
assume A: "A \<in> sets M" |
|
1260 |
then have A': "?A \<in> sets M'" using T' unfolding measurable_def measure_preserving_def |
|
1261 |
by auto |
|
1262 |
from A have "\<nu> A = \<nu>' ?A" using T'[THEN measure_preservingD] by simp |
|
1263 |
also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' x * indicator ?A x \<partial>M'" |
|
1264 |
using A' by (rule M'.RN_deriv(2)[OF \<nu>']) |
|
1265 |
also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator ?A (T x) \<partial>M" |
|
1266 |
proof (rule positive_integral_vimage) |
|
1267 |
show "sigma_algebra M'" by default |
|
1268 |
show "(\<lambda>x. RN_deriv M' \<nu>' x * indicator (T' -` A \<inter> space M') x) \<in> borel_measurable M'" |
|
1269 |
by (auto intro!: A' M'.RN_deriv(1)[OF \<nu>']) |
|
1270 |
qed fact |
|
1271 |
also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M" |
|
1272 |
using T inv by (auto intro!: positive_integral_cong simp: measure_preserving_def measurable_def indicator_def) |
|
1273 |
finally show "\<nu> A = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M" . |
|
1274 |
qed |
|
1275 |
||
40859 | 1276 |
lemma (in sigma_finite_measure) RN_deriv_finite: |
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
|
1277 |
assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<nu>" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1278 |
shows "AE x. RN_deriv M \<nu> x \<noteq> \<infinity>" |
40859 | 1279 |
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
|
1280 |
interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact |
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
|
1281 |
have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default |
40859 | 1282 |
from sfm show ?thesis |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1283 |
using sigma_finite_iff_density_finite[OF \<nu> RN_deriv(1)[OF \<nu> ac]] RN_deriv(2,3)[OF \<nu> ac] by simp |
40859 | 1284 |
qed |
1285 |
||
1286 |
lemma (in sigma_finite_measure) |
|
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
|
1287 |
assumes \<nu>: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" |
40859 | 1288 |
and f: "f \<in> borel_measurable M" |
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
|
1289 |
shows RN_deriv_integrable: "integrable (M\<lparr>measure := \<nu>\<rparr>) f \<longleftrightarrow> |
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
|
1290 |
integrable M (\<lambda>x. real (RN_deriv M \<nu> x) * f x)" (is ?integrable) |
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
|
1291 |
and RN_deriv_integral: "integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) f = |
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
|
1292 |
(\<integral>x. real (RN_deriv M \<nu> x) * f x \<partial>M)" (is ?integral) |
40859 | 1293 |
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
|
1294 |
interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact |
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
|
1295 |
have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default |
43920 | 1296 |
have minus_cong: "\<And>A B A' B'::ereal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp |
40859 | 1297 |
have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1298 |
have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp |
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
|
1299 |
{ fix f :: "'a \<Rightarrow> real" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1300 |
{ fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>" |
43920 | 1301 |
have "ereal (real (RN_deriv M \<nu> x)) * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)" |
40859 | 1302 |
by (simp add: mult_le_0_iff) |
43920 | 1303 |
then have "RN_deriv M \<nu> x * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)" |
1304 |
using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: ereal_real split: split_if_asm) } |
|
1305 |
then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (f x) \<partial>M)" |
|
1306 |
"(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (- f x) \<partial>M)" |
|
1307 |
using RN_deriv_finite[OF \<nu>] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric] |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1308 |
by (auto intro!: positive_integral_cong_AE) } |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1309 |
note * = this |
40859 | 1310 |
show ?integral ?integrable |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1311 |
unfolding lebesgue_integral_def integrable_def * |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1312 |
using f RN_deriv(1)[OF ms \<nu>(2)] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1313 |
by (auto simp: RN_deriv_positive_integral[OF ms \<nu>(2)]) |
40859 | 1314 |
qed |
1315 |
||
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1316 |
lemma (in sigma_finite_measure) real_RN_deriv: |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1317 |
assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>") |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1318 |
assumes ac: "absolutely_continuous \<nu>" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1319 |
obtains D where "D \<in> borel_measurable M" |
43920 | 1320 |
and "AE x. RN_deriv M \<nu> x = ereal (D x)" |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1321 |
and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1322 |
and "\<And>x. 0 \<le> D x" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1323 |
proof |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1324 |
interpret \<nu>: finite_measure ?\<nu> by fact |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1325 |
have ms: "measure_space ?\<nu>" by default |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1326 |
note RN = RN_deriv[OF ms ac] |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1327 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1328 |
let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M \<nu> x = t}" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1329 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1330 |
show "(\<lambda>x. real (RN_deriv M \<nu> x)) \<in> borel_measurable M" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1331 |
using RN by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1332 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1333 |
have "\<nu> (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN \<infinity>) x \<partial>M)" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1334 |
using RN by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1335 |
also have "\<dots> = (\<integral>\<^isup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1336 |
by (intro positive_integral_cong) (auto simp: indicator_def) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1337 |
also have "\<dots> = \<infinity> * \<mu> (?RN \<infinity>)" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1338 |
using RN by (intro positive_integral_cmult_indicator) auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1339 |
finally have eq: "\<nu> (?RN \<infinity>) = \<infinity> * \<mu> (?RN \<infinity>)" . |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1340 |
moreover |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1341 |
have "\<mu> (?RN \<infinity>) = 0" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1342 |
proof (rule ccontr) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1343 |
assume "\<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>} \<noteq> 0" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1344 |
moreover from RN have "0 \<le> \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1345 |
ultimately have "0 < \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1346 |
with eq have "\<nu> (?RN \<infinity>) = \<infinity>" by simp |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1347 |
with \<nu>.finite_measure[of "?RN \<infinity>"] RN show False by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1348 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1349 |
ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1350 |
using RN by (intro AE_iff_measurable[THEN iffD2]) auto |
43920 | 1351 |
then show "AE x. RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))" |
1352 |
using RN(3) by (auto simp: ereal_real) |
|
1353 |
then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))" |
|
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1354 |
using ac absolutely_continuous_AE[OF ms] by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1355 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1356 |
show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)" |
43920 | 1357 |
using RN by (auto intro: real_of_ereal_pos) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1358 |
|
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1359 |
have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1360 |
using RN by auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1361 |
also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1362 |
by (intro positive_integral_cong) (auto simp: indicator_def) |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1363 |
finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0" |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1364 |
using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1365 |
with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)" |
43920 | 1366 |
by (auto simp: zero_less_real_of_ereal le_less) |
43340
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1367 |
qed |
60e181c4eae4
lemma: independence is equal to mutual information = 0
hoelzl
parents:
42866
diff
changeset
|
1368 |
|
38656 | 1369 |
lemma (in sigma_finite_measure) RN_deriv_singleton: |
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
|
1370 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" |
38656 | 1371 |
and ac: "absolutely_continuous \<nu>" |
1372 |
and "{x} \<in> sets M" |
|
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
|
1373 |
shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}" |
38656 | 1374 |
proof - |
1375 |
note deriv = RN_deriv[OF assms(1, 2)] |
|
1376 |
from deriv(2)[OF `{x} \<in> sets M`] |
|
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
|
1377 |
have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)" |
38656 | 1378 |
by (auto simp: indicator_def intro!: positive_integral_cong) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41832
diff
changeset
|
1379 |
thus ?thesis using positive_integral_cmult_indicator[OF _ `{x} \<in> sets M`] deriv(3) |
38656 | 1380 |
by auto |
1381 |
qed |
|
1382 |
||
1383 |
theorem (in finite_measure_space) RN_deriv_finite_measure: |
|
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
|
1384 |
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" |
38656 | 1385 |
and ac: "absolutely_continuous \<nu>" |
1386 |
and "x \<in> space M" |
|
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
|
1387 |
shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}" |
38656 | 1388 |
proof - |
1389 |
have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto |
|
1390 |
from RN_deriv_singleton[OF assms(1,2) this] show ?thesis . |
|
1391 |
qed |
|
1392 |
||
1393 |
end |