| author | blanchet | 
| Mon, 17 Feb 2014 13:31:42 +0100 | |
| changeset 55534 | b18bdcbda41b | 
| parent 55417 | 01fbfb60c33e | 
| child 55642 | 63beb38e9258 | 
| 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  | 
||
| 47694 | 11  | 
definition "diff_measure M N =  | 
12  | 
measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"  | 
|
13  | 
||
14  | 
lemma  | 
|
15  | 
shows space_diff_measure[simp]: "space (diff_measure M N) = space M"  | 
|
16  | 
and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M"  | 
|
17  | 
by (auto simp: diff_measure_def)  | 
|
18  | 
||
19  | 
lemma emeasure_diff_measure:  | 
|
20  | 
assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N"  | 
|
21  | 
assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M"  | 
|
22  | 
shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A")  | 
|
23  | 
unfolding diff_measure_def  | 
|
24  | 
proof (rule emeasure_measure_of_sigma)  | 
|
25  | 
show "sigma_algebra (space M) (sets M)" ..  | 
|
26  | 
show "positive (sets M) ?\<mu>"  | 
|
27  | 
using pos by (simp add: positive_def ereal_diff_positive)  | 
|
28  | 
show "countably_additive (sets M) ?\<mu>"  | 
|
29  | 
proof (rule countably_additiveI)  | 
|
30  | 
fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> sets M" and "disjoint_family A"  | 
|
31  | 
then have suminf:  | 
|
32  | 
"(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"  | 
|
33  | 
"(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)"  | 
|
34  | 
by (simp_all add: suminf_emeasure sets_eq)  | 
|
35  | 
with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =  | 
|
36  | 
(\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"  | 
|
37  | 
using fin  | 
|
38  | 
by (intro suminf_ereal_minus pos emeasure_nonneg)  | 
|
39  | 
(auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)  | 
|
40  | 
then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =  | 
|
41  | 
emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "  | 
|
42  | 
by (simp add: suminf)  | 
|
43  | 
qed  | 
|
44  | 
qed fact  | 
|
45  | 
||
| 38656 | 46  | 
lemma (in sigma_finite_measure) Ex_finite_integrable_function:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
47  | 
shows "\<exists>h\<in>borel_measurable M. integral\<^sup>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 | 48  | 
proof -  | 
49  | 
obtain A :: "nat \<Rightarrow> 'a set" where  | 
|
| 50003 | 50  | 
range[measurable]: "range A \<subseteq> sets M" and  | 
| 38656 | 51  | 
space: "(\<Union>i. A i) = space M" and  | 
| 47694 | 52  | 
measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and  | 
| 38656 | 53  | 
disjoint: "disjoint_family A"  | 
| 47694 | 54  | 
using sigma_finite_disjoint by auto  | 
55  | 
let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"  | 
|
| 38656 | 56  | 
have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"  | 
57  | 
proof  | 
|
| 47694 | 58  | 
fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"  | 
59  | 
using measure[of i] emeasure_nonneg[of M "A i"]  | 
|
| 
51329
 
4a3c453f99a1
split dense into inner_dense_order and no_top/no_bot
 
hoelzl 
parents: 
50244 
diff
changeset
 | 
60  | 
by (auto intro!: dense simp: ereal_0_gt_inverse ereal_zero_le_0_iff)  | 
| 38656 | 61  | 
qed  | 
62  | 
from choice[OF this] obtain n where n: "\<And>i. 0 < n i"  | 
|
| 47694 | 63  | 
"\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
64  | 
  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
 | 
| 46731 | 65  | 
let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x"  | 
| 38656 | 66  | 
show ?thesis  | 
67  | 
proof (safe intro!: bexI[of _ ?h] del: notI)  | 
|
| 39092 | 68  | 
have "\<And>i. A i \<in> sets M"  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
69  | 
using range by fastforce+  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
70  | 
then have "integral\<^sup>P M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
71  | 
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
 | 
72  | 
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
 | 
73  | 
proof (rule suminf_le_pos)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
74  | 
fix N  | 
| 47694 | 75  | 
have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"  | 
76  | 
using n[of N]  | 
|
| 43920 | 77  | 
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
 | 
78  | 
also have "\<dots> \<le> (1 / 2) ^ Suc N"  | 
| 38656 | 79  | 
using measure[of N] n[of N]  | 
| 47694 | 80  | 
by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"])  | 
| 43920 | 81  | 
(simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)  | 
| 47694 | 82  | 
finally show "n N * emeasure M (A N) \<le> (1 / 2) ^ Suc N" .  | 
83  | 
show "0 \<le> n N * emeasure M (A N)" using n[of N] `A N \<in> sets M` by (simp add: emeasure_nonneg)  | 
|
| 38656 | 84  | 
qed  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
85  | 
finally show "integral\<^sup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto  | 
| 38656 | 86  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
87  | 
    { 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
 | 
88  | 
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
 | 
89  | 
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
 | 
90  | 
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
 | 
91  | 
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
 | 
92  | 
note pos = this  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
93  | 
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
 | 
94  | 
proof cases  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
95  | 
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
 | 
96  | 
next  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
97  | 
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
 | 
98  | 
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
 | 
99  | 
qed  | 
| 50003 | 100  | 
qed measurable  | 
| 38656 | 101  | 
qed  | 
102  | 
||
| 40871 | 103  | 
subsection "Absolutely continuous"  | 
104  | 
||
| 47694 | 105  | 
definition absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where  | 
106  | 
"absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"  | 
|
107  | 
||
108  | 
lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"  | 
|
109  | 
unfolding absolutely_continuous_def by (auto simp: null_sets_count_space)  | 
|
| 38656 | 110  | 
|
| 47694 | 111  | 
lemma absolutely_continuousI_density:  | 
112  | 
"f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)"  | 
|
113  | 
by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in)  | 
|
114  | 
||
115  | 
lemma absolutely_continuousI_point_measure_finite:  | 
|
116  | 
"(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)"  | 
|
117  | 
unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff)  | 
|
118  | 
||
119  | 
lemma absolutely_continuous_AE:  | 
|
120  | 
assumes sets_eq: "sets M' = sets M"  | 
|
121  | 
and "absolutely_continuous M M'" "AE x in M. P x"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
122  | 
shows "AE x in M'. P x"  | 
| 40859 | 123  | 
proof -  | 
| 47694 | 124  | 
  from `AE x in M. P x` obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
 | 
