| author | blanchet | 
| Fri, 11 Feb 2011 11:54:24 +0100 | |
| changeset 41752 | 949eaf045e00 | 
| parent 41705 | 1100512e16d8 | 
| child 41832 | 27cb9113b1a0 | 
| permissions | -rw-r--r-- | 
| 38656 | 1 | theory Radon_Nikodym | 
| 2 | imports Lebesgue_Integration | |
| 3 | begin | |
| 4 | ||
| 40859 | 5 | lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)" | 
| 6 | proof safe | |
| 7 | assume "x < \<omega>" | |
| 8 | then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto | |
| 9 | moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto | |
| 10 | ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat) | |
| 11 | qed auto | |
| 12 | ||
| 38656 | 13 | lemma (in sigma_finite_measure) Ex_finite_integrable_function: | 
| 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 | 14 | shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)" | 
| 38656 | 15 | proof - | 
| 16 | obtain A :: "nat \<Rightarrow> 'a set" where | |
| 17 | range: "range A \<subseteq> sets M" and | |
| 18 | space: "(\<Union>i. A i) = space M" and | |
| 19 | measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and | |
| 20 | disjoint: "disjoint_family A" | |
| 21 | using disjoint_sigma_finite by auto | |
| 22 | let "?B i" = "2^Suc i * \<mu> (A i)" | |
| 23 | have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)" | |
| 24 | proof | |
| 25 | fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)" | |
| 26 | proof cases | |
| 27 | assume "\<mu> (A i) = 0" | |
| 28 | then show ?thesis by (auto intro!: exI[of _ 1]) | |
| 29 | next | |
| 30 | assume not_0: "\<mu> (A i) \<noteq> 0" | |
| 31 | then have "?B i \<noteq> \<omega>" using measure[of i] by auto | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 32 | then have "inverse (?B i) \<noteq> 0" unfolding pextreal_inverse_eq_0 by simp | 
| 38656 | 33 | then show ?thesis using measure[of i] not_0 | 
| 34 | by (auto intro!: exI[of _ "inverse (?B i) / 2"] | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 35 | simp: pextreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq) | 
| 38656 | 36 | qed | 
| 37 | qed | |
| 38 | from choice[OF this] obtain n where n: "\<And>i. 0 < n i" | |
| 39 | "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto | |
| 40 | let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x" | |
| 41 | show ?thesis | |
| 42 | proof (safe intro!: bexI[of _ ?h] del: notI) | |
| 39092 | 43 | have "\<And>i. A i \<in> sets M" | 
| 44 | using range by fastsimp+ | |
| 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 | 45 | then have "integral\<^isup>P M ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))" | 
| 39092 | 46 | by (simp add: positive_integral_psuminf positive_integral_cmult_indicator) | 
| 38656 | 47 | also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))" | 
| 48 | proof (rule psuminf_le) | |
| 49 | fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)" | |
| 50 | using measure[of N] n[of N] | |
| 39092 | 51 | by (cases "n N") | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 52 | (auto simp: pextreal_noteq_omega_Ex field_simps zero_le_mult_iff | 
| 39092 | 53 | mult_le_0_iff mult_less_0_iff power_less_zero_eq | 
| 54 | power_le_zero_eq inverse_eq_divide less_divide_eq | |
| 55 | power_divide split: split_if_asm) | |
| 38656 | 56 | qed | 
| 57 | also have "\<dots> = Real 1" | |
| 58 | by (rule suminf_imp_psuminf, rule power_half_series, 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 | 59 | finally show "integral\<^isup>P M ?h \<noteq> \<omega>" by auto | 
| 38656 | 60 | next | 
| 61 | fix x assume "x \<in> space M" | |
| 62 | then obtain i where "x \<in> A i" using space[symmetric] by auto | |
| 63 | from psuminf_cmult_indicator[OF disjoint, OF this] | |
| 64 | have "?h x = n i" by simp | |
| 65 | then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto | |
| 66 | next | |
| 67 | show "?h \<in> borel_measurable M" using range | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 68 | by (auto intro!: borel_measurable_psuminf borel_measurable_pextreal_times) | 
| 38656 | 69 | qed | 
| 70 | qed | |
| 71 | ||
| 40871 | 72 | subsection "Absolutely continuous" | 
| 73 | ||
| 38656 | 74 | definition (in measure_space) | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 75 | "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pextreal))" | 
| 38656 | 76 | |
| 40859 | 77 | lemma (in sigma_finite_measure) 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" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 80 | shows "measure_space.almost_everywhere M' P" | 
| 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 | |
| 85 | show "\<nu>.almost_everywhere P" | |
| 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})"
 | 
| 39097 | 102 | proof (rule v.measure_finitely_additive''[symmetric]) | 
| 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" | |
| 110 |     hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
 | |
| 111 |     hence "\<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 | 112 |     thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
 | 
| 39097 | 113 | 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 | 114 | finally show "\<nu> N = 0" by simp | 
| 39097 | 115 | qed | 
| 116 | ||
| 40871 | 117 | 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 | 118 | assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" | 
| 40871 | 119 | shows "absolutely_continuous \<nu>" | 
| 120 | using assms unfolding absolutely_continuous_def | |
| 121 | by (simp add: positive_integral_null_set) | |
| 122 | ||
| 123 | subsection "Existence of the Radon-Nikodym derivative" | |
| 124 | ||
| 38656 | 125 | lemma (in finite_measure) Radon_Nikodym_aux_epsilon: | 
| 126 | 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 | 127 | assumes "finite_measure (M\<lparr>measure := \<nu>\<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 | 128 | shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<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 | 129 | real (\<mu> A) - real (\<nu> A) \<and> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 130 | (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < real (\<mu> B) - real (\<nu> B))" | 
| 38656 | 131 | 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 | 132 | let "?d A" = "real (\<mu> A) - real (\<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 | 133 | interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact | 
| 38656 | 134 | let "?A A" = "if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B) | 
| 135 |     then {}
 | |
| 136 | else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)" | |
| 137 |   def A \<equiv> "\<lambda>n. ((\<lambda>B. B \<union> ?A B) ^^ n) {}"
 | |
| 138 | have A_simps[simp]: | |
| 139 |     "A 0 = {}"
 | |
