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