125  | 
unfolding eventually_ae_filter by auto  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
126  | 
show "AE x in M'. P x"  | 
| 47694 | 127  | 
proof (rule AE_I')  | 
128  | 
    show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
 | 
|
129  | 
from `absolutely_continuous M M'` show "N \<in> null_sets M'"  | 
|
130  | 
using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto  | 
|
| 40859 | 131  | 
qed  | 
132  | 
qed  | 
|
133  | 
||
| 40871 | 134  | 
subsection "Existence of the Radon-Nikodym derivative"  | 
135  | 
||
| 38656 | 136  | 
lemma (in finite_measure) Radon_Nikodym_aux_epsilon:  | 
137  | 
fixes e :: real assumes "0 < e"  | 
|
| 47694 | 138  | 
assumes "finite_measure N" and sets_eq: "sets N = sets M"  | 
139  | 
shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le> measure M A - measure N A \<and>  | 
|
140  | 
(\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < measure M B - measure N B)"  | 
|
| 38656 | 141  | 
proof -  | 
| 47694 | 142  | 
interpret M': finite_measure N by fact  | 
143  | 
let ?d = "\<lambda>A. measure M A - measure N A"  | 
|
| 46731 | 144  | 
let ?A = "\<lambda>A. if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)  | 
| 38656 | 145  | 
    then {}
 | 
146  | 
else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"  | 
|
147  | 
  def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
 | 
|
148  | 
have A_simps[simp]:  | 
|
149  | 
    "A 0 = {}"
 | 
|
150  | 
"\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all  | 
|
151  | 
  { fix A assume "A \<in> sets M"
 | 
|
152  | 
have "?A A \<in> sets M"  | 
|
153  | 
by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) }  | 
|
154  | 
note A'_in_sets = this  | 
|
155  | 
  { fix n have "A n \<in> sets M"
 | 
|
156  | 
proof (induct n)  | 
|
157  | 
case (Suc n) thus "A (Suc n) \<in> sets M"  | 
|
158  | 
using A'_in_sets[of "A n"] by (auto split: split_if_asm)  | 
|
159  | 
qed (simp add: A_def) }  | 
|
160  | 
note A_in_sets = this  | 
|
161  | 
hence "range A \<subseteq> sets M" by auto  | 
|
162  | 
  { fix n B
 | 
|
163  | 
assume Ex: "\<exists>B. B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> -e"  | 
|
164  | 
hence False: "\<not> (\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B)" by (auto simp: not_less)  | 
|
165  | 
have "?d (A (Suc n)) \<le> ?d (A n) - e" unfolding A_simps if_not_P[OF False]  | 
|
166  | 
proof (rule someI2_ex[OF Ex])  | 
|
167  | 
fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"  | 
|
168  | 
      hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
 | 
|
169  | 
hence "?d (A n \<union> B) = ?d (A n) + ?d B"  | 
|
| 47694 | 170  | 
using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq)  | 
| 38656 | 171  | 
also have "\<dots> \<le> ?d (A n) - e" using dB by simp  | 
172  | 
finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .  | 
|
173  | 
qed }  | 
|
174  | 
note dA_epsilon = this  | 
|
175  | 
  { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
 | 
|
176  | 
proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e")  | 
|
177  | 
case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp  | 
|
178  | 
next  | 
|
179  | 
case False  | 
|
180  | 
hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le)  | 
|
181  | 
thus ?thesis by simp  | 
|
182  | 
qed }  | 
|
183  | 
note dA_mono = this  | 
|
184  | 
show ?thesis  | 
|
185  | 
proof (cases "\<exists>n. \<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B")  | 
|
186  | 
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  | 
|
187  | 
show ?thesis  | 
|
188  | 
proof (safe intro!: bexI[of _ "space M - A n"])  | 
|
189  | 
fix B assume "B \<in> sets M" "B \<subseteq> space M - A n"  | 
|
190  | 
from B[OF this] show "-e < ?d B" .  | 
|
191  | 
next  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
192  | 
show "space M - A n \<in> sets M" by (rule sets.compl_sets) fact  | 
| 38656 | 193  | 
next  | 
194  | 
show "?d (space M) \<le> ?d (space M - A n)"  | 
|
195  | 
proof (induct n)  | 
|
196  | 
fix n assume "?d (space M) \<le> ?d (space M - A n)"  | 
|
197  | 
also have "\<dots> \<le> ?d (space M - A (Suc n))"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
198  | 
using A_in_sets sets.sets_into_space dA_mono[of n] finite_measure_compl M'.finite_measure_compl  | 
| 47694 | 199  | 
by (simp del: A_simps add: sets_eq sets_eq_imp_space_eq[OF sets_eq])  | 
| 38656 | 200  | 
finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .  | 
201  | 
qed simp  | 
|
202  | 
qed  | 
|
203  | 
next  | 
|
204  | 
case False hence B: "\<And>n. \<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"  | 
|
205  | 
by (auto simp add: not_less)  | 
|
206  | 
    { fix n have "?d (A n) \<le> - real n * e"
 | 
|
207  | 
proof (induct n)  | 
|
208  | 
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
 | 
209  | 
next  | 
| 47694 | 210  | 
case 0 with measure_empty 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
 | 
211  | 
qed } note dA_less = this  | 
| 38656 | 212  | 
have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq  | 
213  | 
proof (rule incseq_SucI)  | 
|
214  | 
fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto  | 
|
215  | 
qed  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
216  | 
have A: "incseq A" by (auto intro!: incseq_SucI)  | 
| 47694 | 217  | 
from finite_Lim_measure_incseq[OF _ A] `range A \<subseteq> sets M`  | 
218  | 
M'.finite_Lim_measure_incseq[OF _ A]  | 
|
| 38656 | 219  | 
have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"  | 
| 47694 | 220  | 
by (auto intro!: tendsto_diff simp: sets_eq)  | 
| 38656 | 221  | 
obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto  | 
222  | 
moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]  | 
|
223  | 
have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps)  | 
|
224  | 
ultimately show ?thesis by auto  | 
|
225  | 
qed  | 
|
226  | 
qed  | 
|
227  | 
||
| 47694 | 228  | 
lemma (in finite_measure) Radon_Nikodym_aux:  | 
229  | 
assumes "finite_measure N" and sets_eq: "sets N = sets M"  | 
|
230  | 
shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le>  | 
|
231  | 
measure M A - measure N A \<and>  | 
|
232  | 
(\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> measure M B - measure N B)"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
233  | 
proof -  | 
| 47694 | 234  | 
interpret N: finite_measure N by fact  | 
235  | 
let ?d = "\<lambda>A. measure M A - measure N A"  | 
|
| 46731 | 236  | 
let ?P = "\<lambda>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)"  | 
237  | 
let ?r = "\<lambda>S. restricted_space S"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
238  | 
  { fix S n assume S: "S \<in> sets M"
 | 
| 47694 | 239  | 
then have "finite_measure (density M (indicator S))" "0 < 1 / real (Suc n)"  | 
240  | 
"finite_measure (density N (indicator S))" "sets (density N (indicator S)) = sets (density M (indicator S))"  | 
|
241  | 
by (auto simp: finite_measure_restricted N.finite_measure_restricted sets_eq)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
242  | 
from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this  | 
| 47694 | 243  | 
with S have "?P (S \<inter> X) S n"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
244  | 
by (simp add: measure_restricted sets_eq sets.Int) (metis inf_absorb2)  | 
| 47694 | 245  | 
hence "\<exists>A. ?P A S n" .. }  | 
| 38656 | 246  | 
note Ex_P = this  | 
| 55415 | 247  | 
def A \<equiv> "rec_nat (space M) (\<lambda>n A. SOME B. ?P B A n)"  | 
| 38656 | 248  | 
have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def)  | 
249  | 
have A_0[simp]: "A 0 = space M" unfolding A_def by simp  | 
|
250  | 
  { fix i have "A i \<in> sets M" unfolding A_def
 | 
|
251  | 
proof (induct i)  | 
|
252  | 
case (Suc i)  | 
|
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55415 
diff
changeset
 | 
253  | 
from Ex_P[OF this, of i] show ?case unfolding nat.recs(2)  | 
| 38656 | 254  | 
by (rule someI2_ex) simp  | 
255  | 
qed simp }  | 
|
256  | 
note A_in_sets = this  | 
|
257  | 
  { fix n have "?P (A (Suc n)) (A n) n"
 | 
|
258  | 
using Ex_P[OF A_in_sets] unfolding A_Suc  | 
|
259  | 
by (rule someI2_ex) simp }  | 
|
260  | 
note P_A = this  | 
|
261  | 
have "range A \<subseteq> sets M" using A_in_sets by auto  | 
|
262  | 
have A_mono: "\<And>i. A (Suc i) \<subseteq> A i" using P_A by simp  | 
|
263  | 
have mono_dA: "mono (\<lambda>i. ?d (A i))" using P_A by (simp add: mono_iff_le_Suc)  | 
|
264  | 
have epsilon: "\<And>C i. \<lbrakk> C \<in> sets M; C \<subseteq> A (Suc i) \<rbrakk> \<Longrightarrow> - 1 / real (Suc i) < ?d C"  | 
|
265  | 
using P_A by auto  | 
|
266  | 
show ?thesis  | 
|
267  | 
proof (safe intro!: bexI[of _ "\<Inter>i. A i"])  | 
|
268  | 
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
 | 
269  | 
have A: "decseq A" using A_mono by (auto intro!: decseq_SucI)  | 
| 47694 | 270  | 
from `range A \<subseteq> sets M`  | 
271  | 
finite_Lim_measure_decseq[OF _ A] N.finite_Lim_measure_decseq[OF _ A]  | 
|
272  | 
have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: tendsto_diff simp: sets_eq)  | 
|
| 38656 | 273  | 
thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]  | 
| 47694 | 274  | 
by (rule_tac LIMSEQ_le_const) auto  | 
| 38656 | 275  | 
next  | 
276  | 
fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"  | 
|
277  | 
show "0 \<le> ?d B"  | 
|
278  | 
proof (rule ccontr)  | 
|
279  | 
assume "\<not> 0 \<le> ?d B"  | 
|
280  | 
hence "0 < - ?d B" by auto  | 
|
281  | 
from ex_inverse_of_nat_Suc_less[OF this]  | 
|
282  | 
obtain n where *: "?d B < - 1 / real (Suc n)"  | 
|
283  | 
by (auto simp: real_eq_of_nat inverse_eq_divide field_simps)  | 
|
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
55415 
diff
changeset
 | 
284  | 
have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat.recs(2))  | 
| 38656 | 285  | 
from epsilon[OF B(1) this] *  | 
286  | 
show False by auto  | 
|
287  | 
qed  | 
|
288  | 
qed  | 
|
289  | 
qed  | 
|
290  | 
||
291  | 
lemma (in finite_measure) Radon_Nikodym_finite_measure:  | 
|
| 47694 | 292  | 
assumes "finite_measure N" and sets_eq: "sets N = sets M"  | 
293  | 
assumes "absolutely_continuous M N"  | 
|
294  | 
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"  | 
|
| 38656 | 295  | 
proof -  | 
| 47694 | 296  | 
interpret N: finite_measure N by fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
297  | 
  def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
 | 
| 50003 | 298  | 
  { fix f have "f \<in> G \<Longrightarrow> f \<in> borel_measurable M" by (auto simp: G_def) }
 | 
299  | 
note this[measurable_dest]  | 
|
| 38656 | 300  | 
have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto  | 
301  | 
  hence "G \<noteq> {}" by auto
 | 