| 140 | "\<And>n. A (Suc n) = (A n \<union> ?A (A n))" unfolding A_def by simp_all | |
| 141 |   { fix A assume "A \<in> sets M"
 | |
| 142 | have "?A A \<in> sets M" | |
| 143 | by (auto intro!: someI2[of _ _ "\<lambda>A. A \<in> sets M"] simp: not_less) } | |
| 144 | note A'_in_sets = this | |
| 145 |   { fix n have "A n \<in> sets M"
 | |
| 146 | proof (induct n) | |
| 147 | case (Suc n) thus "A (Suc n) \<in> sets M" | |
| 148 | using A'_in_sets[of "A n"] by (auto split: split_if_asm) | |
| 149 | qed (simp add: A_def) } | |
| 150 | note A_in_sets = this | |
| 151 | hence "range A \<subseteq> sets M" by auto | |
| 152 |   { fix n B
 | |
| 153 | assume Ex: "\<exists>B. B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> -e" | |
| 154 | hence False: "\<not> (\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B)" by (auto simp: not_less) | |
| 155 | have "?d (A (Suc n)) \<le> ?d (A n) - e" unfolding A_simps if_not_P[OF False] | |
| 156 | proof (rule someI2_ex[OF Ex]) | |
| 157 | fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e" | |
| 158 |       hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
 | |
| 159 | hence "?d (A n \<union> B) = ?d (A n) + ?d B" | |
| 160 | using `A n \<in> sets M` real_finite_measure_Union M'.real_finite_measure_Union by simp | |
| 161 | also have "\<dots> \<le> ?d (A n) - e" using dB by simp | |
| 162 | finally show "?d (A n \<union> B) \<le> ?d (A n) - e" . | |
| 163 | qed } | |
| 164 | note dA_epsilon = this | |
| 165 |   { fix n have "?d (A (Suc n)) \<le> ?d (A n)"
 | |
| 166 | proof (cases "\<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e") | |
| 167 | case True from dA_epsilon[OF this] show ?thesis using `0 < e` by simp | |
| 168 | next | |
| 169 | case False | |
| 170 | hence "\<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B" by (auto simp: not_le) | |
| 171 | thus ?thesis by simp | |
| 172 | qed } | |
| 173 | note dA_mono = this | |
| 174 | show ?thesis | |
| 175 | proof (cases "\<exists>n. \<forall>B\<in>sets M. B \<subseteq> space M - A n \<longrightarrow> -e < ?d B") | |
| 176 | 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 | |
| 177 | show ?thesis | |
| 178 | proof (safe intro!: bexI[of _ "space M - A n"]) | |
| 179 | fix B assume "B \<in> sets M" "B \<subseteq> space M - A n" | |
| 180 | from B[OF this] show "-e < ?d B" . | |
| 181 | next | |
| 182 | show "space M - A n \<in> sets M" by (rule compl_sets) fact | |
| 183 | next | |
| 184 | show "?d (space M) \<le> ?d (space M - A n)" | |
| 185 | proof (induct n) | |
| 186 | fix n assume "?d (space M) \<le> ?d (space M - A n)" | |
| 187 | also have "\<dots> \<le> ?d (space M - A (Suc n))" | |
| 188 | using A_in_sets sets_into_space dA_mono[of n] | |
| 189 | real_finite_measure_Diff[of "space M"] | |
| 190 | real_finite_measure_Diff[of "space M"] | |
| 191 | M'.real_finite_measure_Diff[of "space M"] | |
| 192 | M'.real_finite_measure_Diff[of "space M"] | |
| 193 | by (simp del: A_simps) | |
| 194 | finally show "?d (space M) \<le> ?d (space M - A (Suc n))" . | |
| 195 | qed simp | |
| 196 | qed | |
| 197 | next | |
| 198 | case False hence B: "\<And>n. \<exists>B. B\<in>sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e" | |
| 199 | by (auto simp add: not_less) | |
| 200 |     { fix n have "?d (A n) \<le> - real n * e"
 | |
| 201 | proof (induct n) | |
| 202 | case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps) | |
| 203 | qed simp } note dA_less = this | |
| 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 | |
| 208 | from real_finite_continuity_from_below[of A] `range A \<subseteq> sets M` | |
| 209 | M'.real_finite_continuity_from_below[of A] | |
| 210 | have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)" | |
| 211 | by (auto intro!: LIMSEQ_diff) | |
| 212 | obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto | |
| 213 | moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less] | |
| 214 | have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps) | |
| 215 | ultimately show ?thesis by auto | |
| 216 | qed | |
| 217 | qed | |
| 218 | ||
| 219 | 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 | 220 | assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "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 | 221 | shows "\<exists>A\<in>sets M. real (\<mu> (space M)) - real (\<nu> (space M)) \<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 | 222 | real (\<mu> A) - real (\<nu> A) \<and> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 223 | (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> real (\<mu> B) - real (\<nu> B))" | 
| 38656 | 224 | 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 | 225 | let "?d A" = "real (\<mu> A) - real (\<nu> A)" | 
| 38656 | 226 | 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)" | 
| 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 | 227 | 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 | 228 | "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto | 
| 39092 | 229 | let "?r S" = "restricted_space S" | 
| 38656 | 230 |   { fix S n
 | 
| 231 | assume S: "S \<in> sets M" | |
| 232 | hence **: "\<And>X. X \<in> op \<inter> S ` sets M \<longleftrightarrow> X \<in> sets M \<and> X \<subseteq> S" 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 | 233 | have [simp]: "(restricted_space S\<lparr>measure := \<nu>\<rparr>) = M'.restricted_space S" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 234 | by (cases M) 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 | 235 | from M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 236 | have "finite_measure (?r S)" "0 < 1 / real (Suc 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 | 237 | "finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto | 
| 38656 | 238 | from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. | 
| 239 | hence "?P X S n" | |
| 240 | proof (simp add: **, safe) | |
| 241 | fix C assume C: "C \<in> sets M" "C \<subseteq> X" "X \<subseteq> S" and | |
| 242 | *: "\<forall>B\<in>sets M. S \<inter> B \<subseteq> X \<longrightarrow> - (1 / real (Suc n)) < ?d (S \<inter> B)" | |
| 243 | hence "C \<subseteq> S" "C \<subseteq> X" "S \<inter> C = C" by auto | |
| 244 | with *[THEN bspec, OF `C \<in> sets M`] | |
| 245 | show "- (1 / real (Suc n)) < ?d C" by auto | |
| 246 | qed | |
| 247 | hence "\<exists>A. ?P A S n" by auto } | |
| 248 | note Ex_P = this | |
| 249 | def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)" | |
| 250 | have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def) | |
| 251 | have A_0[simp]: "A 0 = space M" unfolding A_def by simp | |
| 252 |   { fix i have "A i \<in> sets M" unfolding A_def
 | |
| 253 | proof (induct i) | |
| 254 | case (Suc i) | |
| 255 | from Ex_P[OF this, of i] show ?case unfolding nat_rec_Suc | |
| 256 | by (rule someI2_ex) simp | |
| 257 | qed simp } | |
| 258 | note A_in_sets = this | |
| 259 |   { fix n have "?P (A (Suc n)) (A n) n"
 | |
| 260 | using Ex_P[OF A_in_sets] unfolding A_Suc | |
| 261 | by (rule someI2_ex) simp } | |
| 262 | note P_A = this | |
| 263 | have "range A \<subseteq> sets M" using A_in_sets by auto | |
| 264 | have A_mono: "\<And>i. A (Suc i) \<subseteq> A i" using P_A by simp | |
| 265 | have mono_dA: "mono (\<lambda>i. ?d (A i))" using P_A by (simp add: mono_iff_le_Suc) | |
| 266 | have epsilon: "\<And>C i. \<lbrakk> C \<in> sets M; C \<subseteq> A (Suc i) \<rbrakk> \<Longrightarrow> - 1 / real (Suc i) < ?d C" | |
| 267 | using P_A by auto | |
| 268 | show ?thesis | |
| 269 | proof (safe intro!: bexI[of _ "\<Inter>i. A i"]) | |
| 270 | show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto | |
| 271 | from `range A \<subseteq> sets M` A_mono | |
| 272 | real_finite_continuity_from_above[of A] | |
| 273 | M'.real_finite_continuity_from_above[of A] | |
| 274 | have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: LIMSEQ_diff) | |
| 275 | thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _] | |
| 276 | by (rule_tac LIMSEQ_le_const) (auto intro!: exI) | |
| 277 | next | |
| 278 | fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)" | |
| 279 | show "0 \<le> ?d B" | |
| 280 | proof (rule ccontr) | |
| 281 | assume "\<not> 0 \<le> ?d B" | |
| 282 | hence "0 < - ?d B" by auto | |
| 283 | from ex_inverse_of_nat_Suc_less[OF this] | |
| 284 | obtain n where *: "?d B < - 1 / real (Suc n)" | |
| 285 | by (auto simp: real_eq_of_nat inverse_eq_divide field_simps) | |
| 286 | have "B \<subseteq> A (Suc n)" using B by (auto simp del: nat_rec_Suc) | |
| 287 | from epsilon[OF B(1) this] * | |
| 288 | show False by auto | |
| 289 | qed | |
| 290 | qed | |
| 291 | qed | |
| 292 | ||
| 293 | 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 | 294 | assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'") | 
| 38656 | 295 | 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 | 296 | 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 | 297 | 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 | 298 | 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 | 299 | 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 | 300 | using assms(1) 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 | 301 |   def G \<equiv> "{g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A}"
 | 
| 38656 | 302 | have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto | 
| 303 |   hence "G \<noteq> {}" by auto
 | |
