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