|
302  | 
  { fix f g assume f: "f \<in> G" and g: "g \<in> G"
 | 
|
303  | 
have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def  | 
|
304  | 
proof safe  | 
|
305  | 
show "?max \<in> borel_measurable M" using f g unfolding G_def by auto  | 
|
306  | 
      let ?A = "{x \<in> space M. f x \<le> g x}"
 | 
|
307  | 
have "?A \<in> sets M" using f g unfolding G_def by auto  | 
|
308  | 
fix A assume "A \<in> sets M"  | 
|
309  | 
hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto  | 
|
| 47694 | 310  | 
hence sets': "?A \<inter> A \<in> sets N" "(space M - ?A) \<inter> A \<in> sets N" by (auto simp: sets_eq)  | 
| 38656 | 311  | 
have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
312  | 
using sets.sets_into_space[OF `A \<in> sets M`] by auto  | 
| 38656 | 313  | 
have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =  | 
314  | 
g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"  | 
|
315  | 
by (auto simp: indicator_def max_def)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
316  | 
hence "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
317  | 
(\<integral>\<^sup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
318  | 
(\<integral>\<^sup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"  | 
| 38656 | 319  | 
using f g sets unfolding G_def  | 
| 46731 | 320  | 
by (auto cong: positive_integral_cong intro!: positive_integral_add)  | 
| 47694 | 321  | 
also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"  | 
| 38656 | 322  | 
using f g sets unfolding G_def by (auto intro!: add_mono)  | 
| 47694 | 323  | 
also have "\<dots> = N A"  | 
324  | 
using plus_emeasure[OF sets'] union by auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
325  | 
finally show "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
326  | 
next  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
327  | 
fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max)  | 
| 38656 | 328  | 
qed }  | 
329  | 
note max_in_G = this  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
330  | 
  { fix f assume  "incseq f" and f: "\<And>i. f i \<in> G"
 | 
| 50003 | 331  | 
then have [measurable]: "\<And>i. f i \<in> borel_measurable M" by (auto simp: G_def)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
332  | 
have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def  | 
| 38656 | 333  | 
proof safe  | 
| 50003 | 334  | 
show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M" by measurable  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
335  | 
      { fix x show "0 \<le> (SUP i. f i x)"
 | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
336  | 
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
 | 
337  | 
next  | 
| 38656 | 338  | 
fix A assume "A \<in> sets M"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
339  | 
have "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
340  | 
(\<integral>\<^sup>+x. (SUP i. f i 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
 | 
341  | 
by (intro positive_integral_cong) (simp split: split_indicator)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
342  | 
also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i 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
 | 
343  | 
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
 | 
344  | 
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
 | 
345  | 
(auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
346  | 
finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
347  | 
using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def)  | 
| 38656 | 348  | 
qed }  | 
349  | 
note SUP_in_G = this  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
350  | 
let ?y = "SUP g : G. integral\<^sup>P M g"  | 
| 47694 | 351  | 
have y_le: "?y \<le> N (space M)" unfolding G_def  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
352  | 
proof (safe intro!: SUP_least)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
353  | 
fix g assume "\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
354  | 
from this[THEN bspec, OF sets.top] show "integral\<^sup>P M g \<le> N (space M)"  | 
| 38656 | 355  | 
by (simp cong: positive_integral_cong)  | 
356  | 
qed  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
357  | 
  from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^sup>P M"] guess ys .. note ys = this
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
358  | 
then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>P M g = ys n"  | 
| 38656 | 359  | 
proof safe  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
360  | 
fix n assume "range ys \<subseteq> integral\<^sup>P M ` G"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
361  | 
hence "ys n \<in> integral\<^sup>P M ` G" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
362  | 
thus "\<exists>g. g\<in>G \<and> integral\<^sup>P M g = ys n" by auto  | 
| 38656 | 363  | 
qed  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
364  | 
from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^sup>P M (gs n) = ys n" by auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
365  | 
hence y_eq: "?y = (SUP i. integral\<^sup>P M (gs i))" using ys by auto  | 
| 46731 | 366  | 
  let ?g = "\<lambda>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
 | 
367  | 
def f \<equiv> "\<lambda>x. SUP i. ?g i x"  | 
| 46731 | 368  | 
let ?F = "\<lambda>A x. f x * indicator A x"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
369  | 
  have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
 | 
| 38656 | 370  | 
  { fix i have "?g i \<in> G"
 | 
371  | 
proof (induct i)  | 
|
372  | 
case 0 thus ?case by simp fact  | 
|
373  | 
next  | 
|
374  | 
case (Suc i)  | 
|
375  | 
with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case  | 
|
376  | 
by (auto simp add: atMost_Suc intro!: max_in_G)  | 
|
377  | 
qed }  | 
|
378  | 
note g_in_G = this  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
379  | 
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
 | 
380  | 
by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)  | 
| 50003 | 381  | 
from SUP_in_G[OF this g_in_G] have [measurable]: "f \<in> G" unfolding f_def .  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
382  | 
then have [simp, intro]: "f \<in> borel_measurable M" unfolding G_def by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
383  | 
have "integral\<^sup>P M f = (SUP i. integral\<^sup>P M (?g i))" unfolding f_def  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
384  | 
using g_in_G `incseq ?g`  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
385  | 
by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def)  | 
| 38656 | 386  | 
also have "\<dots> = ?y"  | 
387  | 
proof (rule antisym)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
388  | 
show "(SUP i. integral\<^sup>P M (?g i)) \<le> ?y"  | 
| 47694 | 389  | 
using g_in_G by (auto intro: Sup_mono simp: SUP_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
390  | 
show "?y \<le> (SUP i. integral\<^sup>P M (?g i))" unfolding y_eq  | 
| 38656 | 391  | 
by (auto intro!: SUP_mono positive_integral_mono Max_ge)  | 
392  | 
qed  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
393  | 
finally have int_f_eq_y: "integral\<^sup>P M f = ?y" .  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
394  | 
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
 | 
395  | 
unfolding f_def using `\<And>i. gs i \<in> G`  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
396  | 
by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
397  | 
let ?t = "\<lambda>A. N A - (\<integral>\<^sup>+x. ?F A x \<partial>M)"  | 
| 47694 | 398  | 
let ?M = "diff_measure N (density M f)"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
399  | 
have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
400  | 
using `f \<in> G` unfolding G_def by auto  | 
| 47694 | 401  | 
have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure ?M A = ?t A"  | 
402  | 
proof (subst emeasure_diff_measure)  | 
|
403  | 
from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)"  | 
|
404  | 
by (auto intro!: finite_measureI simp: emeasure_density cong: positive_integral_cong)  | 
|
405  | 
next  | 
|
406  | 
fix B assume "B \<in> sets N" with f_le_N[of B] show "emeasure (density M f) B \<le> emeasure N B"  | 
|
407  | 
by (auto simp: sets_eq emeasure_density cong: positive_integral_cong)  | 
|
408  | 
qed (auto simp: sets_eq emeasure_density)  | 
|
409  | 
from emeasure_M[of "space M"] N.finite_emeasure_space positive_integral_positive[of M "?F (space M)"]  | 
|
410  | 
interpret M': finite_measure ?M  | 
|
411  | 
by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure ereal_minus_eq_PInfty_iff)  | 
|
412  | 
||
413  | 
have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def  | 
|
| 
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
 | 
414  | 
proof  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
415  | 
fix A assume A_M: "A \<in> null_sets M"  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
416  | 
with `absolutely_continuous M N` have A_N: "A \<in> null_sets N"  | 
| 47694 | 417  | 
unfolding absolutely_continuous_def by auto  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
418  | 
moreover from A_M A_N have "(\<integral>\<^sup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
419  | 
ultimately have "N A - (\<integral>\<^sup>+ x. ?F A x \<partial>M) = 0"  | 
| 47694 | 420  | 
using positive_integral_positive[of M] by (auto intro!: antisym)  | 
421  | 
then show "A \<in> null_sets ?M"  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
422  | 
using A_M by (simp add: emeasure_M null_sets_def sets_eq)  | 
| 38656 | 423  | 
qed  | 
| 47694 | 424  | 
have upper_bound: "\<forall>A\<in>sets M. ?M A \<le> 0"  | 
| 38656 | 425  | 
proof (rule ccontr)  | 
426  | 
assume "\<not> ?thesis"  | 
|
| 47694 | 427  | 
then obtain A where A: "A \<in> sets M" and pos: "0 < ?M A"  | 
| 38656 | 428  | 
by (auto simp: not_le)  | 
429  | 
note pos  | 
|
| 47694 | 430  | 
also have "?M A \<le> ?M (space M)"  | 
431  | 
using emeasure_space[of ?M A] by (simp add: sets_eq[THEN sets_eq_imp_space_eq])  | 
|
432  | 
finally have pos_t: "0 < ?M (space M)" by simp  | 
|
| 38656 | 433  | 
moreover  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
434  | 
from pos_t have "emeasure M (space M) \<noteq> 0"  | 
| 47694 | 435  | 
using ac unfolding absolutely_continuous_def by (auto simp: null_sets_def)  | 
436  | 
then have pos_M: "0 < emeasure M (space M)"  | 
|
437  | 
using emeasure_nonneg[of M "space M"] by (simp add: le_less)  | 
|
| 38656 | 438  | 
moreover  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
439  | 
have "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"  | 
| 38656 | 440  | 
using `f \<in> G` unfolding G_def by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
441  | 
hence "(\<integral>\<^sup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"  | 
| 47694 | 442  | 
using M'.finite_emeasure_space by auto  | 
| 38656 | 443  | 
moreover  | 
| 47694 | 444  | 
def b \<equiv> "?M (space M) / emeasure M (space M) / 2"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
445  | 
ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"  | 
| 47694 | 446  | 
by (auto simp: ereal_divide_eq)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
447  | 
then have b: "b \<noteq> 0" "0 \<le> b" "0 < b" "b \<noteq> \<infinity>" by auto  | 
| 47694 | 448  | 
let ?Mb = "density M (\<lambda>_. b)"  | 
449  | 
have Mb: "finite_measure ?Mb" "sets ?Mb = sets ?M"  | 
|
450  | 
using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI)  | 
|
451  | 
from M'.Radon_Nikodym_aux[OF this] guess A0 ..  | 
|
452  | 
then have "A0 \<in> sets M"  | 
|
453  | 
and space_less_A0: "measure ?M (space M) - real b * measure M (space M) \<le> measure ?M A0 - real b * measure M A0"  | 
|
454  | 
and *: "\<And>B. B \<in> sets M \<Longrightarrow> B \<subseteq> A0 \<Longrightarrow> 0 \<le> measure ?M B - real b * measure M B"  | 
|
455  | 
using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
456  | 
    { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
 | 
| 47694 | 457  | 
with *[OF this] have "b * emeasure M B \<le> ?M B"  | 
458  | 
using b unfolding M'.emeasure_eq_measure emeasure_eq_measure by (cases b) auto }  | 
|
| 38656 | 459  | 
note bM_le_t = this  | 
| 46731 | 460  | 
let ?f0 = "\<lambda>x. f x + b * indicator A0 x"  | 
| 38656 | 461  | 
    { fix A assume A: "A \<in> sets M"
 | 
462  | 
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
463  | 
have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
464  | 
(\<integral>\<^sup>+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
 | 
465  | 
by (auto intro!: positive_integral_cong split: split_indicator)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
466  | 
hence "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
467  | 
(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
468  | 
using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`  | 
| 50003 | 469  | 
by (simp add: positive_integral_add positive_integral_cmult_indicator G_def) }  | 
| 38656 | 470  | 
note f0_eq = this  | 
471  | 
    { fix A assume A: "A \<in> sets M"
 | 
|
472  | 
hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
473  | 
have f_le_v: "(\<integral>\<^sup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto  | 
| 38656 | 474  | 
note f0_eq[OF A]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
475  | 
also have "(\<integral>\<^sup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^sup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"  | 
| 38656 | 476  | 
using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`  | 
477  | 
by (auto intro!: add_left_mono)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
478  | 
also have "\<dots> \<le> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + ?M A"  | 
| 47694 | 479  | 
using emeasure_mono[of "A \<inter> A0" A ?M] `A \<in> sets M` `A0 \<in> sets M`  | 
480  | 
by (auto intro!: add_left_mono simp: sets_eq)  | 
|
481  | 
also have "\<dots> \<le> N A"  | 
|
482  | 
unfolding emeasure_M[OF `A \<in> sets M`]  | 
|
483  | 
using f_le_v N.emeasure_eq_measure[of A] positive_integral_positive[of M "?F A"]  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
484  | 
by (cases "\<integral>\<^sup>+x. ?F A x \<partial>M", cases "N A") auto  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
485  | 
finally have "(\<integral>\<^sup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }  | 
| 50003 | 486  | 
hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G`  | 
487  | 
by (auto intro!: ereal_add_nonneg_nonneg simp: G_def)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
488  | 
have int_f_finite: "integral\<^sup>P M f \<noteq> \<infinity>"  | 
| 47694 | 489  | 
by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le)  | 
490  | 
have "0 < ?M (space M) - emeasure ?Mb (space M)"  | 
|
491  | 
using pos_t  | 
|
492  | 
by (simp add: b emeasure_density_const)  | 
|
493  | 
(simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def)  | 
|
494  | 
also have "\<dots> \<le> ?M A0 - b * emeasure M A0"  | 
|
495  | 
using space_less_A0 `A0 \<in> sets M` b  | 
|
496  | 
by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure)  | 
|
497  | 
finally have 1: "b * emeasure M A0 < ?M A0"  | 
|
498  | 
by (metis M'.emeasure_real `A0 \<in> sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1)  | 
|
499  | 
less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def)  | 
|
500  | 
with b have "0 < ?M A0"  | 
|
501  | 
by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times  | 
|
502  | 
ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def)  | 
|
503  | 
then have "emeasure M A0 \<noteq> 0" using ac `A0 \<in> sets M`  | 
|
504  | 
by (auto simp: absolutely_continuous_def null_sets_def)  | 
|
505  | 
then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto  | 
|
506  | 
hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
507  | 
with int_f_finite have "?y + 0 < integral\<^sup>P M f + b * emeasure M A0" unfolding int_f_eq_y  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
508  | 
using `f \<in> G`  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
509  | 
by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
510  | 
also have "\<dots> = integral\<^sup>P M ?f0" using f0_eq[OF sets.top] `A0 \<in> sets M` sets.sets_into_space  | 
| 38656 | 511  | 
by (simp cong: positive_integral_cong)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
512  | 
finally have "?y < integral\<^sup>P M ?f0" by simp  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
513  | 
moreover from `?f0 \<in> G` have "integral\<^sup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper)  | 
| 38656 | 514  | 
ultimately show False by auto  | 
515  | 
qed  | 
|
| 47694 | 516  | 
let ?f = "\<lambda>x. max 0 (f x)"  | 
| 38656 | 517  | 
show ?thesis  | 
| 47694 | 518  | 
proof (intro bexI[of _ ?f] measure_eqI conjI)  | 
519  | 
show "sets (density M ?f) = sets N"  | 
|
520  | 
by (simp add: sets_eq)  | 
|
521  | 
fix A assume A: "A\<in>sets (density M ?f)"  | 
|
522  | 
then show "emeasure (density M ?f) A = emeasure N A"  | 
|
523  | 
using `f \<in> G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
524  | 
by (cases "integral\<^sup>P M (?F A)")  | 
| 47694 | 525  | 
(auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric])  | 
526  | 
qed auto  | 
|
| 38656 | 527  | 
qed  | 
528  | 
||
| 40859 | 529  | 
lemma (in finite_measure) split_space_into_finite_sets_and_rest:  | 
| 47694 | 530  | 
assumes ac: "absolutely_continuous M N" and sets_eq: "sets N = sets M"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
531  | 
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>  | 
| 47694 | 532  | 
(\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>)) \<and>  | 
533  | 
(\<forall>i. N (B i) \<noteq> \<infinity>)"  | 
|
| 38656 | 534  | 
proof -  | 
| 47694 | 535  | 
  let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}"
 | 
536  | 
let ?a = "SUP Q:?Q. emeasure M Q"  | 
|
537  | 
  have "{} \<in> ?Q" by auto
 | 
|
| 38656 | 538  | 
  then have Q_not_empty: "?Q \<noteq> {}" by blast
 | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
539  | 
have "?a \<le> emeasure M (space M)" using sets.sets_into_space  | 
| 47694 | 540  | 
by (auto intro!: SUP_least emeasure_mono)  | 
541  | 
then have "?a \<noteq> \<infinity>" using finite_emeasure_space  | 
|
| 38656 | 542  | 
by auto  | 
| 47694 | 543  | 
from SUPR_countable_SUPR[OF Q_not_empty, of "emeasure M"]  | 
544  | 
obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"  | 
|
| 38656 | 545  | 
by auto  | 
| 47694 | 546  | 
then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto  | 
547  | 
from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q"  | 
|
| 38656 | 548  | 
by auto  | 
| 47694 | 549  | 
then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp  | 
| 46731 | 550  | 
let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i"  | 
| 47694 | 551  | 
have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)"  | 
552  | 
proof (rule SUP_emeasure_incseq[of ?O])  | 
|
553  | 
show "range ?O \<subseteq> sets M" using Q' by auto  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
554  | 
show "incseq ?O" by (fastforce intro!: incseq_SucI)  | 
| 38656 | 555  | 
qed  | 
556  | 
have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto  | 
|
| 47694 | 557  | 
have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto  | 
| 38656 | 558  | 
then have O_in_G: "\<And>i. ?O i \<in> ?Q"  | 
559  | 
proof (safe del: notI)  | 
|
| 47694 | 560  | 
    fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
 | 
561  | 
then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"  | 
|
562  | 
by (simp add: sets_eq emeasure_subadditive_finite)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
563  | 
also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)  | 
| 47694 | 564  | 
finally show "N (?O i) \<noteq> \<infinity>" by simp  | 
| 38656 | 565  | 
qed auto  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
566  | 
have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce  | 
| 47694 | 567  | 
have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric]  | 
| 38656 | 568  | 
proof (rule antisym)  | 
| 47694 | 569  | 
show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim  | 
570  | 
using Q' by (auto intro!: SUP_mono emeasure_mono)  | 
|
571  | 
show "(SUP i. emeasure M (?O i)) \<le> ?a" unfolding SUP_def  | 
|
| 38656 | 572  | 
proof (safe intro!: Sup_mono, unfold bex_simps)  | 
573  | 
fix i  | 
|
| 
52141
 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 
haftmann 
parents: 
51329 
diff
changeset
 | 
574  | 
      have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto
 | 
| 47694 | 575  | 
then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>  | 
| 
52141
 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 
haftmann 
parents: 
51329 
diff
changeset
 | 
576  | 
        emeasure M (\<Union>(Q' ` {..i})) \<le> emeasure M x"
 | 