| 304 |   { fix f g assume f: "f \<in> G" and g: "g \<in> G"
 | |
| 305 | have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def | |
| 306 | proof safe | |
| 307 | show "?max \<in> borel_measurable M" using f g unfolding G_def by auto | |
| 308 |       let ?A = "{x \<in> space M. f x \<le> g x}"
 | |
| 309 | have "?A \<in> sets M" using f g unfolding G_def by auto | |
| 310 | fix A assume "A \<in> sets M" | |
| 311 | hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto | |
| 312 | have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A" | |
| 313 | using sets_into_space[OF `A \<in> sets M`] by auto | |
| 314 | have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x = | |
| 315 | g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x" | |
| 316 | 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 | 317 | 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 | 318 | (\<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 | 319 | (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)" | 
| 38656 | 320 | using f g sets unfolding G_def | 
| 321 | by (auto cong: positive_integral_cong intro!: positive_integral_add borel_measurable_indicator) | |
| 322 | also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)" | |
| 323 | using f g sets unfolding G_def by (auto intro!: add_mono) | |
| 324 | also have "\<dots> = \<nu> A" | |
| 325 | 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 | 326 | finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" . | 
| 38656 | 327 | qed } | 
| 328 | note max_in_G = this | |
| 329 |   { fix f g assume  "f \<up> g" and f: "\<And>i. f i \<in> G"
 | |
| 330 | have "g \<in> G" unfolding G_def | |
| 331 | proof safe | |
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41095diff
changeset | 332 | from `f \<up> g` have [simp]: "g = (\<lambda>x. SUP i. f i x)" | 
| 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41095diff
changeset | 333 | unfolding isoton_def fun_eq_iff SUPR_apply by simp | 
| 38656 | 334 | have f_borel: "\<And>i. f i \<in> borel_measurable M" using f unfolding G_def by simp | 
| 41097 
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
 hoelzl parents: 
41095diff
changeset | 335 | thus "g \<in> borel_measurable M" by auto | 
| 38656 | 336 | fix A assume "A \<in> sets M" | 
| 337 | hence "\<And>i. (\<lambda>x. f i x * indicator A x) \<in> borel_measurable M" | |
| 338 | using f_borel by (auto intro!: borel_measurable_indicator) | |
| 339 | from positive_integral_isoton[OF isoton_indicator[OF `f \<up> 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 | 340 | have SUP: "(\<integral>\<^isup>+x. g 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 | 341 | (SUP i. (\<integral>\<^isup>+x. f i x * indicator A x \<partial>M))" | 
| 38656 | 342 | unfolding isoton_def 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 | 343 | show "(\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A" unfolding SUP | 
| 38656 | 344 | using f `A \<in> sets M` unfolding G_def by (auto intro!: SUP_leI) | 
| 345 | qed } | |
| 346 | 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 | 347 | let ?y = "SUP g : G. integral\<^isup>P M g" | 
| 38656 | 348 | have "?y \<le> \<nu> (space M)" unfolding G_def | 
| 349 | proof (safe intro!: SUP_leI) | |
| 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 | 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 | 351 | from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)" | 
| 38656 | 352 | by (simp cong: positive_integral_cong) | 
| 353 | qed | |
| 354 | hence "?y \<noteq> \<omega>" using M'.finite_measure_of_space by auto | |
| 355 |   from SUPR_countable_SUPR[OF this `G \<noteq> {}`] guess ys .. note ys = 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 | 356 | hence "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" | 
| 38656 | 357 | 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 | 358 | 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 | 359 | 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 | 360 | thus "\<exists>g. g\<in>G \<and> integral\<^isup>P M g = ys n" by auto | 
| 38656 | 361 | 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 | 362 | 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 | 363 | hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto | 
| 38656 | 364 |   let "?g i x" = "Max ((\<lambda>n. gs n x) ` {..i})"
 | 
| 365 | def f \<equiv> "SUP i. ?g i" | |
| 366 |   have gs_not_empty: "\<And>i. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
 | |
| 367 |   { fix i have "?g i \<in> G"
 | |
| 368 | proof (induct i) | |
| 369 | case 0 thus ?case by simp fact | |
| 370 | next | |
| 371 | case (Suc i) | |
| 372 | with Suc gs_not_empty `gs (Suc i) \<in> G` show ?case | |
| 373 | by (auto simp add: atMost_Suc intro!: max_in_G) | |
| 374 | qed } | |
| 375 | note g_in_G = this | |
| 376 | have "\<And>x. \<forall>i. ?g i x \<le> ?g (Suc i) x" | |
| 377 | using gs_not_empty by (simp add: atMost_Suc) | |
| 378 | hence isoton_g: "?g \<up> f" by (simp add: isoton_def le_fun_def f_def) | |
| 379 | from SUP_in_G[OF this g_in_G] have "f \<in> G" . | |
| 380 | hence [simp, intro]: "f \<in> borel_measurable 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 | 381 | have "(\<lambda>i. integral\<^isup>P M (?g i)) \<up> integral\<^isup>P M f" | 
| 38656 | 382 | using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_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 | 383 | hence "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" | 
| 38656 | 384 | unfolding isoton_def by simp | 
| 385 | also have "\<dots> = ?y" | |
| 386 | 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 | 387 | show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y" | 
| 38656 | 388 | using g_in_G by (auto intro!: exI Sup_mono simp: SUPR_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 | 389 | show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq | 
| 38656 | 390 | by (auto intro!: SUP_mono positive_integral_mono Max_ge) | 
| 391 | 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 | 392 | finally have int_f_eq_y: "integral\<^isup>P M f = ?y" . | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 393 | let "?t A" = "\<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 | 394 | 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 | 395 | 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 | 396 | by (intro sigma_algebra_cong) 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 | 397 | have fmM: "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 | 398 | proof (default, simp_all add: countably_additive_def, safe) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 399 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family 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 | 400 | have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) 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 | 401 | = (\<integral>\<^isup>+x. (\<Sum>\<^isub>\<infinity>n. f x * indicator (A n) 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 | 402 | using `range A \<subseteq> sets 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 | 403 | by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 404 | also have "\<dots> = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) 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 | 405 | apply (rule positive_integral_cong) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 406 | apply (subst psuminf_cmult_right) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 407 | unfolding psuminf_indicator[OF `disjoint_family 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 | 408 | finally have "(\<Sum>\<^isub>\<infinity> n. (\<integral>\<^isup>+x. f x * indicator (A n) 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 | 409 | = (\<integral>\<^isup>+x. f x * indicator (\<Union>n. A n) 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 | 410 | moreover have "(\<Sum>\<^isub>\<infinity>n. \<nu> (A n)) = \<nu> (\<Union>n. A 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 | 411 | using M'.measure_countably_additive A by (simp add: comp_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 | 412 | moreover have "\<And>i. (\<integral>\<^isup>+x. f x * indicator (A i) x \<partial>M) \<le> \<nu> (A 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 | 413 | using A `f \<in> G` unfolding G_def 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 | 414 | moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<omega>" using M'.finite_measure A by (simp add: countable_UN) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 415 |     moreover {
 | 
| 
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 | have "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A 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 | 417 | using A `f \<in> G` unfolding G_def by (auto simp: countable_UN) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 418 | also have "\<nu> (\<Union>i. A i) < \<omega>" using v_fin by (simp add: pextreal_less_\<omega>) | 
| 
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 "(\<integral>\<^isup>+x. f x * indicator (\<Union>i. A i) x \<partial>M) \<noteq> \<omega>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 420 | by (simp add: pextreal_less_\<omega>) } | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 421 | ultimately | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 422 | show "(\<Sum>\<^isub>\<infinity> n. ?t (A n)) = ?t (\<Union>i. A 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 | 423 | apply (subst psuminf_minus) by simp_all | 
| 38656 | 424 | 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 | 425 | 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 | 426 | 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 | 427 | by (simp_all add: fmM) | 
| 38656 | 428 | have ac: "absolutely_continuous ?t" using `absolutely_continuous \<nu>` unfolding absolutely_continuous_def by auto | 
| 429 | have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0" | |
| 430 | proof (rule ccontr) | |
| 431 | assume "\<not> ?thesis" | |
| 432 | then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A" | |
| 433 | by (auto simp: not_le) | |
| 434 | note pos | |
| 435 | also have "?t A \<le> ?t (space M)" | |
| 436 | using M.measure_mono[of A "space M"] A sets_into_space by simp | |
| 437 | finally have pos_t: "0 < ?t (space M)" by simp | |
| 438 | moreover | |
| 439 | hence pos_M: "0 < \<mu> (space M)" | |
| 440 | using ac top unfolding absolutely_continuous_def by auto | |
| 441 | 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 | 442 | have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)" | 
| 38656 | 443 | 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 | 444 | hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<omega>" | 
| 38656 | 445 | using M'.finite_measure_of_space by auto | 
| 446 | moreover | |
| 447 | def b \<equiv> "?t (space M) / \<mu> (space M) / 2" | |
| 448 | ultimately have b: "b \<noteq> 0" "b \<noteq> \<omega>" | |
| 449 | using M'.finite_measure_of_space | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 450 | by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space) | 
| 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 | 451 | 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 | 452 | interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) 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 | 453 | have "finite_measure ?Mb" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 454 | by default | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 455 | (insert finite_measure_of_space b measure_countably_additive, | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 456 | auto simp: psuminf_cmult_right countably_additive_def) | 
| 38656 | 457 | from M.Radon_Nikodym_aux[OF this] | 
| 458 | obtain A0 where "A0 \<in> sets M" and | |
| 459 | space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and | |
| 460 | *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)" by auto | |
| 461 |     { fix B assume "B \<in> sets M" "B \<subseteq> A0"
 | |
| 462 | with *[OF this] have "b * \<mu> B \<le> ?t B" | |
| 463 | using M'.finite_measure b finite_measure | |
| 464 | by (cases "b * \<mu> B", cases "?t B") (auto simp: field_simps) } | |
| 465 | note bM_le_t = this | |
| 466 | let "?f0 x" = "f x + b * indicator A0 x" | |
| 467 |     { fix A assume A: "A \<in> sets M"
 | |
| 468 | 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 | 469 | 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 | 470 | (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)" | 
| 38656 | 471 | by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith) | 
| 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 | 472 | 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 | 473 | (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)" | 
| 38656 | 474 | using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A | 
| 475 | by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) } | |
| 476 | note f0_eq = this | |
| 477 |     { fix A assume A: "A \<in> sets M"
 | |
| 478 | 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 | 479 | have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A" | 
| 38656 | 480 | using `f \<in> G` A unfolding G_def by auto | 
| 481 | 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 | 482 | 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 | 483 | (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t (A \<inter> A0)" | 
| 38656 | 484 | using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M` | 
| 485 | 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 | 486 | also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t A" | 
| 38656 | 487 | using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`] | 
| 488 | by (auto intro!: add_left_mono) | |
| 489 | also have "\<dots> \<le> \<nu> A" | |
| 490 | using f_le_v M'.finite_measure[simplified, OF `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 | 491 | by (cases "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M)", cases "\<nu> A", 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 | 492 | finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . } | 
| 38656 | 493 | hence "?f0 \<in> G" using `A0 \<in> sets M` unfolding G_def | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 494 | by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times) | 
| 38656 | 495 | have real: "?t (space M) \<noteq> \<omega>" "?t A0 \<noteq> \<omega>" | 
| 496 | "b * \<mu> (space M) \<noteq> \<omega>" "b * \<mu> A0 \<noteq> \<omega>" | |
| 497 | using `A0 \<in> sets M` b | |
| 498 | finite_measure[of A0] M.finite_measure[of A0] | |
| 499 | finite_measure_of_space M.finite_measure_of_space | |
| 500 | 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 | 501 | have int_f_finite: "integral\<^isup>P M f \<noteq> \<omega>" | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 502 | using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff | 
| 38656 | 503 | by (auto cong: positive_integral_cong) | 
| 504 | have "?t (space M) > b * \<mu> (space M)" unfolding b_def | |
| 505 | apply (simp add: field_simps) | |
| 506 | apply (subst mult_assoc[symmetric]) | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 507 | apply (subst pextreal_mult_inverse) | 
| 38656 | 508 | using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 509 | using pextreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"] | 
| 38656 | 510 | by simp_all | 
| 511 | hence "0 < ?t (space M) - b * \<mu> (space M)" | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 512 | by (simp add: pextreal_zero_less_diff_iff) | 
| 38656 | 513 | also have "\<dots> \<le> ?t A0 - b * \<mu> A0" | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 514 | using space_less_A0 pos_M pos_t b real[unfolded pextreal_noteq_omega_Ex] by auto | 
| 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 515 | finally have "b * \<mu> A0 < ?t A0" unfolding pextreal_zero_less_diff_iff . | 
| 38656 | 516 | hence "0 < ?t A0" by auto | 
| 517 | hence "0 < \<mu> A0" using ac unfolding absolutely_continuous_def | |
| 518 | using `A0 \<in> sets M` by auto | |
| 519 | hence "0 < b * \<mu> A0" using b by auto | |
| 520 | from int_f_finite 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 | 521 | have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 522 | by (rule pextreal_less_add) | 
| 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 | 523 | also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space | 
| 38656 | 524 | 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 | 525 | finally have "?y < integral\<^isup>P M ?f0" 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 | 526 | moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: le_SUPI) | 
| 38656 | 527 | ultimately show False by auto | 
| 528 | qed | |
| 529 | show ?thesis | |
| 530 | proof (safe intro!: bexI[of _ f]) | |
| 531 | fix A assume "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 | 532 | show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" | 
| 38656 | 533 | 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 | 534 | show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A" | 
| 38656 | 535 | 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 | 536 | show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" | 
| 38656 | 537 | using upper_bound[THEN bspec, OF `A \<in> sets M`] | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 538 | by (simp add: pextreal_zero_le_diff) | 
| 38656 | 539 | qed | 
| 540 | qed simp | |
| 541 | qed | |
| 542 | ||
| 40859 | 543 | 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 | 544 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") | 
| 40859 | 545 | assumes ac: "absolutely_continuous \<nu>" | 
| 546 | shows "\<exists>\<Omega>0\<in>sets M. \<exists>\<Omega>::nat\<Rightarrow>'a set. disjoint_family \<Omega> \<and> range \<Omega> \<subseteq> sets M \<and> \<Omega>0 = space M - (\<Union>i. \<Omega> i) \<and> | |
| 547 | (\<forall>A\<in>sets M. A \<subseteq> \<Omega>0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<omega>)) \<and> | |
| 548 | (\<forall>i. \<nu> (\<Omega> i) \<noteq> \<omega>)" | |
| 38656 | 549 | 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 | 550 | 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 | 551 | 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 | 552 | by fact auto | 
| 38656 | 553 |   let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
 | 
| 554 | let ?a = "SUP Q:?Q. \<mu> Q" | |
| 555 |   have "{} \<in> ?Q" using v.empty_measure by auto
 | |
| 556 |   then have Q_not_empty: "?Q \<noteq> {}" by blast
 | |
| 557 | have "?a \<le> \<mu> (space M)" using sets_into_space | |
| 558 | by (auto intro!: SUP_leI measure_mono top) | |
| 559 | then have "?a \<noteq> \<omega>" using finite_measure_of_space | |
| 560 | by auto | |
| 561 | from SUPR_countable_SUPR[OF this Q_not_empty] | |
| 562 | obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" | |
| 563 | by auto | |
| 564 | then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto | |
| 565 | from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q" | |
| 566 | by auto | |
| 567 | then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp | |
| 568 | let "?O n" = "\<Union>i\<le>n. Q' i" | |
| 569 | have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)" | |
| 570 | proof (rule continuity_from_below[of ?O]) | |
| 571 | show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN) | |
| 572 | show "\<And>i. ?O i \<subseteq> ?O (Suc i)" by fastsimp | |
| 573 | qed | |
| 574 | have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto | |
| 575 | have O_sets: "\<And>i. ?O i \<in> sets M" | |
| 576 | using Q' by (auto intro!: finite_UN Un) | |
| 577 | then have O_in_G: "\<And>i. ?O i \<in> ?Q" | |
| 578 | proof (safe del: notI) | |
| 579 |     fix i have "Q' ` {..i} \<subseteq> sets M"
 | |