| 38656 | 577  | 
using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])  | 
578  | 
qed  | 
|
579  | 
qed  | 
|
| 46731 | 580  | 
let ?O_0 = "(\<Union>i. ?O i)"  | 
| 38656 | 581  | 
have "?O_0 \<in> sets M" using Q' by auto  | 
| 40859 | 582  | 
def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"  | 
| 38656 | 583  | 
  { fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) }
 | 
584  | 
note Q_sets = this  | 
|
| 40859 | 585  | 
show ?thesis  | 
586  | 
proof (intro bexI exI conjI ballI impI allI)  | 
|
587  | 
show "disjoint_family Q"  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
588  | 
by (fastforce simp: disjoint_family_on_def Q_def  | 
| 40859 | 589  | 
split: nat.split_asm)  | 
590  | 
show "range Q \<subseteq> sets M"  | 
|
591  | 
using Q_sets by auto  | 
|
592  | 
    { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
 | 
|
| 47694 | 593  | 
show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"  | 
| 40859 | 594  | 
proof (rule disjCI, simp)  | 
| 47694 | 595  | 
assume *: "0 < emeasure M A \<longrightarrow> N A \<noteq> \<infinity>"  | 
596  | 
show "emeasure M A = 0 \<and> N A = 0"  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
597  | 
proof (cases "emeasure M A = 0")  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
598  | 
case True  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
599  | 
with ac A have "N A = 0"  | 
| 40859 | 600  | 
unfolding absolutely_continuous_def by auto  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
601  | 
with True show ?thesis by simp  | 
| 40859 | 602  | 
next  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
603  | 
case False  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
604  | 
with * have "N A \<noteq> \<infinity>" using emeasure_nonneg[of M A] by auto  | 
| 47694 | 605  | 
with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
606  | 
using Q' by (auto intro!: plus_emeasure sets.countable_UN)  | 
| 47694 | 607  | 
also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"  | 
608  | 
proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])  | 
|
| 40859 | 609  | 
show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"  | 
| 47694 | 610  | 
using `N A \<noteq> \<infinity>` O_sets A by auto  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
611  | 
qed (fastforce intro!: incseq_SucI)  | 
| 40859 | 612  | 
also have "\<dots> \<le> ?a"  | 
| 
44928
 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 
hoelzl 
parents: 
44918 
diff
changeset
 | 
613  | 
proof (safe intro!: SUP_least)  | 
| 40859 | 614  | 
fix i have "?O i \<union> A \<in> ?Q"  | 
615  | 
proof (safe del: notI)  | 
|
616  | 
show "?O i \<union> A \<in> sets M" using O_sets A by auto  | 
|
| 47694 | 617  | 
from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"  | 
618  | 
using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq)  | 
|
619  | 
with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"  | 
|
620  | 
using `N A \<noteq> \<infinity>` by auto  | 
|
| 40859 | 621  | 
qed  | 
| 47694 | 622  | 
then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)  | 
| 40859 | 623  | 
qed  | 
| 47694 | 624  | 
finally have "emeasure M A = 0"  | 
625  | 
unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure)  | 
|
626  | 
with `emeasure M A \<noteq> 0` show ?thesis by auto  | 
|
| 40859 | 627  | 
qed  | 
628  | 
qed }  | 
|
| 47694 | 629  | 
    { fix i show "N (Q i) \<noteq> \<infinity>"
 | 
| 40859 | 630  | 
proof (cases i)  | 
631  | 
case 0 then show ?thesis  | 
|
632  | 
unfolding Q_def using Q'[of 0] by simp  | 
|
633  | 
next  | 
|
634  | 
case (Suc n)  | 
|
| 47694 | 635  | 
with `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`  | 
636  | 
emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\<Union> x\<le>n. Q' x)"]  | 
|
637  | 
show ?thesis  | 
|
638  | 
by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def)  | 
|
| 40859 | 639  | 
qed }  | 
640  | 
show "space M - ?O_0 \<in> sets M" using Q'_sets by auto  | 
|
641  | 
    { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
 | 
|
642  | 
proof (induct j)  | 
|
643  | 
case 0 then show ?case by (simp add: Q_def)  | 
|
644  | 
next  | 
|
645  | 
case (Suc j)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44568 
diff
changeset
 | 
646  | 
have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastforce  | 
| 40859 | 647  | 
        have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
 | 
648  | 
then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"  | 
|
649  | 
by (simp add: UN_Un[symmetric] Q_def del: UN_Un)  | 
|
650  | 
then show ?case using Suc by (auto simp add: eq atMost_Suc)  | 
|
651  | 
qed }  | 
|
652  | 
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
 | 