| 580 | using Q' by (auto intro: finite_UN) | |
| 581 |     with v.measure_finitely_subadditive[of "{.. i}" Q']
 | |
| 582 | have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 583 | also have "\<dots> < \<omega>" unfolding setsum_\<omega> pextreal_less_\<omega> using Q' by auto | 
| 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 584 | finally show "\<nu> (?O i) \<noteq> \<omega>" unfolding pextreal_less_\<omega> by auto | 
| 38656 | 585 | qed auto | 
| 586 | have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastsimp | |
| 587 | have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric] | |
| 588 | proof (rule antisym) | |
| 589 | show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim | |
| 590 | using Q' by (auto intro!: SUP_mono measure_mono finite_UN) | |
| 591 | show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUPR_def | |
| 592 | proof (safe intro!: Sup_mono, unfold bex_simps) | |
| 593 | fix i | |
| 594 |       have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
 | |
| 595 | then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<omega>) \<and> | |
| 596 |         \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
 | |
| 597 | using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) | |
| 598 | qed | |
| 599 | qed | |
| 600 | let "?O_0" = "(\<Union>i. ?O i)" | |
| 601 | have "?O_0 \<in> sets M" using Q' by auto | |
| 40859 | 602 | def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n" | 
| 38656 | 603 |   { fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) }
 | 
| 604 | note Q_sets = this | |
| 40859 | 605 | show ?thesis | 
| 606 | proof (intro bexI exI conjI ballI impI allI) | |
| 607 | show "disjoint_family Q" | |
| 608 | by (fastsimp simp: disjoint_family_on_def Q_def | |
| 609 | split: nat.split_asm) | |
| 610 | show "range Q \<subseteq> sets M" | |
| 611 | using Q_sets by auto | |
| 612 |     { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
 | |
| 613 | show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>" | |
| 614 | proof (rule disjCI, simp) | |
| 615 | assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<omega>" | |
| 616 | show "\<mu> A = 0 \<and> \<nu> A = 0" | |
| 617 | proof cases | |
| 618 | assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0" | |
| 619 | unfolding absolutely_continuous_def by auto | |
| 620 | ultimately show ?thesis by simp | |
| 621 | next | |
| 622 | assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<omega>" by auto | |
| 623 | with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)" | |
| 624 | using Q' by (auto intro!: measure_additive countable_UN) | |
| 625 | also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))" | |
| 626 | proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified]) | |
| 627 | show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M" | |
| 628 | using `\<nu> A \<noteq> \<omega>` O_sets A by auto | |
| 629 | qed fastsimp | |
| 630 | also have "\<dots> \<le> ?a" | |
| 631 | proof (safe intro!: SUPR_bound) | |
| 632 | fix i have "?O i \<union> A \<in> ?Q" | |
| 633 | proof (safe del: notI) | |
| 634 | show "?O i \<union> A \<in> sets M" using O_sets A by auto | |
| 635 | from O_in_G[of i] | |
| 636 | moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A" | |
| 637 | using v.measure_subadditive[of "?O i" A] A O_sets by auto | |
| 638 | ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>" | |
| 639 | using `\<nu> A \<noteq> \<omega>` by auto | |
| 640 | qed | |
| 641 | then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI) | |
| 642 | qed | |
| 643 | finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`] | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 644 | by (cases "\<mu> A") (auto simp: pextreal_noteq_omega_Ex) | 
| 40859 | 645 | with `\<mu> A \<noteq> 0` show ?thesis by auto | 
| 646 | qed | |
| 647 | qed } | |
| 648 |     { fix i show "\<nu> (Q i) \<noteq> \<omega>"
 | |
| 649 | proof (cases i) | |
| 650 | case 0 then show ?thesis | |
| 651 | unfolding Q_def using Q'[of 0] by simp | |
| 652 | next | |
| 653 | case (Suc n) | |
| 654 | then show ?thesis unfolding Q_def | |
| 655 | using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono | |
| 656 | using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto | |
| 657 | qed } | |
| 658 | show "space M - ?O_0 \<in> sets M" using Q'_sets by auto | |
| 659 |     { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
 | |
| 660 | proof (induct j) | |
| 661 | case 0 then show ?case by (simp add: Q_def) | |
| 662 | next | |
| 663 | case (Suc j) | |
| 664 | have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp | |
| 665 |         have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
 | |
| 666 | then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)" | |
| 667 | by (simp add: UN_Un[symmetric] Q_def del: UN_Un) | |
| 668 | then show ?case using Suc by (auto simp add: eq atMost_Suc) | |
| 669 | qed } | |
| 670 | then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp | |
| 671 | then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastsimp | |
| 672 | qed | |
| 673 | qed | |
| 674 | ||
| 675 | 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 | 676 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") | 
| 40859 | 677 | 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 | 678 | shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" | 
| 40859 | 679 | 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 | 680 | 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 | 681 | 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 | 682 | by fact auto | 
| 40859 | 683 | from split_space_into_finite_sets_and_rest[OF assms] | 
| 684 | obtain Q0 and Q :: "nat \<Rightarrow> 'a set" | |
| 685 | where Q: "disjoint_family Q" "range Q \<subseteq> sets M" | |
| 686 | and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)" | |
| 687 | 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 = \<omega>" | |
| 688 | and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<omega>" by force | |
| 689 | from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto | |
| 38656 | 690 | have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<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 | 691 | \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))" | 
| 38656 | 692 | proof | 
| 693 | fix i | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 694 | have indicator_eq: "\<And>f x A. (f x :: pextreal) * indicator (Q i \<inter> A) x * indicator (Q i) x | 
| 38656 | 695 | = (f x * indicator (Q i) x) * indicator A x" | 
| 696 | 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 | 697 | 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 | 698 | (is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]]) | 
| 38656 | 699 | 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 | 700 | have fmv: "finite_measure (restricted_space (Q i) \<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?Q") | 
| 38656 | 701 | unfolding finite_measure_def finite_measure_axioms_def | 
| 702 | 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 | 703 | show "measure_space ?Q" | 
| 38656 | 704 | using v.restricted_measure_space Q_sets[of i] 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 | 705 | show "measure ?Q (space ?Q) \<noteq> \<omega>" using Q_fin by simp | 
| 38656 | 706 | qed | 
| 707 | have "R.absolutely_continuous \<nu>" | |
| 708 | using `absolutely_continuous \<nu>` `Q i \<in> sets M` | |
| 709 | 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 | 710 | from R.Radon_Nikodym_finite_measure[OF fmv this] | 
| 38656 | 711 | 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 | 712 | 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 | 713 | unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`] | 
| 714 | positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq) | |
| 715 | then show "\<exists>f. f\<in>borel_measurable M \<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 | 716 | \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))" | 
| 38656 | 717 | by (fastsimp intro!: exI[of _ "\<lambda>x. f x * indicator (Q i) x"] positive_integral_cong | 
| 718 | simp: indicator_def) | |
| 719 | qed | |
| 720 | from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M" | |
| 721 | 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 | 722 | \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" | 
| 38656 | 723 | by auto | 
| 724 | let "?f x" = | |
| 40859 | 725 | "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator Q0 x" | 
| 38656 | 726 | show ?thesis | 
| 727 | proof (safe intro!: bexI[of _ ?f]) | |
| 728 | show "?f \<in> borel_measurable M" | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 729 | by (safe intro!: borel_measurable_psuminf borel_measurable_pextreal_times | 
| 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 730 | borel_measurable_pextreal_add borel_measurable_indicator | 
| 40859 | 731 | borel_measurable_const borel Q_sets Q0 Diff countable_UN) | 
| 38656 | 732 | fix A assume "A \<in> sets M" | 
| 40859 | 733 | have *: | 
| 38656 | 734 | "\<And>x i. indicator A x * (f i x * indicator (Q i) x) = | 
| 735 | f i x * indicator (Q i \<inter> A) x" | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 736 | "\<And>x i. (indicator A x * indicator Q0 x :: pextreal) = | 
| 40859 | 737 | indicator (Q0 \<inter> A) x" by (auto 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 | 738 | have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = | 
| 40859 | 739 | (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)" | 
| 38656 | 740 | unfolding f[OF `A \<in> sets M`] | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 741 | apply (simp del: pextreal_times(2) add: field_simps *) | 
| 38656 | 742 | apply (subst positive_integral_add) | 
| 40859 | 743 | apply (fastsimp intro: Q0 `A \<in> sets M`) | 
| 744 | apply (fastsimp intro: Q_sets `A \<in> sets M` borel_measurable_psuminf borel) | |
| 745 | apply (subst positive_integral_cmult_indicator) | |
| 746 | apply (fastsimp intro: Q0 `A \<in> sets M`) | |
| 38656 | 747 | unfolding psuminf_cmult_right[symmetric] | 
| 748 | apply (subst positive_integral_psuminf) | |
| 40859 | 749 | apply (fastsimp intro: `A \<in> sets M` Q_sets borel) | 
| 750 | apply (simp add: *) | |
| 751 | done | |
| 38656 | 752 | moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)" | 
| 40859 | 753 | using Q Q_sets `A \<in> sets M` | 
| 754 | by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified]) | |
| 755 | (auto simp: disjoint_family_on_def) | |
| 756 | moreover have "\<omega> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)" | |
| 757 | proof - | |
| 758 | have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast | |
| 759 | from in_Q0[OF this] show ?thesis by auto | |
| 38656 | 760 | qed | 
| 40859 | 761 | moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M" | 
| 762 | using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN) | |
| 763 |     moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
 | |