653  | 
then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastforce  | 
| 40859 | 654  | 
qed  | 
655  | 
qed  | 
|
656  | 
||
657  | 
lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:  | 
|
| 47694 | 658  | 
assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M"  | 
659  | 
shows "\<exists>f\<in>borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"  | 
|
| 40859 | 660  | 
proof -  | 
661  | 
from split_space_into_finite_sets_and_rest[OF assms]  | 
|
662  | 
obtain Q0 and Q :: "nat \<Rightarrow> 'a set"  | 
|
663  | 
where Q: "disjoint_family Q" "range Q \<subseteq> sets M"  | 
|
664  | 
and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"  | 
|
| 47694 | 665  | 
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"  | 
666  | 
and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force  | 
|
| 40859 | 667  | 
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto  | 
| 47694 | 668  | 
let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))"  | 
669  | 
have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). (\<forall>x. 0 \<le> f x) \<and> density (?M i) f = ?N i"  | 
|
670  | 
proof (intro allI finite_measure.Radon_Nikodym_finite_measure)  | 
|
| 38656 | 671  | 
fix i  | 
| 47694 | 672  | 
from Q show "finite_measure (?M i)"  | 
673  | 
by (auto intro!: finite_measureI cong: positive_integral_cong  | 
|
674  | 
simp add: emeasure_density subset_eq sets_eq)  | 
|
675  | 
from Q have "emeasure (?N i) (space N) = emeasure N (Q i)"  | 
|
676  | 
by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: positive_integral_cong)  | 
|
677  | 
with Q_fin show "finite_measure (?N i)"  | 
|
678  | 
by (auto intro!: finite_measureI)  | 
|
679  | 
show "sets (?N i) = sets (?M i)" by (simp add: sets_eq)  | 
|
| 50003 | 680  | 
have [measurable]: "\<And>A. A \<in> sets M \<Longrightarrow> A \<in> sets N" by (simp add: sets_eq)  | 
| 47694 | 681  | 
show "absolutely_continuous (?M i) (?N i)"  | 
682  | 
using `absolutely_continuous M N` `Q i \<in> sets M`  | 
|
683  | 
by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq  | 
|
684  | 
intro!: absolutely_continuous_AE[OF sets_eq])  | 
|
| 38656 | 685  | 
qed  | 
| 47694 | 686  | 
from choice[OF this[unfolded Bex_def]]  | 
687  | 
obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"  | 
|
688  | 
and f_density: "\<And>i. density (?M i) (f i) = ?N i"  | 
|
| 
54776
 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 
immler 
parents: 
53374 
diff
changeset
 | 
689  | 
by force  | 
| 47694 | 690  | 
  { fix A i assume A: "A \<in> sets M"
 | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
691  | 
with Q borel have "(\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"  | 
| 47694 | 692  | 
by (auto simp add: emeasure_density positive_integral_density subset_eq  | 
693  | 
intro!: positive_integral_cong split: split_indicator)  | 
|
694  | 
also have "\<dots> = emeasure N (Q i \<inter> A)"  | 
|
695  | 
using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
696  | 
finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }  | 
| 47694 | 697  | 
note integral_eq = this  | 
| 46731 | 698  | 
let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"  | 
| 38656 | 699  | 
show ?thesis  | 
700  | 
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
 | 
701  | 
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
 | 
702  | 
by (auto intro!: measurable_If)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
703  | 
show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def)  | 
| 47694 | 704  | 
show "density M ?f = N"  | 
705  | 
proof (rule measure_eqI)  | 
|
706  | 
fix A assume "A \<in> sets (density M ?f)"  | 
|
707  | 
then have "A \<in> sets M" by simp  | 
|
708  | 
have Qi: "\<And>i. Q i \<in> sets M" using Q by auto  | 
|
709  | 
have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"  | 
|
710  | 
"\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"  | 
|
711  | 
using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
712  | 
have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"  | 
| 47694 | 713  | 
using borel by (intro positive_integral_cong) (auto simp: indicator_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
714  | 
also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"  | 
| 47694 | 715  | 
using borel Qi Q0(1) `A \<in> sets M`  | 
716  | 
by (subst positive_integral_add) (auto simp del: ereal_infty_mult  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
717  | 
simp add: positive_integral_cmult_indicator sets.Int intro!: suminf_0_le)  | 
| 47694 | 718  | 
also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"  | 
719  | 
by (subst integral_eq[OF `A \<in> sets M`], subst positive_integral_suminf) auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
720  | 
finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .  | 
| 47694 | 721  | 
moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"  | 
722  | 
using Q Q_sets `A \<in> sets M`  | 
|
723  | 
by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)  | 
|
724  | 
moreover have "\<infinity> * emeasure M (Q0 \<inter> A) = N (Q0 \<inter> A)"  | 
|
725  | 
proof -  | 
|
726  | 
have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast  | 
|
727  | 
from in_Q0[OF this] show ?thesis by auto  | 
|
728  | 
qed  | 
|
729  | 
moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"  | 
|
730  | 
using Q_sets `A \<in> sets M` Q0(1) by auto  | 
|
731  | 
      moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