| 764 | 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 | 765 | ultimately show "\<nu> A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)" | 
| 40859 | 766 | using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"] | 
| 767 | by simp | |
| 38656 | 768 | qed | 
| 769 | qed | |
| 770 | ||
| 771 | 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 | 772 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N") | 
| 38656 | 773 | 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 | 774 | 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 | 775 | proof - | 
| 776 | from Ex_finite_integrable_function | |
| 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 | obtain h where finite: "integral\<^isup>P M h \<noteq> \<omega>" and | 
| 38656 | 778 | borel: "h \<in> borel_measurable M" and | 
| 779 | pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and | |
| 780 | "\<And>x. x \<in> space M \<Longrightarrow> h x < \<omega>" 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 | 781 | 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 | 782 | let ?MT = "M\<lparr> measure := ?T \<rparr>" | 
| 38656 | 783 | from measure_space_density[OF borel] 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 | 784 | 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 | 785 | where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T" | 
| 38656 | 786 | unfolding finite_measure_def finite_measure_axioms_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 | 787 | by (simp_all cong: positive_integral_cong) | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 788 |   have "\<And>N. N \<in> sets M \<Longrightarrow> {x \<in> space M. h x \<noteq> 0 \<and> indicator N x \<noteq> (0::pextreal)} = N"
 | 
| 38656 | 789 | using sets_into_space pos by (force simp: indicator_def) | 
| 790 | then have "T.absolutely_continuous \<nu>" using assms(2) borel | |
| 791 | unfolding T.absolutely_continuous_def absolutely_continuous_def | |
| 792 | by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff) | |
| 793 | from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this] | |
| 794 | obtain f where f_borel: "f \<in> borel_measurable M" 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 | 795 | 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 | 796 | by (auto simp: measurable_def) | 
| 38656 | 797 | show ?thesis | 
| 798 | proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"]) | |
| 799 | show "(\<lambda>x. h x * f x) \<in> borel_measurable M" | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 800 | using borel f_borel by (auto intro: borel_measurable_pextreal_times) | 
| 38656 | 801 | fix A assume "A \<in> sets M" | 
| 802 | then have "(\<lambda>x. f x * indicator A x) \<in> borel_measurable M" | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 803 | using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator) | 
| 38656 | 804 | from positive_integral_translated_density[OF borel 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 | 805 | show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)" | 
| 38656 | 806 | unfolding fT[OF `A \<in> sets M`] by (simp add: field_simps) | 
| 807 | qed | |
| 808 | qed | |
| 809 | ||
| 40859 | 810 | section "Uniqueness of densities" | 
| 811 | ||
| 812 | lemma (in measure_space) finite_density_unique: | |
| 813 | assumes borel: "f \<in> borel_measurable M" "g \<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 | 814 | and fin: "integral\<^isup>P M f < \<omega>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 815 | 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 | 816 | \<longleftrightarrow> (AE x. f x = g x)" | 
| 817 | (is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _") | |
| 818 | proof (intro iffI ballI) | |
| 819 | fix A assume eq: "AE x. f x = g x" | |
| 41705 | 820 | then show "?P f A = ?P g A" | 
| 821 | by (auto intro: positive_integral_cong_AE) | |
| 40859 | 822 | next | 
| 823 | assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" | |
| 824 | from this[THEN bspec, OF top] 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 | 825 | have g_fin: "integral\<^isup>P M g < \<omega>" by (simp cong: positive_integral_cong) | 
| 40859 | 826 |   { fix f g assume borel: "f \<in> borel_measurable M" "g \<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 | 827 | and g_fin: "integral\<^isup>P M g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A" | 
| 40859 | 828 |     let ?N = "{x\<in>space M. g x < f x}"
 | 
| 829 | have N: "?N \<in> sets M" using borel 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 | 830 | 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 | 831 | by (auto intro!: positive_integral_cong simp: indicator_def) | 
| 832 | also have "\<dots> = ?P f ?N - ?P g ?N" | |
| 833 | proof (rule positive_integral_diff) | |
| 834 | show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M" | |
| 835 | using borel 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 | 836 | have "?P g ?N \<le> integral\<^isup>P M g" | 
| 40859 | 837 | by (auto intro!: positive_integral_mono simp: indicator_def) | 
| 838 | then show "?P g ?N \<noteq> \<omega>" using g_fin by auto | |
| 839 | fix x assume "x \<in> space M" | |
| 840 | show "g x * indicator ?N x \<le> f x * indicator ?N x" | |
| 841 | by (auto simp: indicator_def) | |
| 842 | qed | |
| 843 | also have "\<dots> = 0" | |
| 844 | using eq[THEN bspec, OF N] by simp | |
| 845 |     finally have "\<mu> {x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = 0"
 | |
| 846 | using borel N by (subst (asm) positive_integral_0_iff) auto | |
| 847 |     moreover have "{x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = ?N"
 | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 848 | by (auto simp: pextreal_zero_le_diff) | 
| 40859 | 849 | ultimately have "?N \<in> null_sets" using N by simp } | 
| 850 | from this[OF borel g_fin eq] this[OF borel(2,1) fin] | |
| 851 |   have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} \<in> null_sets"
 | |
| 852 | using eq by (intro null_sets_Un) auto | |
| 853 |   also have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} = {x\<in>space M. f x \<noteq> g x}"
 | |
| 854 | by auto | |
| 855 | finally show "AE x. f x = g x" | |
| 856 | unfolding almost_everywhere_def by auto | |
| 857 | qed | |
| 858 | ||
| 859 | lemma (in finite_measure) density_unique_finite_measure: | |
| 860 | assumes borel: "f \<in> borel_measurable M" "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 | 861 | 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 | 862 | (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A") | 
| 863 | shows "AE x. f x = f' x" | |
| 864 | proof - | |
| 865 | let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A" | |
| 866 | 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 | 867 | interpret M: measure_space "M\<lparr> measure := ?\<nu>\<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 | 868 | using borel(1) by (rule measure_space_density) simp | 
| 40859 | 869 | have ac: "absolutely_continuous ?\<nu>" | 
| 870 | 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 | 871 | from split_space_into_finite_sets_and_rest[OF `measure_space (M\<lparr> measure := ?\<nu>\<rparr>)` ac] | 
| 40859 | 872 | obtain Q0 and Q :: "nat \<Rightarrow> 'a set" | 
| 873 | where Q: "disjoint_family Q" "range Q \<subseteq> sets M" | |
| 874 | and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)" | |
| 875 | 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 = \<omega>" | |
| 876 | and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<omega>" by force | |
| 877 | from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto | |
| 878 |   let ?N = "{x\<in>space M. f x \<noteq> f' x}"
 | |
| 879 | have "?N \<in> sets M" using borel by auto | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 880 | have *: "\<And>i x A. \<And>y::pextreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x" | 
| 40859 | 881 | unfolding indicator_def by auto | 
| 41705 | 882 | have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q | 
| 40859 | 883 | by (intro finite_density_unique[THEN iffD1] allI) | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 884 | (auto intro!: borel_measurable_pextreal_times f Int simp: *) | 
| 41705 | 885 | moreover have "AE x. ?f Q0 x = ?f' Q0 x" | 
| 40859 | 886 | proof (rule AE_I') | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 887 |     { fix f :: "'a \<Rightarrow> pextreal" 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 | 888 | and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)" | 
| 40859 | 889 |       let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
 | 
| 890 | have "(\<Union>i. ?A i) \<in> null_sets" | |
| 891 | proof (rule null_sets_UN) | |
| 892 | fix i have "?A i \<in> sets M" | |
| 893 | using borel Q0(1) 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 | 894 | have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. of_nat i * indicator (?A i) x \<partial>M)" | 
| 40859 | 895 | unfolding eq[OF `?A i \<in> sets M`] | 
| 896 | by (auto intro!: positive_integral_mono simp: indicator_def) | |
| 897 | also have "\<dots> = of_nat i * \<mu> (?A i)" | |
| 898 | using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator) | |
| 899 | also have "\<dots> < \<omega>" | |
| 900 | using `?A i \<in> sets M`[THEN finite_measure] by auto | |
| 901 | finally have "?\<nu> (?A i) \<noteq> \<omega>" by simp | |
| 902 | then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto | |
| 903 | qed | |
| 904 |       also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x < \<omega>}"
 | |
| 905 | by (auto simp: less_\<omega>_Ex_of_nat) | |
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 906 |       finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" by (simp add: pextreal_less_\<omega>) }
 | 
| 40859 | 907 | from this[OF borel(1) refl] this[OF borel(2) f] | 
| 908 |     have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>} \<in> null_sets" by simp_all
 | |
| 909 |     then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>}) \<in> null_sets" by (rule null_sets_Un)
 | |
| 910 |     show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq>
 | |
| 911 |       (Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>})" by (auto simp: indicator_def)
 | |
| 912 | qed | |
| 41705 | 913 | moreover have "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow> | 
| 40859 | 914 | ?f (space M) x = ?f' (space M) x" | 
| 915 | by (auto simp: indicator_def Q0) | |
| 41705 | 916 | ultimately have "AE x. ?f (space M) x = ?f' (space M) x" | 
| 917 | by (auto simp: all_AE_countable) | |
| 918 | then show "AE x. f x = f' x" by auto | |
| 40859 | 919 | qed | 
| 920 | ||
| 921 | lemma (in sigma_finite_measure) density_unique: | |
| 922 | assumes borel: "f \<in> borel_measurable M" "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 | 923 | 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 | 924 | (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A") | 
| 925 | shows "AE x. f x = f' x" | |
| 926 | proof - | |
| 927 | obtain h where h_borel: "h \<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 | 928 | and fin: "integral\<^isup>P M h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>" | 
| 40859 | 929 | using Ex_finite_integrable_function 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 | 930 | interpret h: measure_space "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h 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 | 931 | using h_borel by (rule measure_space_density) 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 | 932 | interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>" | 
| 40859 | 933 | 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 | 934 | 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 | 935 | interpret f: measure_space ?fM | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 936 | using borel(1) by (rule measure_space_density) 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 | 937 | 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 | 938 | interpret f': measure_space ?f'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 | 939 | using borel(2) by (rule measure_space_density) simp | 
| 40859 | 940 |   { fix A assume "A \<in> sets M"
 | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 941 |     then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pextreal)} = A"
 | 