732  | 
using `A \<in> sets M` sets.sets_into_space Q0 by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
733  | 
ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)"  | 
| 47694 | 734  | 
using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "Q0 \<inter> A"] by (simp add: sets_eq)  | 
735  | 
with `A \<in> sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A"  | 
|
| 50003 | 736  | 
by (auto simp: subset_eq emeasure_density)  | 
| 47694 | 737  | 
qed (simp add: sets_eq)  | 
| 38656 | 738  | 
qed  | 
739  | 
qed  | 
|
740  | 
||
741  | 
lemma (in sigma_finite_measure) Radon_Nikodym:  | 
|
| 47694 | 742  | 
assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M"  | 
743  | 
shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"  | 
|
| 38656 | 744  | 
proof -  | 
745  | 
from Ex_finite_integrable_function  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
746  | 
obtain h where finite: "integral\<^sup>P M h \<noteq> \<infinity>" and  | 
| 38656 | 747  | 
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
 | 
748  | 
nn: "\<And>x. 0 \<le> h x" and  | 
| 38656 | 749  | 
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
 | 
750  | 
"\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
751  | 
let ?T = "\<lambda>A. (\<integral>\<^sup>+x. h x * indicator A x \<partial>M)"  | 
| 47694 | 752  | 
let ?MT = "density M h"  | 
753  | 
from borel finite nn interpret T: finite_measure ?MT  | 
|
754  | 
by (auto intro!: finite_measureI cong: positive_integral_cong simp: emeasure_density)  | 
|
755  | 
have "absolutely_continuous ?MT N" "sets N = sets ?MT"  | 
|
756  | 
proof (unfold absolutely_continuous_def, safe)  | 
|
757  | 
fix A assume "A \<in> null_sets ?MT"  | 
|
758  | 
with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0"  | 
|
759  | 
by (auto simp add: null_sets_density_iff)  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
760  | 
with pos sets.sets_into_space have "AE x in M. x \<notin> A"  | 
| 47694 | 761  | 
by (elim eventually_elim1) (auto simp: not_le[symmetric])  | 
762  | 
then have "A \<in> null_sets M"  | 
|
763  | 
using `A \<in> sets M` by (simp add: AE_iff_null_sets)  | 
|
764  | 
with ac show "A \<in> null_sets N"  | 
|
765  | 
by (auto simp: absolutely_continuous_def)  | 
|
766  | 
qed (auto simp add: sets_eq)  | 
|
767  | 
from T.Radon_Nikodym_finite_measure_infinite[OF this]  | 
|
768  | 
obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "density ?MT f = N" by auto  | 
|
769  | 
with nn borel show ?thesis  | 
|
770  | 
by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)  | 
|
| 38656 | 771  | 
qed  | 
772  | 
||
| 40859 | 773  | 
section "Uniqueness of densities"  | 
774  | 
||
| 47694 | 775  | 
lemma finite_density_unique:  | 
| 40859 | 776  | 
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 47694 | 777  | 
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
778  | 
and fin: "integral\<^sup>P M f \<noteq> \<infinity>"  | 
| 49785 | 779  | 
shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"  | 
| 40859 | 780  | 
proof (intro iffI ballI)  | 
| 47694 | 781  | 
fix A assume eq: "AE x in M. f x = g x"  | 
| 49785 | 782  | 
with borel show "density M f = density M g"  | 
783  | 
by (auto intro: density_cong)  | 
|
| 40859 | 784  | 
next  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
785  | 
let ?P = "\<lambda>f A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M"  | 
| 49785 | 786  | 
assume "density M f = density M g"  | 
787  | 
with borel have eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"  | 
|
788  | 
by (simp add: emeasure_density[symmetric])  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
789  | 
from this[THEN bspec, OF sets.top] fin  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
790  | 
have g_fin: "integral\<^sup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)  | 
| 40859 | 791  | 
  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
 | 
| 47694 | 792  | 
and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
793  | 
and g_fin: "integral\<^sup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"  | 
| 40859 | 794  | 
    let ?N = "{x\<in>space M. g x < f x}"
 | 
795  | 
have N: "?N \<in> sets M" using borel by simp  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
796  | 
have "?P g ?N \<le> integral\<^sup>P M g" using pos  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
797  | 
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
 | 
798  | 
then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
799  | 
have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"  | 
| 40859 | 800  | 
by (auto intro!: positive_integral_cong simp: indicator_def)  | 
801  | 
also have "\<dots> = ?P f ?N - ?P g ?N"  | 
|
802  | 
proof (rule positive_integral_diff)  | 
|
803  | 
show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"  | 
|
804  | 
using borel N by auto  | 
|
| 47694 | 805  | 
show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"  | 
806  | 
"AE x in M. 0 \<le> g x * indicator ?N x"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
807  | 
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
 | 
808  | 
qed fact  | 
| 40859 | 809  | 
also have "\<dots> = 0"  | 
| 47694 | 810  | 
unfolding eq[THEN bspec, OF N] using positive_integral_positive[of M] Pg_fin by auto  | 
811  | 
finally have "AE x in M. f x \<le> g x"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
812  | 
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
 | 
813  | 
by (subst (asm) positive_integral_0_iff_AE)  | 
| 43920 | 814  | 
(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
 | 
815  | 
from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq  | 
| 47694 | 816  | 
show "AE x in M. f x = g x" by auto  | 
| 40859 | 817  | 
qed  | 
818  | 
||
819  | 
lemma (in finite_measure) density_unique_finite_measure:  | 
|
820  | 
assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"  | 
|
| 47694 | 821  | 
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
822  | 
assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. f' x * indicator A x \<partial>M)"  | 
| 40859 | 823  | 
(is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")  | 
| 47694 | 824  | 
shows "AE x in M. f x = f' x"  | 
| 40859 | 825  | 
proof -  | 
| 47694 | 826  | 
let ?D = "\<lambda>f. density M f"  | 
827  | 
let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A"  | 
|
| 46731 | 828  | 
let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x"  | 
| 47694 | 829  | 
|
830  | 
have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M"  | 
|
831  | 
using borel by (auto intro!: absolutely_continuousI_density)  | 
|
832  | 
from split_space_into_finite_sets_and_rest[OF this]  | 
|
| 40859 | 833  | 
obtain Q0 and Q :: "nat \<Rightarrow> 'a set"  | 
834  | 
where Q: "disjoint_family Q" "range Q \<subseteq> sets M"  | 
|
835  | 
and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"  | 
|
| 47694 | 836  | 
and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>"  | 
837  | 
and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force  | 
|
838  | 
with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>"  | 
|
839  | 
and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq)  | 
|
840  | 
||
| 40859 | 841  | 
from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto  | 
| 47694 | 842  | 
  let ?D = "{x\<in>space M. f x \<noteq> f' x}"
 | 
843  | 
have "?D \<in> sets M" using borel by auto  | 
|
| 43920 | 844  | 
have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"  | 
| 40859 | 845  | 
unfolding indicator_def by auto  | 
| 47694 | 846  | 
have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos  | 
| 40859 | 847  | 
by (intro finite_density_unique[THEN iffD1] allI)  | 
| 50003 | 848  | 
(auto intro!: f measure_eqI simp: emeasure_density * subset_eq)  | 
| 47694 | 849  | 
moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"  | 
| 40859 | 850  | 
proof (rule AE_I')  | 
| 43920 | 851  | 
    { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
 | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
852  | 
and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"  | 
| 46731 | 853  | 
      let ?A = "\<lambda>i. Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
 | 
| 47694 | 854  | 
have "(\<Union>i. ?A i) \<in> null_sets M"  | 
| 40859 | 855  | 
proof (rule null_sets_UN)  | 
| 43923 | 856  | 
fix i ::nat have "?A i \<in> sets M"  | 
| 40859 | 857  | 
using borel Q0(1) by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
858  | 
have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"  | 
| 40859 | 859  | 
unfolding eq[OF `?A i \<in> sets M`]  | 
860  | 
by (auto intro!: positive_integral_mono simp: indicator_def)  | 
|
| 47694 | 861  | 
also have "\<dots> = i * emeasure M (?A i)"  | 
| 40859 | 862  | 
using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)  | 
| 47694 | 863  | 
also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by simp  | 
864  | 
finally have "?N (?A i) \<noteq> \<infinity>" by simp  | 
|
865  | 
then show "?A i \<in> null_sets M" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto  | 
|
| 40859 | 866  | 
qed  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
867  | 
      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
 | 
868  | 
by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat)  | 
| 47694 | 869  | 
      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
 | 
| 40859 | 870  | 
from this[OF borel(1) refl] this[OF borel(2) f]  | 
| 47694 | 871  | 
    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
 | 
872  | 
    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 M" by (rule null_sets.Un)
 | 
|
| 40859 | 873  | 
    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
 | 
874  | 
      (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 | 875  | 
qed  | 
| 47694 | 876  | 
moreover have "AE x in M. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>  | 
| 40859 | 877  | 
?f (space M) x = ?f' (space M) x"  | 
878  | 
by (auto simp: indicator_def Q0)  | 
|
| 47694 | 879  | 
ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x"  | 
880  | 
unfolding AE_all_countable[symmetric]  | 
|
881  | 
by eventually_elim (auto intro!: AE_I2 split: split_if_asm simp: indicator_def)  | 
|
882  | 
then show "AE x in M. f x = f' x" by auto  | 
|
| 40859 | 883  | 
qed  | 
884  | 
||
885  | 
lemma (in sigma_finite_measure) density_unique:  | 
|
| 47694 | 886  | 
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"  | 
887  | 
assumes f': "f' \<in> borel_measurable M" "AE x in M. 0 \<le> f' x"  | 
|
888  | 
assumes density_eq: "density M f = density M f'"  | 
|
889  | 
shows "AE x in M. f x = f' x"  | 
|
| 40859 | 890  | 
proof -  | 
891  | 
obtain h where h_borel: "h \<in> borel_measurable M"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
892  | 
and fin: "integral\<^sup>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 | 893  | 
using Ex_finite_integrable_function by auto  | 
| 47694 | 894  | 
then have h_nn: "AE x in M. 0 \<le> h x" by auto  | 
895  | 
let ?H = "density M h"  | 
|
896  | 
interpret h: finite_measure ?H  | 
|
897  | 
using fin h_borel pos  | 
|
898  | 
by (intro finite_measureI) (simp cong: positive_integral_cong emeasure_density add: fin)  | 
|
899  | 
let ?fM = "density M f"  | 
|
900  | 
let ?f'M = "density M f'"  | 
|
| 40859 | 901  | 
  { 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
 | 
902  | 
    then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
 | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
903  | 
using pos(1) sets.sets_into_space by (force simp: indicator_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
904  | 
then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
905  | 
using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto }  | 
| 40859 | 906  | 
note h_null_sets = this  | 
907  | 
  { fix A assume "A \<in> sets M"
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
908  | 
have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
909  | 
using `A \<in> sets M` h_borel h_nn f f'  | 
| 47694 | 910  | 
by (intro positive_integral_density[symmetric]) auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
911  | 
also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)"  | 
| 47694 | 912  | 
by (simp_all add: density_eq)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
913  | 
also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h 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
 | 
914  | 
using `A \<in> sets M` h_borel h_nn f f'  | 
| 47694 | 915  | 
by (intro positive_integral_density) auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
916  | 
finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
917  | 
by (simp add: ac_simps)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
918  | 
then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
919  | 
using `A \<in> sets M` h_borel h_nn f f'  | 
| 47694 | 920  | 
by (subst (asm) (1 2) positive_integral_density[symmetric]) auto }  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
921  | 
then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'  | 
| 47694 | 922  | 
by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M])  | 
923  | 
(auto simp add: AE_density)  | 
|
924  | 
then show "AE x in M. f x = f' x"  | 
|
925  | 
unfolding eventually_ae_filter using h_borel pos  | 
|
926  | 
by (auto simp add: h_null_sets null_sets_density_iff not_less[symmetric]  | 
|
| 
50021
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
927  | 
AE_iff_null_sets[symmetric]) blast  | 
| 40859 | 928  | 
qed  | 
929  | 
||
| 47694 | 930  | 
lemma (in sigma_finite_measure) density_unique_iff:  | 
931  | 
assumes f: "f \<in> borel_measurable M" and "AE x in M. 0 \<le> f x"  | 
|
932  | 
assumes f': "f' \<in> borel_measurable M" and "AE x in M. 0 \<le> f' x"  | 
|
933  | 
shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"  | 
|
934  | 
using density_unique[OF assms] density_cong[OF f f'] by auto  | 
|
935  | 
||
| 49785 | 936  | 
lemma sigma_finite_density_unique:  | 
937  | 
assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
|
938  | 
assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"  | 
|
939  | 
and fin: "sigma_finite_measure (density M f)"  | 
|
940  | 
shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"  | 
|
941  | 
proof  | 
|
942  | 
assume "AE x in M. f x = g x" with borel show "density M f = density M g"  | 
|
943  | 
by (auto intro: density_cong)  | 
|
944  | 
next  | 
|
945  | 
assume eq: "density M f = density M g"  | 
|
946  | 
interpret f!: sigma_finite_measure "density M f" by fact  | 
|
947  | 
from f.sigma_finite_incseq guess A . note cover = this  | 
|
948  | 
||
949  | 
have "AE x in M. \<forall>i. x \<in> A i \<longrightarrow> f x = g x"  | 
|
950  | 
unfolding AE_all_countable  | 
|
951  | 
proof  | 
|
952  | 
fix i  | 
|
953  | 
have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))"  | 
|
954  | 
unfolding eq ..  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
955  | 
moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"  | 
| 49785 | 956  | 
using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq)  | 
957  | 
ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x"  | 
|
958  | 
using borel pos cover(1) pos  | 
|
959  | 
by (intro finite_density_unique[THEN iffD1])  | 
|
960  | 
(auto simp: density_density_eq subset_eq)  | 
|
961  | 
then show "AE x in M. x \<in> A i \<longrightarrow> f x = g x"  | 
|
962  | 
by auto  | 
|
963  | 
qed  | 
|
964  | 
with AE_space show "AE x in M. f x = g x"  | 
|
965  | 
apply eventually_elim  | 
|
966  | 
using cover(2)[symmetric]  | 
|
967  | 
apply auto  | 
|
968  | 
done  | 
|
969  | 
qed  | 
|
970  | 
||
| 
49778
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
971  | 
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite':  | 
| 47694 | 972  | 
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"  | 
973  | 
shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"  | 
|
974  | 
(is "sigma_finite_measure ?N \<longleftrightarrow> _")  | 
|
| 40859 | 975  | 
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
 | 
976  | 
assume "sigma_finite_measure ?N"  | 
| 47694 | 977  | 
then interpret N: sigma_finite_measure ?N .  | 
978  | 
from N.Ex_finite_integrable_function obtain h where  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
979  | 
h: "h \<in> borel_measurable M" "integral\<^sup>P ?N h \<noteq> \<infinity>" and  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
980  | 
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
 | 
981  | 
fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto  | 
| 47694 | 982  | 
have "AE x in M. f x * h x \<noteq> \<infinity>"  | 
| 40859 | 983  | 
proof (rule AE_I')  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
984  | 
have "integral\<^sup>P ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)" using f h h_nn  | 
| 47694 | 985  | 
by (auto intro!: positive_integral_density)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
986  | 
then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"  | 
| 40859 | 987  | 
using h(2) by simp  | 
| 47694 | 988  | 
    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
 | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
989  | 
using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage)  | 
| 40859 | 990  | 
qed auto  | 
| 47694 | 991  | 
then show "AE x in M. f x \<noteq> \<infinity>"  | 
| 41705 | 992  | 
using fin by (auto elim!: AE_Ball_mp)  | 
| 40859 | 993  | 
next  | 
| 47694 | 994  | 
assume AE: "AE x in M. f x \<noteq> \<infinity>"  | 
| 40859 | 995  | 
from sigma_finite guess Q .. note Q = this  | 
| 43923 | 996  | 
  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(of_nat (Suc n))}) \<inter> space M"
 | 
| 40859 | 997  | 
  { fix i j have "A i \<inter> Q j \<in> sets M"
 | 
998  | 
unfolding A_def using f Q  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
999  | 
apply (rule_tac sets.Int)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1000  | 
by (cases i) (auto intro: measurable_sets[OF f(1)]) }  | 
| 40859 | 1001  | 
note A_in_sets = this  | 
| 46731 | 1002  | 
let ?A = "\<lambda>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
 | 
1003  | 
show "sigma_finite_measure ?N"  | 
| 40859 | 1004  | 
proof (default, intro exI conjI subsetI allI)  | 
1005  | 
fix x assume "x \<in> range ?A"  | 
|
1006  | 
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
 | 
1007  | 
then show "x \<in> sets ?N" using A_in_sets by (cases "prod_decode n") auto  | 
| 40859 | 1008  | 
next  | 
1009  | 
have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)"  | 
|
1010  | 
proof safe  | 
|
1011  | 
fix x i j assume "x \<in> A i" "x \<in> Q j"  | 
|
1012  | 
then show "x \<in> (\<Union>i. case prod_decode i of (i, j) \<Rightarrow> A i \<inter> Q j)"  | 
|
1013  | 
by (intro UN_I[of "prod_encode (i,j)"]) auto  | 
|
1014  | 
qed auto  | 
|
1015  | 
also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto  | 
|
1016  | 
also have "(\<Union>i. A i) = space M"  | 
|
1017  | 
proof safe  | 
|
1018  | 
fix x assume x: "x \<in> space M"  | 
|
1019  | 
show "x \<in> (\<Union>i. A i)"  | 
|
1020  | 
proof (cases "f x")  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1021  | 
case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])  | 
| 40859 | 1022  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1023  | 
case (real r)  | 
| 43923 | 1024  | 
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
 | 
1025  | 
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
 | 
1026  | 
next  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1027  | 
case MInf with x show ?thesis  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1028  | 
unfolding A_def by (auto intro!: exI[of _ "Suc 0"])  | 
| 40859 | 1029  | 
qed  | 
1030  | 
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
 | 
1031  | 
finally show "(\<Union>i. ?A i) = space ?N" by simp  | 
| 40859 | 1032  | 
next  | 
1033  | 
fix n obtain i j where  | 
|
1034  | 
[simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1035  | 
have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"  | 
| 40859 | 1036  | 
proof (cases i)  | 
1037  | 
case 0  | 
|
| 47694 | 1038  | 
have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"  | 
| 41705 | 1039  | 
using AE by (auto simp: A_def `i = 0`)  | 
1040  | 
from positive_integral_cong_AE[OF this] show ?thesis by simp  | 
|
| 40859 | 1041  | 
next  | 
1042  | 
case (Suc n)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1043  | 
then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1044  | 
(\<integral>\<^sup>+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
 | 
1045  | 
by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat)  | 
| 47694 | 1046  | 
also have "\<dots> = Suc n * emeasure M (Q j)"  | 
| 40859 | 1047  | 
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
 | 
1048  | 
also have "\<dots> < \<infinity>"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1049  | 
using Q by (auto simp: real_eq_of_nat[symmetric])  | 
| 40859 | 1050  | 
finally show ?thesis by simp  | 
1051  | 
qed  | 
|
| 47694 | 1052  | 
then show "emeasure ?N (?A n) \<noteq> \<infinity>"  | 
1053  | 
using A_in_sets Q f by (auto simp: emeasure_density)  | 
|
| 40859 | 1054  | 
qed  | 
1055  | 
qed  | 
|
1056  | 
||
| 
49778
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1057  | 
lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1058  | 
"f \<in> borel_measurable M \<Longrightarrow> sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1059  | 
apply (subst density_max_0)  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1060  | 
apply (subst sigma_finite_iff_density_finite')  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1061  | 
apply (auto simp: max_def intro!: measurable_If)  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1062  | 
done  | 
| 
 
bbbc0f492780
sigma_finite_iff_density_finite does not require a positive density function
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1063  | 
|
| 40871 | 1064  | 
section "Radon-Nikodym derivative"  | 
| 38656 | 1065  | 
|
| 
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
 | 
1066  | 
definition  | 
| 47694 | 1067  | 
"RN_deriv M N \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N"  | 
| 38656 | 1068  | 
|
| 47694 | 1069  | 
lemma  | 
1070  | 
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"  | 
|
1071  | 
shows borel_measurable_RN_deriv_density: "RN_deriv M (density M f) \<in> borel_measurable M" (is ?borel)  | 
|
1072  | 
and density_RN_deriv_density: "density M (RN_deriv M (density M f)) = density M f" (is ?density)  | 
|
1073  | 
and RN_deriv_density_nonneg: "0 \<le> RN_deriv M (density M f) x" (is ?pos)  | 
|
| 40859 | 1074  | 
proof -  | 
| 47694 | 1075  | 
let ?f = "\<lambda>x. max 0 (f x)"  | 
1076  | 
let ?P = "\<lambda>g. g \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> g x) \<and> density M g = density M f"  | 
|
1077  | 
from f have "?P ?f" using f by (auto intro!: density_cong simp: split: split_max)  | 
|
1078  | 
then have "?P (RN_deriv M (density M f))"  | 
|
1079  | 
unfolding RN_deriv_def by (rule someI[where P="?P"])  | 
|
1080  | 
then show ?borel ?density ?pos by auto  | 
|
| 40859 | 1081  | 
qed  | 
1082  | 
||
| 38656 | 1083  | 
lemma (in sigma_finite_measure) RN_deriv:  | 
| 47694 | 1084  | 
assumes "absolutely_continuous M N" "sets N = sets M"  | 
| 50003 | 1085  | 
shows borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M" (is ?borel)  | 
| 47694 | 1086  | 
and density_RN_deriv: "density M (RN_deriv M N) = N" (is ?density)  | 
1087  | 
and RN_deriv_nonneg: "0 \<le> RN_deriv M N x" (is ?pos)  | 
|
| 38656 | 1088  | 
proof -  | 
1089  | 
note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]  | 
|
| 47694 | 1090  | 
from Ex show ?borel unfolding RN_deriv_def by (rule someI2_ex) simp  | 
1091  | 
from Ex show ?density unfolding RN_deriv_def by (rule someI2_ex) simp  | 
|
1092  | 
from Ex show ?pos unfolding RN_deriv_def by (rule someI2_ex) simp  | 
|
| 38656 | 1093  | 
qed  | 
1094  | 
||
| 40859 | 1095  | 
lemma (in sigma_finite_measure) RN_deriv_positive_integral:  | 
| 47694 | 1096  | 
assumes N: "absolutely_continuous M N" "sets N = sets M"  | 
| 40859 | 1097  | 
and f: "f \<in> borel_measurable M"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1098  | 
shows "integral\<^sup>P N f = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"  | 
| 40859 | 1099  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1100  | 
have "integral\<^sup>P N f = integral\<^sup>P (density M (RN_deriv M N)) f"  | 
| 47694 | 1101  | 
using N by (simp add: density_RN_deriv)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1102  | 
also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"  | 
| 47694 | 1103  | 
using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density)  | 
1104  | 
finally show ?thesis by simp  | 
|
| 40859 | 1105  | 
qed  | 
1106  | 
||
| 47694 | 1107  | 
lemma null_setsD_AE: "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"  | 
1108  | 
using AE_iff_null_sets[of N M] by auto  | 
|
1109  | 
||
1110  | 
lemma (in sigma_finite_measure) RN_deriv_unique:  | 
|
1111  | 
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"  | 
|
1112  | 
and eq: "density M f = N"  | 
|
1113  | 
shows "AE x in M. f x = RN_deriv M N x"  | 
|
| 49785 | 1114  | 
unfolding eq[symmetric]  | 
1115  | 
by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv_density  | 
|
1116  | 
RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])  | 
|
1117  | 
||
1118  | 
lemma RN_deriv_unique_sigma_finite:  | 
|
1119  | 
assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"  | 
|
1120  | 
and eq: "density M f = N" and fin: "sigma_finite_measure N"  | 
|
1121  | 
shows "AE x in M. f x = RN_deriv M N x"  | 
|
1122  | 
using fin unfolding eq[symmetric]  | 
|
1123  | 
by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv_density  | 
|
1124  | 
RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])  | 
|
| 47694 | 1125  | 
|
1126  | 
lemma (in sigma_finite_measure) RN_deriv_distr:  | 
|
1127  | 
fixes T :: "'a \<Rightarrow> 'b"  | 
|
1128  | 
assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"  | 
|
1129  | 
and inv: "\<forall>x\<in>space M. T' (T x) = x"  | 
|
| 
50021
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1130  | 
and ac[simp]: "absolutely_continuous (distr M M' T) (distr N M' T)"  | 
| 47694 | 1131  | 
and N: "sets N = sets M"  | 
1132  | 
shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x"  | 
|
| 41832 | 1133  | 
proof (rule RN_deriv_unique)  | 
| 47694 | 1134  | 
have [simp]: "sets N = sets M" by fact  | 
1135  | 
note sets_eq_imp_space_eq[OF N, simp]  | 
|
1136  | 
have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def)  | 
|
1137  | 
  { fix A assume "A \<in> sets M"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
1138  | 
with inv T T' sets.sets_into_space[OF this]  | 
| 47694 | 1139  | 
have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A"  | 
1140  | 
by (auto simp: measurable_def) }  | 
|
1141  | 
note eq = this[simp]  | 
|
1142  | 
  { fix A assume "A \<in> sets M"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
1143  | 
with inv T T' sets.sets_into_space[OF this]  | 
| 47694 | 1144  | 
have "(T' \<circ> T) -` A \<inter> space M = A"  | 
1145  | 
by (auto simp: measurable_def) }  | 
|
1146  | 
note eq2 = this[simp]  | 
|
1147  | 
let ?M' = "distr M M' T" and ?N' = "distr N M' T"  | 
|
1148  | 
interpret M': sigma_finite_measure ?M'  | 
|
| 41832 | 1149  | 
proof  | 
1150  | 
from sigma_finite guess F .. note F = this  | 
|
| 47694 | 1151  | 
show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets ?M' \<and> (\<Union>i. A i) = space ?M' \<and> (\<forall>i. emeasure ?M' (A i) \<noteq> \<infinity>)"  | 
| 41832 | 1152  | 
proof (intro exI conjI allI)  | 
| 47694 | 1153  | 
show *: "range (\<lambda>i. T' -` F i \<inter> space ?M') \<subseteq> sets ?M'"  | 
1154  | 
using F T' by (auto simp: measurable_def)  | 
|
1155  | 
show "(\<Union>i. T' -` F i \<inter> space ?M') = space ?M'"  | 
|
1156  | 
using F T' by (force simp: measurable_def)  | 
|
| 41832 | 1157  | 
fix i  | 
1158  | 
have "T' -` F i \<inter> space M' \<in> sets M'" using * by auto  | 
|
1159  | 
moreover  | 
|
1160  | 
have Fi: "F i \<in> sets M" using F by auto  | 
|
| 47694 | 1161  | 
ultimately show "emeasure ?M' (T' -` F i \<inter> space ?M') \<noteq> \<infinity>"  | 
1162  | 
using F T T' by (simp add: emeasure_distr)  | 
|
| 41832 | 1163  | 
qed  | 
1164  | 
qed  | 
|
| 47694 | 1165  | 
have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"  | 
| 
50021
 
d96a3f468203
add support for function application to measurability prover
 
hoelzl 
parents: 
50003 
diff
changeset
 | 
1166  | 
using T ac by measurable  | 
| 47694 | 1167  | 
then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"  | 
| 41832 | 1168  | 
by (simp add: comp_def)  | 
| 47694 | 1169  | 
show "AE x in M. 0 \<le> RN_deriv ?M' ?N' (T x)" using M'.RN_deriv_nonneg[OF ac] by auto  | 
1170  | 
||
1171  | 
have "N = distr N M (T' \<circ> T)"  | 
|
1172  | 
by (subst measure_of_of_measure[of N, symmetric])  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50021 
diff
changeset
 | 
1173  | 
(auto simp add: distr_def sets.sigma_sets_eq intro!: measure_of_eq sets.space_closed)  | 
| 47694 | 1174  | 
also have "\<dots> = distr (distr N M' T) M T'"  | 
1175  | 
using T T' by (simp add: distr_distr)  | 
|
1176  | 
also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'"  | 
|
1177  | 
using ac by (simp add: M'.density_RN_deriv)  | 
|
1178  | 
also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)"  | 
|
1179  | 
using M'.borel_measurable_RN_deriv[OF ac] by (simp add: distr_density_distr[OF T T', OF inv])  | 
|
1180  | 
finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N"  | 
|
1181  | 
by (simp add: comp_def)  | 
|
| 41832 | 1182  | 
qed  | 
1183  | 
||
| 40859 | 1184  | 
lemma (in sigma_finite_measure) RN_deriv_finite:  | 
| 47694 | 1185  | 
assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"  | 
1186  | 
shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>"  | 
|
| 40859 | 1187  | 
proof -  | 
| 47694 | 1188  | 
interpret N: sigma_finite_measure N by fact  | 
1189  | 
from N show ?thesis  | 
|
1190  | 
using sigma_finite_iff_density_finite[OF RN_deriv(1)[OF ac]] RN_deriv(2,3)[OF ac] by simp  | 
|
| 40859 | 1191  | 
qed  | 
1192  | 
||
1193  | 
lemma (in sigma_finite_measure)  | 
|
| 47694 | 1194  | 
assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"  | 
| 40859 | 1195  | 
and f: "f \<in> borel_measurable M"  | 
| 47694 | 1196  | 
shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>  | 
1197  | 
integrable M (\<lambda>x. real (RN_deriv M N x) * f x)" (is ?integrable)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1198  | 
and RN_deriv_integral: "integral\<^sup>L N f =  | 
| 47694 | 1199  | 
(\<integral>x. real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)  | 
| 40859 | 1200  | 
proof -  | 
| 47694 | 1201  | 
note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]  | 
1202  | 
interpret N: sigma_finite_measure N by fact  | 
|
| 43920 | 1203  | 
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 | 1204  | 
have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto  | 
| 47694 | 1205  | 
have Nf: "f \<in> borel_measurable N" using f by (simp add: measurable_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
 | 
1206  | 
  { fix f :: "'a \<Rightarrow> real"
 | 
| 47694 | 1207  | 
    { fix x assume *: "RN_deriv M N x \<noteq> \<infinity>"
 | 
1208  | 
have "ereal (real (RN_deriv M N x)) * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"  | 
|
| 40859 | 1209  | 
by (simp add: mult_le_0_iff)  | 
| 47694 | 1210  | 
then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"  | 
1211  | 
using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) }  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1212  | 
then have "(\<integral>\<^sup>+x. ereal (real (RN_deriv M N x) * f x) \<partial>M) = (\<integral>\<^sup>+x. RN_deriv M N x * ereal (f x) \<partial>M)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1213  | 
"(\<integral>\<^sup>+x. ereal (- (real (RN_deriv M N x) * f x)) \<partial>M) = (\<integral>\<^sup>+x. RN_deriv M N x * ereal (- f x) \<partial>M)"  | 
| 47694 | 1214  | 
using RN_deriv_finite[OF N ac] 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
 | 
1215  | 
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
 | 
1216  | 
note * = this  | 
| 40859 | 1217  | 
show ?integral ?integrable  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41832 
diff
changeset
 | 
1218  | 
unfolding lebesgue_integral_def integrable_def *  | 
| 47694 | 1219  | 
using Nf f RN_deriv(1)[OF ac]  | 
1220  | 
by (auto simp: RN_deriv_positive_integral[OF ac])  | 
|
| 40859 | 1221  | 
qed  | 
1222  | 
||
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1223  | 
lemma (in sigma_finite_measure) real_RN_deriv:  | 
| 47694 | 1224  | 
assumes "finite_measure N"  | 
1225  | 
assumes ac: "absolutely_continuous M N" "sets N = sets M"  | 
|
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1226  | 
obtains D where "D \<in> borel_measurable M"  | 
| 47694 | 1227  | 
and "AE x in M. RN_deriv M N x = ereal (D x)"  | 
1228  | 
and "AE x in N. 0 < D x"  | 
|
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1229  | 
and "\<And>x. 0 \<le> D x"  | 
| 
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1230  | 
proof  | 
| 47694 | 1231  | 
interpret N: finite_measure N by fact  | 
1232  | 
||
1233  | 
note RN = RN_deriv[OF ac]  | 
|
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1234  | 
|
| 47694 | 1235  | 
  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
 | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1236  | 
|
| 47694 | 1237  | 
show "(\<lambda>x. real (RN_deriv M N x)) \<in> borel_measurable M"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1238  | 
using RN by auto  | 
| 
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1239  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1240  | 
have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"  | 
| 47694 | 1241  | 
using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1242  | 
also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1243  | 
by (intro positive_integral_cong) (auto simp: indicator_def)  | 
| 47694 | 1244  | 
also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1245  | 
using RN by (intro positive_integral_cmult_indicator) auto  | 
| 47694 | 1246  | 
finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" .  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1247  | 
moreover  | 
| 47694 | 1248  | 
have "emeasure M (?RN \<infinity>) = 0"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1249  | 
proof (rule ccontr)  | 
| 47694 | 1250  | 
    assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
 | 
1251  | 
    moreover from RN have "0 \<le> emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
 | 
|
1252  | 
    ultimately have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
 | 
|
1253  | 
with eq have "N (?RN \<infinity>) = \<infinity>" by simp  | 
|
1254  | 
with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto  | 
|
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1255  | 
qed  | 
| 47694 | 1256  | 
ultimately have "AE x in M. RN_deriv M N x < \<infinity>"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1257  | 
using RN by (intro AE_iff_measurable[THEN iffD2]) auto  | 
| 47694 | 1258  | 
then show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))"  | 
| 43920 | 1259  | 
using RN(3) by (auto simp: ereal_real)  | 
| 47694 | 1260  | 
then have eq: "AE x in N. RN_deriv M N x = ereal (real (RN_deriv M N x))"  | 
1261  | 
using ac absolutely_continuous_AE by auto  | 
|
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1262  | 
|
| 47694 | 1263  | 
show "\<And>x. 0 \<le> real (RN_deriv M N x)"  | 
| 43920 | 1264  | 
using RN by (auto intro: real_of_ereal_pos)  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1265  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1266  | 
have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"  | 
| 47694 | 1267  | 
using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1268  | 
also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)"  | 
| 
43340
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1269  | 
by (intro positive_integral_cong) (auto simp: indicator_def)  | 
| 47694 | 1270  | 
finally have "AE x in N. RN_deriv M N x \<noteq> 0"  | 
1271  | 
using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)  | 
|
1272  | 
with RN(3) eq show "AE x in N. 0 < real (RN_deriv M N x)"  | 
|
| 43920 | 1273  | 
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
 | 
1274  | 
qed  | 
| 
 
60e181c4eae4
lemma: independence is equal to mutual information = 0
 
hoelzl 
parents: 
42866 
diff
changeset
 | 
1275  | 
|
| 38656 | 1276  | 
lemma (in sigma_finite_measure) RN_deriv_singleton:  | 
| 47694 | 1277  | 
assumes ac: "absolutely_continuous M N" "sets N = sets M"  | 
1278  | 
  and x: "{x} \<in> sets M"
 | 
|
1279  | 
  shows "N {x} = RN_deriv M N x * emeasure M {x}"
 | 
|
| 38656 | 1280  | 
proof -  | 
| 47694 | 1281  | 
note deriv = RN_deriv[OF ac]  | 
1282  | 
  from deriv(1,3) `{x} \<in> sets M`
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
52141 
diff
changeset
 | 
1283  | 
  have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
 | 
| 47694 | 1284  | 
by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong)  | 
1285  | 
with x deriv show ?thesis  | 
|
1286  | 
by (auto simp: positive_integral_cmult_indicator)  | 
|
| 38656 | 1287  | 
qed  | 
1288  | 
||
1289  | 
end  |