| 40859 | 942 | using pos 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 | 943 | then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets" | 
| 40859 | 944 | using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) } | 
| 945 | note h_null_sets = this | |
| 946 |   { fix A assume "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 | 947 | have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 948 | using `A \<in> sets M` h_borel 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 | 949 | by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 950 | also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'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 | 951 | by (rule f'.positive_integral_cong_measure) (simp_all add: 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 | 952 | also have "\<dots> = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)" | 
| 40859 | 953 | using `A \<in> sets M` h_borel borel | 
| 954 | by (simp add: positive_integral_translated_density ac_simps 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 | 955 | 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)" . } | 
| 40859 | 956 | then have "h.almost_everywhere (\<lambda>x. f x = f' x)" | 
| 957 | using h_borel borel | |
| 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 | 958 | apply (intro h.density_unique_finite_measure) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 959 | apply (simp add: measurable_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 | 960 | apply (simp add: measurable_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 | 961 | by (simp add: positive_integral_translated_density) | 
| 40859 | 962 | then show "AE x. f x = f' x" | 
| 963 | unfolding h.almost_everywhere_def almost_everywhere_def | |
| 964 | by (auto simp add: h_null_sets) | |
| 965 | qed | |
| 966 | ||
| 967 | 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 | 968 | assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "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 | 969 | and f: "f \<in> borel_measurable 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 | 970 | 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 | 971 | shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)" | 
| 40859 | 972 | 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 | 973 | 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 | 974 | 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 | 975 | 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 | 976 | and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def) | 
| 40859 | 977 | from \<nu>.Ex_finite_integrable_function obtain h where | 
| 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 | 978 | h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<omega>" | 
| 41705 | 979 | and fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>" by auto | 
| 40859 | 980 | have "AE x. f x * h x \<noteq> \<omega>" | 
| 981 | proof (rule AE_I') | |
| 41705 | 982 | have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h | 
| 983 | by (subst \<nu>.positive_integral_cong_measure[symmetric, | |
| 984 | of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"]) | |
| 985 | (auto intro!: positive_integral_translated_density simp: eq) | |
| 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 | 986 | then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<omega>" | 
| 40859 | 987 | using h(2) by simp | 
| 988 |     then show "(\<lambda>x. f x * h x) -` {\<omega>} \<inter> space M \<in> null_sets"
 | |
| 989 | using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage) | |
| 990 | qed auto | |
| 991 | then show "AE x. f x \<noteq> \<omega>" | |
| 41705 | 992 | using fin by (auto elim!: AE_Ball_mp) | 
| 40859 | 993 | next | 
| 994 | assume AE: "AE x. f x \<noteq> \<omega>" | |
| 995 | 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 | 996 | 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 | 997 | 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 | 998 | and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def) | 
| 40859 | 999 |   def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<omega>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
 | 
| 1000 |   { fix i j have "A i \<inter> Q j \<in> sets M"
 | |
| 1001 | unfolding A_def using f Q | |
| 1002 | apply (rule_tac Int) | |
| 1003 | by (cases i) (auto intro: measurable_sets[OF f]) } | |
| 1004 | note A_in_sets = this | |
| 1005 | 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 | 1006 | show "sigma_finite_measure ?N" | 
| 40859 | 1007 | proof (default, intro exI conjI subsetI allI) | 
| 1008 | fix x assume "x \<in> range ?A" | |
| 1009 | 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 | 1010 | then show "x \<in> sets ?N" using A_in_sets by (cases "prod_decode n") auto | 
| 40859 | 1011 | next | 
| 1012 | have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)" | |
| 1013 | proof safe | |
| 1014 | fix x i j assume "x \<in> A i" "x \<in> Q j" | |
| 1015 | then show "x \<in> (\<Union>i. case prod_decode i of (i, j) \<Rightarrow> A i \<inter> Q j)" | |
| 1016 | by (intro UN_I[of "prod_encode (i,j)"]) auto | |
| 1017 | qed auto | |
| 1018 | also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto | |
| 1019 | also have "(\<Union>i. A i) = space M" | |
| 1020 | proof safe | |
| 1021 | fix x assume x: "x \<in> space M" | |
| 1022 | show "x \<in> (\<Union>i. A i)" | |
| 1023 | proof (cases "f x") | |
| 1024 | case infinite then show ?thesis using x unfolding A_def by (auto intro: exI[of _ 0]) | |
| 1025 | next | |
| 1026 | case (preal r) | |
| 1027 | with less_\<omega>_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by auto | |
| 1028 | then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"]) | |
| 1029 | qed | |
| 1030 | qed (auto simp: A_def) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1031 | finally show "(\<Union>i. ?A i) = space ?N" by simp | 
| 40859 | 1032 | next | 
| 1033 | fix n obtain i j where | |
| 1034 | [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto | |
| 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 | 1035 | have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<omega>" | 
| 40859 | 1036 | proof (cases i) | 
| 1037 | case 0 | |
| 1038 | have "AE x. f x * indicator (A i \<inter> Q j) x = 0" | |
| 41705 | 1039 | using AE by (auto simp: A_def `i = 0`) | 
| 1040 | from positive_integral_cong_AE[OF this] show ?thesis by simp | |
| 40859 | 1041 | next | 
| 1042 | case (Suc n) | |
| 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 | 1043 | then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<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 | 1044 | (\<integral>\<^isup>+x. of_nat (Suc n) * indicator (Q j) x \<partial>M)" | 
| 40859 | 1045 | by (auto intro!: positive_integral_mono simp: indicator_def A_def) | 
| 1046 | also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)" | |
| 1047 | using Q by (auto intro!: positive_integral_cmult_indicator) | |
| 1048 | also have "\<dots> < \<omega>" | |
| 1049 | using Q by auto | |
| 1050 | finally show ?thesis by simp | |
| 1051 | 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 | 1052 | then show "measure ?N (?A n) \<noteq> \<omega>" | 
| 40859 | 1053 | using A_in_sets Q eq by auto | 
| 1054 | qed | |
| 1055 | qed | |
| 1056 | ||
| 40871 | 1057 | section "Radon-Nikodym derivative" | 
| 38656 | 1058 | |
| 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 | 1059 | definition | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1060 | "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1061 | (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))" | 
| 38656 | 1062 | |
| 40859 | 1063 | 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 | 1064 | 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 | 1065 | 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 | 1066 | shows "RN_deriv M' \<nu>' x = RN_deriv M \<nu> x" | 
| 40859 | 1067 | 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 | 1068 | 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 | 1069 | using cong by (rule sigma_finite_measure_cong) | 
| 40859 | 1070 | 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 | 1071 | 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 | 1072 | by (simp add: cong \<nu> positive_integral_cong_measure[OF cong] measurable_def) | 
| 40859 | 1073 | qed | 
| 1074 | ||
| 38656 | 1075 | 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 | 1076 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 38656 | 1077 | 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 | 1078 | 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 | 1079 | and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)" | 
| 38656 | 1080 | (is "\<And>A. _ \<Longrightarrow> ?int A") | 
| 1081 | proof - | |
| 1082 | note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] | |
| 1083 | thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto | |
| 1084 | fix A assume "A \<in> sets M" | |
| 1085 | from Ex show "?int A" unfolding RN_deriv_def | |
| 1086 | by (rule someI2_ex) (simp add: `A \<in> sets M`) | |
| 1087 | qed | |
| 1088 | ||
| 40859 | 1089 | 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 | 1090 | assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" | 
| 40859 | 1091 | 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 | 1092 | shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)" | 
| 40859 | 1093 | 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 | 1094 | interpret \<nu>: measure_space "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 | 1095 | have "integral\<^isup>P (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 | 1096 | integral\<^isup>P (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<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 | 1097 | by (intro \<nu>.positive_integral_cong_measure[symmetric]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1098 | (simp_all add: RN_deriv(2)[OF \<nu>, symmetric]) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1099 | also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f 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 | 1100 | by (intro positive_integral_translated_density) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1101 | (simp_all add: RN_deriv[OF \<nu>] f) | 
| 40859 | 1102 | finally show ?thesis . | 
| 1103 | qed | |
| 1104 | ||
| 1105 | 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 | 1106 | assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" | 
| 40859 | 1107 | 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 | 1108 | 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 | 1109 | shows "AE x. f x = RN_deriv M \<nu> x" | 
| 40859 | 1110 | proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]]) | 
| 1111 | 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 | 1112 | 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 | 1113 | unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] .. | 
| 1114 | qed | |
| 1115 | ||
| 1116 | 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 | 1117 | assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<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 | 1118 | shows "AE x. RN_deriv M \<nu> x \<noteq> \<omega>" | 
| 40859 | 1119 | 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 | 1120 | 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 | 1121 | have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 40859 | 1122 | from sfm show ?thesis | 
| 1123 | using sigma_finite_iff_density_finite[OF \<nu> RN_deriv[OF \<nu> ac]] by simp | |
| 1124 | qed | |
| 1125 | ||
| 1126 | 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 | 1127 | assumes \<nu>: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>" | 
| 40859 | 1128 | 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 | 1129 | 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 | 1130 | 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 | 1131 | 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 | 1132 | (\<integral>x. real (RN_deriv M \<nu> x) * f x \<partial>M)" (is ?integral) | 
| 40859 | 1133 | 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 | 1134 | 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 | 1135 | have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default | 
| 41023 
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
 hoelzl parents: 
40871diff
changeset | 1136 | have minus_cong: "\<And>A B A' B'::pextreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp | 
| 40859 | 1137 | have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f 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 | 1138 | have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f unfolding measurable_def 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 | 1139 |   { fix f :: "'a \<Rightarrow> real"
 | 
| 
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 |     { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<omega>"
 | 
| 
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 | have "Real (real (RN_deriv M \<nu> x)) * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)" | 
| 40859 | 1142 | by (simp add: mult_le_0_iff) | 
| 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 | then have "RN_deriv M \<nu> x * Real (f x) = Real (real (RN_deriv M \<nu> x) * f x)" | 
| 40859 | 1144 | using * by (simp add: Real_real) } | 
| 41705 | 1145 | then have "(\<integral>\<^isup>+x. RN_deriv M \<nu> x * Real (f x) \<partial>M) = (\<integral>\<^isup>+x. Real (real (RN_deriv M \<nu> x) * f x) \<partial>M)" | 
| 1146 | using RN_deriv_finite[OF \<nu>] by (auto intro: positive_integral_cong_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 | 1147 | with this this f f' Nf | 
| 40859 | 1148 | show ?integral ?integrable | 
| 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 | 1149 | unfolding lebesgue_integral_def integrable_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 | 1150 | by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong | 
| 
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 | simp: RN_deriv_positive_integral[OF ms \<nu>(2)]) | 
| 40859 | 1152 | qed | 
| 1153 | ||
| 38656 | 1154 | 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 | 1155 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 38656 | 1156 | and ac: "absolutely_continuous \<nu>" | 
| 1157 |   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 | 1158 |   shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
 | 
| 38656 | 1159 | proof - | 
| 1160 | note deriv = RN_deriv[OF assms(1, 2)] | |
| 1161 |   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 | 1162 |   have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)"
 | 
| 38656 | 1163 | by (auto simp: indicator_def intro!: positive_integral_cong) | 
| 1164 |   thus ?thesis using positive_integral_cmult_indicator[OF `{x} \<in> sets M`]
 | |
| 1165 | by auto | |
| 1166 | qed | |
| 1167 | ||
| 1168 | 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 | 1169 | assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" | 
| 38656 | 1170 | and ac: "absolutely_continuous \<nu>" | 
| 1171 | 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 | 1172 |   shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
 | 
| 38656 | 1173 | proof - | 
| 1174 |   have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
 | |
| 1175 | from RN_deriv_singleton[OF assms(1,2) this] show ?thesis . | |
| 1176 | qed | |
| 1177 | ||
| 1178 | end |