| author | wenzelm | 
| Tue, 21 Feb 2012 13:37:03 +0100 | |
| changeset 46570 | 9c504481d270 | 
| parent 45777 | c36637603821 | 
| child 46731 | 5302e932d1e5 | 
| permissions | -rw-r--r-- | 
| 42067 | 1 | (* Title: HOL/Probability/Measure.thy | 
| 2 | Author: Lawrence C Paulson | |
| 3 | Author: Johannes Hölzl, TU München | |
| 4 | Author: Armin Heller, TU München | |
| 5 | *) | |
| 6 | ||
| 7 | header {* Properties about measure spaces *}
 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 8 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 9 | theory Measure | 
| 38656 | 10 | imports Caratheodory | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 11 | begin | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 12 | |
| 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 | 13 | lemma measure_algebra_more[simp]: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 14 | "\<lparr> space = A, sets = B, \<dots> = algebra.more M \<rparr> \<lparr> measure := m \<rparr> = | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 15 | \<lparr> space = A, sets = B, \<dots> = algebra.more (M \<lparr> measure := m \<rparr>) \<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 16 | by (cases M) simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 17 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 18 | lemma measure_algebra_more_eq[simp]: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 19 | "\<And>X. measure \<lparr> space = T, sets = A, \<dots> = algebra.more X \<rparr> = measure X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 20 | unfolding measure_space.splits by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 21 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 22 | lemma measure_sigma[simp]: "measure (sigma A) = measure A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 23 | unfolding sigma_def by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 24 | |
| 41831 | 25 | lemma algebra_measure_update[simp]: | 
| 26 | "algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> algebra M'" | |
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 27 | unfolding algebra_iff_Un by simp | 
| 41831 | 28 | |
| 29 | lemma sigma_algebra_measure_update[simp]: | |
| 30 | "sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> sigma_algebra M'" | |
| 31 | unfolding sigma_algebra_def sigma_algebra_axioms_def by simp | |
| 32 | ||
| 33 | lemma finite_sigma_algebra_measure_update[simp]: | |
| 34 | "finite_sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> finite_sigma_algebra M'" | |
| 35 | unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp | |
| 36 | ||
| 37 | lemma measurable_cancel_measure[simp]: | |
| 38 | "measurable M1 (M2\<lparr>measure := m2\<rparr>) = measurable M1 M2" | |
| 39 | "measurable (M2\<lparr>measure := m1\<rparr>) M1 = measurable M2 M1" | |
| 40 | unfolding measurable_def by auto | |
| 41 | ||
| 40859 | 42 | lemma inj_on_image_eq_iff: | 
| 43 | assumes "inj_on f S" | |
| 44 | assumes "A \<subseteq> S" "B \<subseteq> S" | |
| 45 | shows "(f ` A = f ` B) \<longleftrightarrow> (A = B)" | |
| 46 | proof - | |
| 47 | have "inj_on f (A \<union> B)" | |
| 48 | using assms by (auto intro: subset_inj_on) | |
| 49 | from inj_on_Un_image_eq_iff[OF this] | |
| 50 | show ?thesis . | |
| 51 | qed | |
| 52 | ||
| 53 | lemma image_vimage_inter_eq: | |
| 54 | assumes "f ` S = T" "X \<subseteq> T" | |
| 55 | shows "f ` (f -` X \<inter> S) = X" | |
| 56 | proof (intro antisym) | |
| 57 | have "f ` (f -` X \<inter> S) \<subseteq> f ` (f -` X)" by auto | |
| 58 | also have "\<dots> = X \<inter> range f" by simp | |
| 59 | also have "\<dots> = X" using assms by auto | |
| 60 | finally show "f ` (f -` X \<inter> S) \<subseteq> X" by auto | |
| 61 | next | |
| 62 | show "X \<subseteq> f ` (f -` X \<inter> S)" | |
| 63 | proof | |
| 64 | fix x assume "x \<in> X" | |
| 65 | then have "x \<in> T" using `X \<subseteq> T` by auto | |
| 66 | then obtain y where "x = f y" "y \<in> S" | |
| 67 | using assms by auto | |
| 68 |     then have "{y} \<subseteq> f -` X \<inter> S" using `x \<in> X` by auto
 | |
| 69 |     moreover have "x \<in> f ` {y}" using `x = f y` by auto
 | |
| 70 | ultimately show "x \<in> f ` (f -` X \<inter> S)" by auto | |
| 71 | qed | |
| 72 | qed | |
| 73 | ||
| 74 | text {*
 | |
| 75 | This formalisation of measure theory is based on the work of Hurd/Coble wand | |
| 76 | was later translated by Lawrence Paulson to Isabelle/HOL. Later it was | |
| 77 | modified to use the positive infinite reals and to prove the uniqueness of | |
| 78 | cut stable measures. | |
| 79 | *} | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 80 | |
| 38656 | 81 | section {* Equations for the measure function @{text \<mu>} *}
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 82 | |
| 38656 | 83 | lemma (in measure_space) measure_countably_additive: | 
| 84 | assumes "range A \<subseteq> sets M" "disjoint_family A" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 85 | shows "(\<Sum>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 86 | proof - | 
| 38656 | 87 | have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN) | 
| 88 | with ca assms show ?thesis by (simp add: countably_additive_def) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 89 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 90 | |
| 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 | 91 | lemma (in sigma_algebra) sigma_algebra_cong: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 92 | assumes "space N = space M" "sets N = sets M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 93 | shows "sigma_algebra N" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 94 | by default (insert sets_into_space, auto simp: assms) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 95 | |
| 38656 | 96 | lemma (in measure_space) measure_space_cong: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 97 | assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 98 | shows "measure_space N" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 99 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 100 | interpret N: sigma_algebra N by (intro sigma_algebra_cong assms) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 101 | show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 102 | proof | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 103 | show "positive N (measure N)" using assms by (auto simp: positive_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 104 | show "countably_additive N (measure N)" unfolding countably_additive_def | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 105 | proof safe | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 106 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets N" "disjoint_family A" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 107 | then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" unfolding assms by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 108 | from measure_countably_additive[of A] A this[THEN assms(1)] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 109 | show "(\<Sum>n. measure N (A n)) = measure N (UNION UNIV A)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 110 | unfolding assms by simp | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 111 | qed | 
| 38656 | 112 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 113 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 114 | |
| 38656 | 115 | lemma (in measure_space) additive: "additive M \<mu>" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 116 | using ca by (auto intro!: countably_additive_additive simp: positive_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 117 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 118 | lemma (in measure_space) measure_additive: | 
| 38656 | 119 |      "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {}
 | 
| 120 | \<Longrightarrow> \<mu> a + \<mu> b = \<mu> (a \<union> b)" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 121 | by (metis additiveD additive) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 122 | |
| 36624 | 123 | lemma (in measure_space) measure_mono: | 
| 124 | assumes "a \<subseteq> b" "a \<in> sets M" "b \<in> sets M" | |
| 38656 | 125 | shows "\<mu> a \<le> \<mu> b" | 
| 36624 | 126 | proof - | 
| 127 | have "b = a \<union> (b - a)" using assms by auto | |
| 128 |   moreover have "{} = a \<inter> (b - a)" by auto
 | |
| 38656 | 129 | ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 130 | using measure_additive[of a "b - a"] Diff[of b a] assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 131 | moreover have "\<mu> a + 0 \<le> \<mu> a + \<mu> (b - a)" using assms by (intro add_mono) auto | 
| 38656 | 132 | ultimately show "\<mu> a \<le> \<mu> b" by auto | 
| 36624 | 133 | qed | 
| 134 | ||
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 135 | lemma (in measure_space) measure_top: | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 136 | "A \<in> sets M \<Longrightarrow> \<mu> A \<le> \<mu> (space M)" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 137 | using sets_into_space[of A] by (auto intro!: measure_mono) | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 138 | |
| 38656 | 139 | lemma (in measure_space) measure_compl: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 140 | assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<infinity>" | 
| 38656 | 141 | shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 142 | proof - | 
| 38656 | 143 | have s_less_space: "\<mu> s \<le> \<mu> (space M)" | 
| 144 | using s by (auto intro!: measure_mono sets_into_space) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 145 | from s have "0 \<le> \<mu> s" by auto | 
| 38656 | 146 | have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s | 
| 147 | by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) | |
| 148 | also have "... = \<mu> s + \<mu> (space M - s)" | |
| 149 | by (rule additiveD [OF additive]) (auto simp add: s) | |
| 150 | finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" . | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 151 | then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 152 | using fin `0 \<le> \<mu> s` | 
| 43920 | 153 | unfolding ereal_eq_minus_iff by (auto simp: ac_simps) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 154 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 155 | |
| 38656 | 156 | lemma (in measure_space) measure_Diff: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 157 | assumes finite: "\<mu> B \<noteq> \<infinity>" | 
| 38656 | 158 | and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A" | 
| 159 | shows "\<mu> (A - B) = \<mu> A - \<mu> B" | |
| 160 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 161 | have "0 \<le> \<mu> B" using assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 162 | have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 163 | then have "\<mu> A = \<mu> ((A - B) \<union> B)" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 164 | also have "\<dots> = \<mu> (A - B) + \<mu> B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 165 | using measurable by (subst measure_additive[symmetric]) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 166 | finally show "\<mu> (A - B) = \<mu> A - \<mu> B" | 
| 43920 | 167 | unfolding ereal_eq_minus_iff | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 168 | using finite `0 \<le> \<mu> B` by auto | 
| 38656 | 169 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 170 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 171 | lemma (in measure_space) measure_countable_increasing: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 172 | assumes A: "range A \<subseteq> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 173 |       and A0: "A 0 = {}"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 174 | and ASuc: "\<And>n. A n \<subseteq> A (Suc n)" | 
| 38656 | 175 | shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 176 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 177 |   { fix n
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 178 | have "\<mu> (A n) = (\<Sum>i<n. \<mu> (A (Suc i) - A i))" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 179 | proof (induct n) | 
| 37032 | 180 | case 0 thus ?case by (auto simp add: A0) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 181 | next | 
| 38656 | 182 | case (Suc m) | 
| 33536 | 183 | have "A (Suc m) = A m \<union> (A (Suc m) - A m)" | 
| 184 | by (metis ASuc Un_Diff_cancel Un_absorb1) | |
| 38656 | 185 | hence "\<mu> (A (Suc m)) = | 
| 186 | \<mu> (A m) + \<mu> (A (Suc m) - A m)" | |
| 187 | by (subst measure_additive) | |
| 188 | (auto simp add: measure_additive range_subsetD [OF A]) | |
| 33536 | 189 | with Suc show ?case | 
| 190 | by simp | |
| 38656 | 191 | qed } | 
| 192 | note Meq = this | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 193 | have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)" | 
| 38656 | 194 | proof (rule UN_finite2_eq [where k=1], simp) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 195 | fix i | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 196 |       show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
 | 
| 33536 | 197 | proof (induct i) | 
| 198 | case 0 thus ?case by (simp add: A0) | |
| 199 | next | |
| 38656 | 200 | case (Suc i) | 
| 33536 | 201 | thus ?case | 
| 202 | by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc]) | |
| 203 | qed | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 204 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 205 | have A1: "\<And>i. A (Suc i) - A i \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 206 | by (metis A Diff range_subsetD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 207 | have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M" | 
| 37032 | 208 | by (blast intro: range_subsetD [OF A]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 209 | have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))" | 
| 43920 | 210 | using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 211 | also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)" | 
| 38656 | 212 | by (rule measure_countably_additive) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 213 | (auto simp add: disjoint_family_Suc ASuc A1 A2) | 
| 38656 | 214 | also have "... = \<mu> (\<Union>i. A i)" | 
| 215 | by (simp add: Aeq) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 216 | finally have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 217 | then show ?thesis by (auto simp add: Meq) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 218 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 219 | |
| 38656 | 220 | lemma (in measure_space) continuity_from_below: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 221 | assumes A: "range A \<subseteq> sets M" and "incseq A" | 
| 38656 | 222 | shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)" | 
| 223 | proof - | |
| 224 |   have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 225 | using A by (auto intro!: SUPR_eq exI split: nat.split) | 
| 38656 | 226 |   have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
 | 
| 227 | by (auto simp add: split: nat.splits) | |
| 228 |   have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
 | |
| 229 | by simp | |
| 230 |   have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 231 | using range_subsetD[OF A] incseq_SucD[OF `incseq A`] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 232 | by (force split: nat.splits intro!: measure_countable_increasing) | 
| 38656 | 233 |   also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
 | 
| 234 | by (simp add: ueq) | |
| 235 |   finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
 | |
| 236 | thus ?thesis unfolding meq * comp_def . | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 237 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 238 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 239 | lemma (in measure_space) measure_incseq: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 240 | assumes "range B \<subseteq> sets M" "incseq B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 241 | shows "incseq (\<lambda>i. \<mu> (B i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 242 | using assms by (auto simp: incseq_def intro!: measure_mono) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 243 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 244 | lemma (in measure_space) continuity_from_below_Lim: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 245 | assumes A: "range A \<subseteq> sets M" "incseq A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 246 | shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)" | 
| 43920 | 247 | using LIMSEQ_ereal_SUPR[OF measure_incseq, OF A] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 248 | continuity_from_below[OF A] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 249 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 250 | lemma (in measure_space) measure_decseq: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 251 | assumes "range B \<subseteq> sets M" "decseq B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 252 | shows "decseq (\<lambda>i. \<mu> (B i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 253 | using assms by (auto simp: decseq_def intro!: measure_mono) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 254 | |
| 38656 | 255 | lemma (in measure_space) continuity_from_above: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 256 | assumes A: "range A \<subseteq> sets M" and "decseq A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 257 | and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 38656 | 258 | shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 259 | proof - | 
| 38656 | 260 | have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)" | 
| 261 | using A by (auto intro!: measure_mono) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 262 | hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 263 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 264 | have A0: "0 \<le> \<mu> (A 0)" using A by auto | 
| 38656 | 265 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 266 | have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))" | 
| 43920 | 267 | by (simp add: ereal_SUPR_uminus minus_ereal_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 268 | also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))" | 
| 43920 | 269 | unfolding minus_ereal_def using A0 assms | 
| 270 | by (subst SUPR_ereal_add) (auto simp add: measure_decseq) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 271 | also have "\<dots> = (SUP n. \<mu> (A 0 - A n))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 272 | using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto | 
| 38656 | 273 | also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)" | 
| 274 | proof (rule continuity_from_below) | |
| 275 | show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M" | |
| 276 | using A by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 277 | show "incseq (\<lambda>n. A 0 - A n)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 278 | using `decseq A` by (auto simp add: incseq_def decseq_def) | 
| 38656 | 279 | qed | 
| 280 | also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 281 | using A finite * by (simp, subst measure_Diff) auto | 
| 38656 | 282 | finally show ?thesis | 
| 43920 | 283 | unfolding ereal_minus_eq_minus_iff using finite A0 by auto | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 284 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 285 | |
| 38656 | 286 | lemma (in measure_space) measure_insert: | 
| 287 |   assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
 | |
| 288 |   shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 289 | proof - | 
| 38656 | 290 |   have "{x} \<inter> A = {}" using `x \<notin> A` by auto
 | 
| 291 | from measure_additive[OF sets this] show ?thesis by simp | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 292 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 293 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 294 | lemma (in measure_space) measure_setsum: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 295 | assumes "finite S" and "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 296 | assumes disj: "disjoint_family_on A S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 297 | shows "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>i\<in>S. A i)" | 
| 38656 | 298 | using assms proof induct | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 299 | case (insert i S) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 300 | then have "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>a\<in>S. A a)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 301 | by (auto intro: disjoint_family_on_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 302 |   moreover have "A i \<inter> (\<Union>a\<in>S. A a) = {}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 303 | using `disjoint_family_on A (insert i S)` `i \<notin> S` | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 304 | by (auto simp: disjoint_family_on_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 305 | ultimately show ?case using insert | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 306 | by (auto simp: measure_additive finite_UN) | 
| 38656 | 307 | qed simp | 
| 35582 | 308 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 309 | lemma (in measure_space) measure_finite_singleton: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 310 |   assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 311 |   shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 312 |   using measure_setsum[of S "\<lambda>x. {x}", OF assms]
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 313 | by (auto simp: disjoint_family_on_def) | 
| 35582 | 314 | |
| 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 | 315 | lemma finite_additivity_sufficient: | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 316 | assumes "sigma_algebra M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 317 | assumes fin: "finite (space M)" and pos: "positive M (measure M)" and add: "additive M (measure M)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 318 | shows "measure_space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 319 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 320 | interpret sigma_algebra M by fact | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 321 | show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 322 | proof | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 323 | show [simp]: "positive M (measure M)" using pos by (simp add: positive_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 324 | show "countably_additive M (measure M)" | 
| 38656 | 325 | proof (auto simp add: countably_additive_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 326 | fix A :: "nat \<Rightarrow> 'a set" | 
| 38656 | 327 | assume A: "range A \<subseteq> sets M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 328 | and disj: "disjoint_family A" | 
| 38656 | 329 | and UnA: "(\<Union>i. A i) \<in> sets M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 330 |       def I \<equiv> "{i. A i \<noteq> {}}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 331 | have "Union (A ` I) \<subseteq> space M" using A | 
| 38656 | 332 | by auto (metis range_subsetD subsetD sets_into_space) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 333 | hence "finite (A ` I)" | 
| 38656 | 334 | by (metis finite_UnionD finite_subset fin) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 335 | moreover have "inj_on A I" using disj | 
| 38656 | 336 | by (auto simp add: I_def disjoint_family_on_def inj_on_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 337 | ultimately have finI: "finite I" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 338 | by (metis finite_imageD) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 339 |       hence "\<exists>N. \<forall>m\<ge>N. A m = {}"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 340 |         proof (cases "I = {}")
 | 
| 38656 | 341 | case True thus ?thesis by (simp add: I_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 342 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 343 | case False | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 344 | hence "\<forall>i\<in>I. i < Suc(Max I)" | 
| 38656 | 345 | by (simp add: Max_less_iff [symmetric] finI) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 346 |           hence "\<forall>m \<ge> Suc(Max I). A m = {}"
 | 
| 38656 | 347 | by (simp add: I_def) (metis less_le_not_le) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 348 | thus ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 349 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 350 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 351 |       then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 352 | then have "\<forall>m\<ge>N. measure M (A m) = 0" using pos[unfolded positive_def] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 353 | then have "(\<Sum>n. measure M (A n)) = (\<Sum>m<N. measure M (A m))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 354 | by (simp add: suminf_finite) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 355 | also have "... = measure M (\<Union>i<N. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 356 | proof (induct N) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 357 | case 0 thus ?case using pos[unfolded positive_def] by simp | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 358 | next | 
| 38656 | 359 | case (Suc n) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 360 | have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)" | 
| 38656 | 361 | proof (rule Caratheodory.additiveD [OF add]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 362 |               show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
 | 
| 35582 | 363 | by (auto simp add: disjoint_family_on_def nat_less_le) blast | 
| 38656 | 364 | show "A n \<in> sets M" using A | 
| 365 | by force | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 366 | show "(\<Union>i<n. A i) \<in> sets M" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 367 | proof (induct n) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 368 | case 0 thus ?case by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 369 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 370 | case (Suc n) thus ?case using A | 
| 38656 | 371 | by (simp add: lessThan_Suc Un range_subsetD) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 372 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 373 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 374 | thus ?case using Suc | 
| 38656 | 375 | by (simp add: lessThan_Suc) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 376 | qed | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 377 | also have "... = measure M (\<Union>i. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 378 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 379 | have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 380 | by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 381 | thus ?thesis by simp | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 382 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 383 | finally show "(\<Sum>n. measure M (A n)) = measure M (\<Union>i. A i)" . | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 384 | qed | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 385 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 386 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 387 | |
| 35692 | 388 | lemma (in measure_space) measure_setsum_split: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 389 | assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 390 | assumes "(\<Union>i\<in>S. B i) = space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 391 | assumes "disjoint_family_on B S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 392 | shows "\<mu> A = (\<Sum>i\<in>S. \<mu> (A \<inter> (B i)))" | 
| 35692 | 393 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 394 | have *: "\<mu> A = \<mu> (\<Union>i\<in>S. A \<inter> B i)" | 
| 35692 | 395 | using assms by auto | 
| 396 | show ?thesis unfolding * | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 397 | proof (rule measure_setsum[symmetric]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 398 | show "disjoint_family_on (\<lambda>i. A \<inter> B i) S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 399 | using `disjoint_family_on B S` | 
| 35692 | 400 | unfolding disjoint_family_on_def by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 401 | qed (insert assms, auto) | 
| 35692 | 402 | qed | 
| 403 | ||
| 38656 | 404 | lemma (in measure_space) measure_subadditive: | 
| 405 | assumes measurable: "A \<in> sets M" "B \<in> sets M" | |
| 406 | shows "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B" | |
| 407 | proof - | |
| 408 | from measure_additive[of A "B - A"] | |
| 409 | have "\<mu> (A \<union> B) = \<mu> A + \<mu> (B - A)" | |
| 410 | using assms by (simp add: Diff) | |
| 411 | also have "\<dots> \<le> \<mu> A + \<mu> B" | |
| 412 | using assms by (auto intro!: add_left_mono measure_mono) | |
| 413 | finally show ?thesis . | |
| 414 | qed | |
| 415 | ||
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 416 | lemma (in measure_space) measure_subadditive_finite: | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 417 | assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 418 | shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 419 | using assms proof induct | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 420 | case (insert i I) | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 421 | then have "\<mu> (\<Union>i\<in>insert i I. A i) = \<mu> (A i \<union> (\<Union>i\<in>I. A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 422 | by simp | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 423 | also have "\<dots> \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 424 | using insert by (intro measure_subadditive finite_UN) auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 425 | also have "\<dots> \<le> \<mu> (A i) + (\<Sum>i\<in>I. \<mu> (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 426 | using insert by (intro add_mono) auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 427 | also have "\<dots> = (\<Sum>i\<in>insert i I. \<mu> (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 428 | using insert by auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 429 | finally show ?case . | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 430 | qed simp | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 431 | |
| 40859 | 432 | lemma (in measure_space) measure_eq_0: | 
| 433 | assumes "N \<in> sets M" and "\<mu> N = 0" and "K \<subseteq> N" and "K \<in> sets M" | |
| 434 | shows "\<mu> K = 0" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 435 | using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto | 
| 40859 | 436 | |
| 39092 | 437 | lemma (in measure_space) measure_finitely_subadditive: | 
| 438 | assumes "finite I" "A ` I \<subseteq> sets M" | |
| 439 | shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))" | |
| 440 | using assms proof induct | |
| 441 | case (insert i I) | |
| 442 | then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN) | |
| 443 | then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)" | |
| 444 | using insert by (simp add: measure_subadditive) | |
| 445 | also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))" | |
| 446 | using insert by (auto intro!: add_left_mono) | |
| 447 | finally show ?case . | |
| 448 | qed simp | |
| 449 | ||
| 40859 | 450 | lemma (in measure_space) measure_countably_subadditive: | 
| 38656 | 451 | assumes "range f \<subseteq> sets M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 452 | shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>i. \<mu> (f i))" | 
| 38656 | 453 | proof - | 
| 454 | have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)" | |
| 455 | unfolding UN_disjointed_eq .. | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 456 | also have "\<dots> = (\<Sum>i. \<mu> (disjointed f i))" | 
| 38656 | 457 | using range_disjointed_sets[OF assms] measure_countably_additive | 
| 458 | by (simp add: disjoint_family_disjointed comp_def) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 459 | also have "\<dots> \<le> (\<Sum>i. \<mu> (f i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 460 | using range_disjointed_sets[OF assms] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 461 | by (auto intro!: suminf_le_pos measure_mono positive_measure disjointed_subset) | 
| 38656 | 462 | finally show ?thesis . | 
| 463 | qed | |
| 464 | ||
| 40859 | 465 | lemma (in measure_space) measure_UN_eq_0: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 466 | assumes "\<And>i::nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M" | 
| 40859 | 467 | shows "\<mu> (\<Union> i. N i) = 0" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 468 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 469 | have "0 \<le> \<mu> (\<Union> i. N i)" using assms by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 470 | moreover have "\<mu> (\<Union> i. N i) \<le> 0" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 471 | using measure_countably_subadditive[OF assms(2)] assms(1) by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 472 | ultimately show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 473 | qed | 
| 40859 | 474 | |
| 39092 | 475 | lemma (in measure_space) measure_inter_full_set: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 476 | assumes "S \<in> sets M" "T \<in> sets M" and fin: "\<mu> (T - S) \<noteq> \<infinity>" | 
| 39092 | 477 | assumes T: "\<mu> T = \<mu> (space M)" | 
| 478 | shows "\<mu> (S \<inter> T) = \<mu> S" | |
| 479 | proof (rule antisym) | |
| 480 | show " \<mu> (S \<inter> T) \<le> \<mu> S" | |
| 481 | using assms by (auto intro!: measure_mono) | |
| 482 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 483 | have pos: "0 \<le> \<mu> (T - S)" using assms by auto | 
| 39092 | 484 | show "\<mu> S \<le> \<mu> (S \<inter> T)" | 
| 485 | proof (rule ccontr) | |
| 486 | assume contr: "\<not> ?thesis" | |
| 487 | have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))" | |
| 488 | unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"]) | |
| 489 | also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)" | |
| 490 | using assms by (auto intro!: measure_subadditive) | |
| 491 | also have "\<dots> < \<mu> (T - S) + \<mu> S" | |
| 43920 | 492 | using fin contr pos by (intro ereal_less_add) auto | 
| 39092 | 493 | also have "\<dots> = \<mu> (T \<union> S)" | 
| 494 | using assms by (subst measure_additive) auto | |
| 495 | also have "\<dots> \<le> \<mu> (space M)" | |
| 496 | using assms sets_into_space by (auto intro!: measure_mono) | |
| 497 | finally show False .. | |
| 498 | qed | |
| 499 | qed | |
| 500 | ||
| 40859 | 501 | lemma measure_unique_Int_stable: | 
| 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 | 502 |   fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 503 | assumes "Int_stable E" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 504 | and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 505 | and M: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<mu>\<rparr>" (is "measure_space ?M") | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 506 | and N: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<nu>\<rparr>" (is "measure_space ?N") | 
| 40859 | 507 | and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 508 | and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 509 | assumes "X \<in> sets (sigma E)" | 
| 40859 | 510 | shows "\<mu> X = \<nu> X" | 
| 511 | 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 | 512 |   let "?D F" = "{D. D \<in> sets (sigma E) \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 513 | interpret M: measure_space ?M | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 514 | where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \<mu>" by (simp_all add: M) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 515 | interpret N: measure_space ?N | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 516 | where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \<nu>" by (simp_all add: N) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 517 |   { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<infinity>"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 518 | then have [intro]: "F \<in> sets (sigma E)" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 519 | have "\<nu> F \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` `F \<in> sets E` eq by simp | 
| 40859 | 520 | interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>" | 
| 521 | proof (rule dynkin_systemI, simp_all) | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 522 | fix A assume "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 523 | then show "A \<subseteq> space E" using M.sets_into_space by auto | 
| 40859 | 524 | next | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 525 | have "F \<inter> space E = F" using `F \<in> sets E` by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 526 | then show "\<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 527 | using `F \<in> sets E` eq by auto | 
| 40859 | 528 | next | 
| 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 | 529 | fix A assume *: "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 530 | then have **: "F \<inter> (space (sigma E) - A) = F - (F \<inter> A)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 531 | and [intro]: "F \<inter> A \<in> sets (sigma E)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 532 | using `F \<in> sets E` M.sets_into_space by auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 533 | have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: N.measure_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 534 | then have "\<nu> (F \<inter> A) \<noteq> \<infinity>" using `\<nu> F \<noteq> \<infinity>` by auto | 
| 40859 | 535 | have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 536 | then have "\<mu> (F \<inter> A) \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` by auto | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 537 | then have "\<mu> (F \<inter> (space (sigma E) - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding ** | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 538 | using `F \<inter> A \<in> sets (sigma E)` by (auto intro!: M.measure_Diff) | 
| 40859 | 539 | also have "\<dots> = \<nu> F - \<nu> (F \<inter> A)" using eq `F \<in> sets E` * by simp | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 540 | also have "\<dots> = \<nu> (F \<inter> (space (sigma E) - A))" unfolding ** | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 541 | using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<infinity>` by (auto intro!: N.measure_Diff[symmetric]) | 
| 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 | 542 | finally show "space E - A \<in> sets (sigma E) \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 543 | using * by auto | 
| 40859 | 544 | next | 
| 545 | fix A :: "nat \<Rightarrow> 'a set" | |
| 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 | 546 |       assume "disjoint_family A" "range A \<subseteq> {X \<in> sets (sigma E). \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}"
 | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 547 | then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets (sigma E)" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 548 | "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. \<mu> (F \<inter> A i) = \<nu> (F \<inter> A i)" "range A \<subseteq> sets (sigma E)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 549 | by (auto simp: disjoint_family_on_def subset_eq) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 550 | then show "(\<Union>x. A x) \<in> sets (sigma E) \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))" | 
| 40859 | 551 | by (auto simp: M.measure_countably_additive[symmetric] | 
| 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 | 552 | N.measure_countably_additive[symmetric] | 
| 40859 | 553 | simp del: UN_simps) | 
| 554 | qed | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 555 | have *: "sets (sigma E) = sets \<lparr>space = space E, sets = ?D F\<rparr>" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 556 | using `F \<in> sets E` `Int_stable E` | 
| 40859 | 557 | by (intro D.dynkin_lemma) | 
| 558 | (auto simp add: sets_sigma Int_stable_def eq intro: sigma_sets.Basic) | |
| 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 | 559 | have "\<And>D. D \<in> sets (sigma E) \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 560 | by (subst (asm) *) auto } | 
| 40859 | 561 | note * = this | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 562 | let "?A i" = "A i \<inter> X" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 563 | have A': "range ?A \<subseteq> sets (sigma E)" "incseq ?A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 564 | using A(1,2) `X \<in> sets (sigma E)` by (auto simp: incseq_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 565 |   { fix i have "\<mu> (?A i) = \<nu> (?A i)"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 566 | using *[of "A i" X] `X \<in> sets (sigma E)` A finite by auto } | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 567 | with M.continuity_from_below[OF A'] N.continuity_from_below[OF A'] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 568 | show ?thesis using A(3) `X \<in> sets (sigma E)` by auto | 
| 40859 | 569 | qed | 
| 570 | ||
| 571 | section "@{text \<mu>}-null sets"
 | |
| 572 | ||
| 573 | abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
 | |
| 574 | ||
| 42866 | 575 | sublocale measure_space \<subseteq> nullsets!: ring_of_sets "\<lparr> space = space M, sets = null_sets \<rparr>" | 
| 576 | where "space \<lparr> space = space M, sets = null_sets \<rparr> = space M" | |
| 577 | and "sets \<lparr> space = space M, sets = null_sets \<rparr> = null_sets" | |
| 578 | proof - | |
| 579 |   { fix A B assume sets: "A \<in> sets M" "B \<in> sets M"
 | |
| 580 | moreover then have "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B" "\<mu> (A - B) \<le> \<mu> A" | |
| 581 | by (auto intro!: measure_subadditive measure_mono) | |
| 582 | moreover assume "\<mu> B = 0" "\<mu> A = 0" | |
| 583 | ultimately have "\<mu> (A - B) = 0" "\<mu> (A \<union> B) = 0" | |
| 584 | by (auto intro!: antisym) } | |
| 585 | note null = this | |
| 586 | show "ring_of_sets \<lparr> space = space M, sets = null_sets \<rparr>" | |
| 587 | by default (insert sets_into_space null, auto) | |
| 588 | qed simp_all | |
| 40859 | 589 | |
| 590 | lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))" | |
| 591 | proof - | |
| 592 | have "(\<Union>i. N i) = (\<Union>i. (N \<circ> Countable.from_nat) i)" | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 593 | unfolding SUP_def image_compose | 
| 40859 | 594 | unfolding surj_from_nat .. | 
| 595 | then show ?thesis by simp | |
| 596 | qed | |
| 597 | ||
| 598 | lemma (in measure_space) null_sets_UN[intro]: | |
| 599 | assumes "\<And>i::'i::countable. N i \<in> null_sets" | |
| 600 | shows "(\<Union>i. N i) \<in> null_sets" | |
| 601 | proof (intro conjI CollectI) | |
| 602 | show "(\<Union>i. N i) \<in> sets M" using assms by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 603 | then have "0 \<le> \<mu> (\<Union>i. N i)" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 604 | moreover have "\<mu> (\<Union>i. N i) \<le> (\<Sum>n. \<mu> (N (Countable.from_nat n)))" | 
| 40859 | 605 | unfolding UN_from_nat[of N] | 
| 606 | using assms by (intro measure_countably_subadditive) auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 607 | ultimately show "\<mu> (\<Union>i. N i) = 0" using assms by auto | 
| 40859 | 608 | qed | 
| 609 | ||
| 610 | lemma (in measure_space) null_set_Int1: | |
| 611 | assumes "B \<in> null_sets" "A \<in> sets M" shows "A \<inter> B \<in> null_sets" | |
| 612 | using assms proof (intro CollectI conjI) | |
| 613 | show "\<mu> (A \<inter> B) = 0" using assms by (intro measure_eq_0[of B "A \<inter> B"]) auto | |
| 614 | qed auto | |
| 615 | ||
| 616 | lemma (in measure_space) null_set_Int2: | |
| 617 | assumes "B \<in> null_sets" "A \<in> sets M" shows "B \<inter> A \<in> null_sets" | |
| 618 | using assms by (subst Int_commute) (rule null_set_Int1) | |
| 619 | ||
| 620 | lemma (in measure_space) measure_Diff_null_set: | |
| 621 | assumes "B \<in> null_sets" "A \<in> sets M" | |
| 622 | shows "\<mu> (A - B) = \<mu> A" | |
| 623 | proof - | |
| 624 | have *: "A - B = (A - (A \<inter> B))" by auto | |
| 625 | have "A \<inter> B \<in> null_sets" using assms by (rule null_set_Int1) | |
| 626 | then show ?thesis | |
| 627 | unfolding * using assms | |
| 628 | by (subst measure_Diff) auto | |
| 629 | qed | |
| 630 | ||
| 631 | lemma (in measure_space) null_set_Diff: | |
| 632 | assumes "B \<in> null_sets" "A \<in> sets M" shows "B - A \<in> null_sets" | |
| 633 | using assms proof (intro CollectI conjI) | |
| 634 | show "\<mu> (B - A) = 0" using assms by (intro measure_eq_0[of B "B - A"]) auto | |
| 635 | qed auto | |
| 636 | ||
| 637 | lemma (in measure_space) measure_Un_null_set: | |
| 638 | assumes "A \<in> sets M" "B \<in> null_sets" | |
| 639 | shows "\<mu> (A \<union> B) = \<mu> A" | |
| 640 | proof - | |
| 641 | have *: "A \<union> B = A \<union> (B - A)" by auto | |
| 642 | have "B - A \<in> null_sets" using assms(2,1) by (rule null_set_Diff) | |
| 643 | then show ?thesis | |
| 644 | unfolding * using assms | |
| 645 | by (subst measure_additive[symmetric]) auto | |
| 646 | qed | |
| 647 | ||
| 40871 | 648 | section "Formalise almost everywhere" | 
| 649 | ||
| 650 | definition (in measure_space) | |
| 651 |   almost_everywhere :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "AE " 10) where
 | |
| 652 |   "almost_everywhere P \<longleftrightarrow> (\<exists>N\<in>null_sets. { x \<in> space M. \<not> P x } \<subseteq> N)"
 | |
| 653 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 654 | syntax | 
| 45342 
5c760e1692b3
proper syntactic category for abstraction syntax, to avoid low-level exception for malformed "\<integral> x y. f \<partial>M", for example;
 wenzelm parents: 
44928diff
changeset | 655 |   "_almost_everywhere" :: "pttrn \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 656 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 657 | translations | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 658 | "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 659 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 660 | lemma (in measure_space) AE_cong_measure: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 661 | assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 662 | shows "(AE x in N. P x) \<longleftrightarrow> (AE x. P x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 663 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 664 | interpret N: measure_space N | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 665 | by (rule measure_space_cong) fact+ | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 666 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 667 | unfolding N.almost_everywhere_def almost_everywhere_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 668 | by (auto simp: assms) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 669 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 670 | |
| 40871 | 671 | lemma (in measure_space) AE_I': | 
| 672 |   "N \<in> null_sets \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x. P x)"
 | |
| 673 | unfolding almost_everywhere_def by auto | |
| 674 | ||
| 675 | lemma (in measure_space) AE_iff_null_set: | |
| 676 |   assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
 | |
| 677 |   shows "(AE x. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets"
 | |
| 678 | proof | |
| 679 | assume "AE x. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "\<mu> N = 0" | |
| 680 | unfolding almost_everywhere_def by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 681 | have "0 \<le> \<mu> ?P" using assms by simp | 
| 40871 | 682 | moreover have "\<mu> ?P \<le> \<mu> N" | 
| 683 | using assms N(1,2) by (auto intro: measure_mono) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 684 | ultimately have "\<mu> ?P = 0" unfolding `\<mu> N = 0` by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 685 | then show "?P \<in> null_sets" using assms by simp | 
| 40871 | 686 | next | 
| 687 | assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I') | |
| 688 | qed | |
| 689 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 690 | lemma (in measure_space) AE_iff_measurable: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 691 |   "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x. P x) \<longleftrightarrow> \<mu> N = 0"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 692 | using AE_iff_null_set[of P] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 693 | |
| 40859 | 694 | lemma (in measure_space) AE_True[intro, simp]: "AE x. True" | 
| 695 | unfolding almost_everywhere_def by auto | |
| 696 | ||
| 697 | lemma (in measure_space) AE_E[consumes 1]: | |
| 698 | assumes "AE x. P x" | |
| 699 |   obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M"
 | |
| 700 | using assms unfolding almost_everywhere_def by auto | |
| 701 | ||
| 41705 | 702 | lemma (in measure_space) AE_E2: | 
| 703 |   assumes "AE x. P x" "{x\<in>space M. P x} \<in> sets M"
 | |
| 704 |   shows "\<mu> {x\<in>space M. \<not> P x} = 0" (is "\<mu> ?P = 0")
 | |
| 705 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 706 |   have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 707 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 708 | with AE_iff_null_set[of P] assms show ?thesis by auto | 
| 41705 | 709 | qed | 
| 710 | ||
| 40859 | 711 | lemma (in measure_space) AE_I: | 
| 712 |   assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M"
 | |
| 713 | shows "AE x. P x" | |
| 714 | using assms unfolding almost_everywhere_def by auto | |
| 715 | ||
| 41705 | 716 | lemma (in measure_space) AE_mp[elim!]: | 
| 40859 | 717 | assumes AE_P: "AE x. P x" and AE_imp: "AE x. P x \<longrightarrow> Q x" | 
| 718 | shows "AE x. Q x" | |
| 719 | proof - | |
| 720 |   from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
 | |
| 721 | and A: "A \<in> sets M" "\<mu> A = 0" | |
| 722 | by (auto elim!: AE_E) | |
| 723 | ||
| 724 |   from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
 | |
| 725 | and B: "B \<in> sets M" "\<mu> B = 0" | |
| 726 | by (auto elim!: AE_E) | |
| 727 | ||
| 728 | show ?thesis | |
| 729 | proof (intro AE_I) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 730 | have "0 \<le> \<mu> (A \<union> B)" using A B by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 731 | moreover have "\<mu> (A \<union> B) \<le> 0" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 732 | using measure_subadditive[of A B] A B by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 733 | ultimately show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B by auto | 
| 40859 | 734 |     show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
 | 
| 735 | using P imp by auto | |
| 736 | qed | |
| 737 | qed | |
| 738 | ||
| 41705 | 739 | lemma (in measure_space) | 
| 740 | shows AE_iffI: "AE x. P x \<Longrightarrow> AE x. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x. Q x" | |
| 741 | and AE_disjI1: "AE x. P x \<Longrightarrow> AE x. P x \<or> Q x" | |
| 742 | and AE_disjI2: "AE x. Q x \<Longrightarrow> AE x. P x \<or> Q x" | |
| 743 | and AE_conjI: "AE x. P x \<Longrightarrow> AE x. Q x \<Longrightarrow> AE x. P x \<and> Q x" | |
| 744 | and AE_conj_iff[simp]: "(AE x. P x \<and> Q x) \<longleftrightarrow> (AE x. P x) \<and> (AE x. Q x)" | |
| 745 | by auto | |
| 40859 | 746 | |
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 747 | lemma (in measure_space) AE_measure: | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 748 |   assumes AE: "AE x. P x" and sets: "{x\<in>space M. P x} \<in> sets M"
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 749 |   shows "\<mu> {x\<in>space M. P x} = \<mu> (space M)"
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 750 | proof - | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 751 | from AE_E[OF AE] guess N . note N = this | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 752 |   with sets have "\<mu> (space M) \<le> \<mu> ({x\<in>space M. P x} \<union> N)"
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 753 | by (intro measure_mono) auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 754 |   also have "\<dots> \<le> \<mu> {x\<in>space M. P x} + \<mu> N"
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 755 | using sets N by (intro measure_subadditive) auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 756 |   also have "\<dots> = \<mu> {x\<in>space M. P x}" using N by simp
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 757 |   finally show "\<mu> {x\<in>space M. P x} = \<mu> (space M)"
 | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 758 | using measure_top[OF sets] by auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 759 | qed | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 760 | |
| 41705 | 761 | lemma (in measure_space) AE_space: "AE x. x \<in> space M" | 
| 40859 | 762 |   by (rule AE_I[where N="{}"]) auto
 | 
| 763 | ||
| 41705 | 764 | lemma (in measure_space) AE_I2[simp, intro]: | 
| 765 | "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x. P x" | |
| 766 | using AE_space by auto | |
| 767 | ||
| 768 | lemma (in measure_space) AE_Ball_mp: | |
| 769 | "\<forall>x\<in>space M. P x \<Longrightarrow> AE x. P x \<longrightarrow> Q x \<Longrightarrow> AE x. Q x" | |
| 770 | by auto | |
| 771 | ||
| 772 | lemma (in measure_space) AE_cong[cong]: | |
| 773 | "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x. P x) \<longleftrightarrow> (AE x. Q x)" | |
| 774 | by auto | |
| 40859 | 775 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 776 | lemma (in measure_space) AE_all_countable: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 777 | "(AE x. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x. P i x)" | 
| 40859 | 778 | proof | 
| 779 | assume "\<forall>i. AE x. P i x" | |
| 780 | from this[unfolded almost_everywhere_def Bex_def, THEN choice] | |
| 781 |   obtain N where N: "\<And>i. N i \<in> null_sets" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
 | |
| 782 |   have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
 | |
| 783 | also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto | |
| 784 |   finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
 | |
| 785 | moreover from N have "(\<Union>i. N i) \<in> null_sets" | |
| 786 | by (intro null_sets_UN) auto | |
| 787 | ultimately show "AE x. \<forall>i. P i x" | |
| 788 | unfolding almost_everywhere_def by auto | |
| 41705 | 789 | qed auto | 
| 40859 | 790 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 791 | lemma (in measure_space) AE_finite_all: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 792 | assumes f: "finite S" shows "(AE x. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x. P i x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 793 | using f by induct auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 794 | |
| 38656 | 795 | lemma (in measure_space) restricted_measure_space: | 
| 796 | assumes "S \<in> sets M" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 797 | shows "measure_space (restricted_space S)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 798 | (is "measure_space ?r") | 
| 38656 | 799 | unfolding measure_space_def measure_space_axioms_def | 
| 800 | proof safe | |
| 801 | show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] . | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 802 | show "positive ?r (measure ?r)" using `S \<in> sets M` by (auto simp: positive_def) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 803 | |
| 
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 | show "countably_additive ?r (measure ?r)" | 
| 38656 | 805 | unfolding countably_additive_def | 
| 806 | proof safe | |
| 807 | fix A :: "nat \<Rightarrow> 'a set" | |
| 808 | assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A" | |
| 809 | from restriction_in_sets[OF assms *[simplified]] ** | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 810 | show "(\<Sum>n. measure ?r (A n)) = measure ?r (\<Union>i. A i)" | 
| 38656 | 811 | using measure_countably_additive by simp | 
| 812 | qed | |
| 813 | qed | |
| 814 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 815 | lemma (in measure_space) AE_restricted: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 816 | assumes "A \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 817 | shows "(AE x in restricted_space A. P x) \<longleftrightarrow> (AE x. x \<in> A \<longrightarrow> P x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 818 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 819 | interpret R: measure_space "restricted_space A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 820 | by (rule restricted_measure_space[OF `A \<in> sets M`]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 821 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 822 | proof | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 823 | assume "AE x in restricted_space A. P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 824 | from this[THEN R.AE_E] guess N' . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 825 |     then obtain N where "{x \<in> A. \<not> P x} \<subseteq> A \<inter> N" "\<mu> (A \<inter> N) = 0" "N \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 826 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 827 |     moreover then have "{x \<in> space M. \<not> (x \<in> A \<longrightarrow> P x)} \<subseteq> A \<inter> N"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 828 | using `A \<in> sets M` sets_into_space by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 829 | ultimately show "AE x. x \<in> A \<longrightarrow> P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 830 | using `A \<in> sets M` by (auto intro!: AE_I[where N="A \<inter> N"]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 831 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 832 | assume "AE x. x \<in> A \<longrightarrow> P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 833 | from this[THEN AE_E] guess N . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 834 | then show "AE x in restricted_space A. P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 835 | using null_set_Int1[OF _ `A \<in> sets M`] `A \<in> sets M`[THEN sets_into_space] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 836 | by (auto intro!: R.AE_I[where N="A \<inter> N"] simp: subset_eq) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 837 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 838 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 839 | |
| 39092 | 840 | lemma (in measure_space) measure_space_subalgebra: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 841 | assumes "sigma_algebra N" and "sets N \<subseteq> sets M" "space N = space M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 842 | and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 843 | shows "measure_space N" | 
| 39092 | 844 | proof - | 
| 41545 | 845 | interpret N: sigma_algebra N by fact | 
| 39092 | 846 | show ?thesis | 
| 847 | proof | |
| 41545 | 848 | from `sets N \<subseteq> sets M` have "\<And>A. range A \<subseteq> sets N \<Longrightarrow> range A \<subseteq> sets M" by blast | 
| 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 | 849 | then show "countably_additive N (measure N)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 850 | by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 851 | show "positive N (measure_space.measure N)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 852 | using assms(2) by (auto simp add: positive_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 853 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 854 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 855 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 856 | lemma (in measure_space) AE_subalgebra: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 857 | assumes ae: "AE x in N. P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 858 | and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 859 | and sa: "sigma_algebra N" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 860 | shows "AE x. P x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 861 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 862 | interpret N: measure_space N using measure_space_subalgebra[OF sa N] . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 863 | from ae[THEN N.AE_E] guess N . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 864 | with N show ?thesis unfolding almost_everywhere_def by auto | 
| 39092 | 865 | qed | 
| 866 | ||
| 38656 | 867 | section "@{text \<sigma>}-finite Measures"
 | 
| 868 | ||
| 869 | locale sigma_finite_measure = measure_space + | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 870 | assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)" | 
| 38656 | 871 | |
| 872 | lemma (in sigma_finite_measure) restricted_sigma_finite_measure: | |
| 873 | assumes "S \<in> sets M" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 874 | shows "sigma_finite_measure (restricted_space S)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 875 | (is "sigma_finite_measure ?r") | 
| 38656 | 876 | unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def | 
| 877 | proof safe | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 878 | show "measure_space ?r" using restricted_measure_space[OF assms] . | 
| 38656 | 879 | next | 
| 880 | obtain A :: "nat \<Rightarrow> 'a set" where | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 881 | "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 38656 | 882 | using sigma_finite by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 883 | show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<infinity>)" | 
| 38656 | 884 | proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI) | 
| 885 | fix i | |
| 886 | show "A i \<inter> S \<in> sets ?r" | |
| 887 | using `range A \<subseteq> sets M` `S \<in> sets M` by auto | |
| 888 | next | |
| 889 | fix x i assume "x \<in> S" thus "x \<in> space ?r" by simp | |
| 890 | next | |
| 891 | fix x assume "x \<in> space ?r" thus "x \<in> (\<Union>i. A i \<inter> S)" | |
| 892 | using `(\<Union>i. A i) = space M` `S \<in> sets M` by auto | |
| 893 | next | |
| 894 | fix i | |
| 895 | have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)" | |
| 896 | using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 897 | then show "measure ?r (A i \<inter> S) \<noteq> \<infinity>" using `\<mu> (A i) \<noteq> \<infinity>` by auto | 
| 38656 | 898 | qed | 
| 899 | qed | |
| 900 | ||
| 40859 | 901 | lemma (in sigma_finite_measure) sigma_finite_measure_cong: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 902 | assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 903 | shows "sigma_finite_measure M'" | 
| 40859 | 904 | 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 | 905 | interpret M': measure_space M' by (intro measure_space_cong cong) | 
| 40859 | 906 | from sigma_finite guess A .. note A = this | 
| 907 | then have "\<And>i. A i \<in> sets M" by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 908 | with A have fin: "\<forall>i. measure M' (A i) \<noteq> \<infinity>" using cong by auto | 
| 40859 | 909 | show ?thesis | 
| 910 | apply default | |
| 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 | 911 | using A fin cong by auto | 
| 40859 | 912 | qed | 
| 913 | ||
| 39092 | 914 | lemma (in sigma_finite_measure) disjoint_sigma_finite: | 
| 915 | "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 916 | (\<forall>i. \<mu> (A i) \<noteq> \<infinity>) \<and> disjoint_family A" | 
| 39092 | 917 | proof - | 
| 918 | obtain A :: "nat \<Rightarrow> 'a set" where | |
| 919 | range: "range A \<subseteq> sets M" and | |
| 920 | space: "(\<Union>i. A i) = space M" and | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 921 | measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 39092 | 922 | using sigma_finite by auto | 
| 923 | note range' = range_disjointed_sets[OF range] range | |
| 924 |   { fix i
 | |
| 925 | have "\<mu> (disjointed A i) \<le> \<mu> (A i)" | |
| 926 | using range' disjointed_subset[of A i] by (auto intro!: measure_mono) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 927 | then have "\<mu> (disjointed A i) \<noteq> \<infinity>" | 
| 39092 | 928 | using measure[of i] by auto } | 
| 929 | with disjoint_family_disjointed UN_disjointed_eq[of A] space range' | |
| 930 | show ?thesis by (auto intro!: exI[of _ "disjointed A"]) | |
| 931 | qed | |
| 932 | ||
| 40859 | 933 | lemma (in sigma_finite_measure) sigma_finite_up: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 934 | "\<exists>F. range F \<subseteq> sets M \<and> incseq F \<and> (\<Union>i. F i) = space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<infinity>)" | 
| 40859 | 935 | proof - | 
| 936 | obtain F :: "nat \<Rightarrow> 'a set" where | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 937 | F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<infinity>" | 
| 40859 | 938 | using sigma_finite by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 939 | then show ?thesis | 
| 40859 | 940 | proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI) | 
| 941 | from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto | |
| 942 | then show "(\<Union>n. \<Union> i\<le>n. F i) = space M" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
43920diff
changeset | 943 | using F by fastforce | 
| 40859 | 944 | next | 
| 945 | fix n | |
| 946 | have "\<mu> (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. \<mu> (F i))" using F | |
| 947 | by (auto intro!: measure_finitely_subadditive) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 948 | also have "\<dots> < \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 949 | using F by (auto simp: setsum_Pinfty) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 950 | finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 951 | qed (force simp: incseq_def)+ | 
| 40859 | 952 | qed | 
| 953 | ||
| 41831 | 954 | section {* Measure preserving *}
 | 
| 955 | ||
| 956 | definition "measure_preserving A B = | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 957 |     {f \<in> measurable A B. (\<forall>y \<in> sets B. measure B y = measure A (f -` y \<inter> space A))}"
 | 
| 41831 | 958 | |
| 959 | lemma measure_preservingI[intro?]: | |
| 960 | assumes "f \<in> measurable A B" | |
| 961 | and "\<And>y. y \<in> sets B \<Longrightarrow> measure A (f -` y \<inter> space A) = measure B y" | |
| 962 | shows "f \<in> measure_preserving A B" | |
| 963 | unfolding measure_preserving_def using assms by auto | |
| 964 | ||
| 965 | lemma (in measure_space) measure_space_vimage: | |
| 966 |   fixes M' :: "('c, 'd) measure_space_scheme"
 | |
| 967 | assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" | |
| 968 | shows "measure_space M'" | |
| 969 | proof - | |
| 970 | interpret M': sigma_algebra M' by fact | |
| 971 | show ?thesis | |
| 972 | proof | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 973 | show "positive M' (measure M')" using T | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 974 | by (auto simp: measure_preserving_def positive_def measurable_sets) | 
| 41831 | 975 | |
| 976 | show "countably_additive M' (measure M')" | |
| 977 | proof (intro countably_additiveI) | |
| 978 | fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A" | |
| 979 | then have A: "\<And>i. A i \<in> sets M'" "(\<Union>i. A i) \<in> sets M'" by auto | |
| 980 | then have *: "range (\<lambda>i. T -` (A i) \<inter> space M) \<subseteq> sets M" | |
| 981 | using T by (auto simp: measurable_def measure_preserving_def) | |
| 982 | moreover have "(\<Union>i. T -` A i \<inter> space M) \<in> sets M" | |
| 983 | using * by blast | |
| 984 | moreover have **: "disjoint_family (\<lambda>i. T -` A i \<inter> space M)" | |
| 985 | using `disjoint_family A` by (auto simp: disjoint_family_on_def) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 986 | ultimately show "(\<Sum>i. measure M' (A i)) = measure M' (\<Union>i. A i)" | 
| 41831 | 987 | using measure_countably_additive[OF _ **] A T | 
| 988 | by (auto simp: comp_def vimage_UN measure_preserving_def) | |
| 989 | qed | |
| 990 | qed | |
| 991 | qed | |
| 992 | ||
| 993 | lemma (in measure_space) almost_everywhere_vimage: | |
| 994 | assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'" | |
| 995 | and AE: "measure_space.almost_everywhere M' P" | |
| 996 | shows "AE x. P (T x)" | |
| 997 | proof - | |
| 998 | interpret M': measure_space M' using T by (rule measure_space_vimage) | |
| 999 | from AE[THEN M'.AE_E] guess N . | |
| 1000 | then show ?thesis | |
| 1001 | unfolding almost_everywhere_def M'.almost_everywhere_def | |
| 1002 | using T(2) unfolding measurable_def measure_preserving_def | |
| 1003 | by (intro bexI[of _ "T -` N \<inter> space M"]) auto | |
| 1004 | qed | |
| 1005 | ||
| 1006 | lemma measure_unique_Int_stable_vimage: | |
| 1007 | fixes A :: "nat \<Rightarrow> 'a set" | |
| 1008 | assumes E: "Int_stable E" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1009 | and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure M (A i) \<noteq> \<infinity>" | 
| 41831 | 1010 | and N: "measure_space N" "T \<in> measurable N M" | 
| 1011 | and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M" | |
| 1012 | and eq: "\<And>X. X \<in> sets E \<Longrightarrow> measure M X = measure N (T -` X \<inter> space N)" | |
| 1013 | assumes X: "X \<in> sets (sigma E)" | |
| 1014 | shows "measure M X = measure N (T -` X \<inter> space N)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1015 | proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X]) | 
| 41831 | 1016 | interpret M: measure_space M by fact | 
| 1017 | interpret N: measure_space N by fact | |
| 1018 | let "?T X" = "T -` X \<inter> space N" | |
| 1019 | show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = measure M\<rparr>" | |
| 1020 | by (rule M.measure_space_cong) (auto simp: M) | |
| 1021 | show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<lambda>X. measure N (?T X)\<rparr>" (is "measure_space ?E") | |
| 1022 | proof (rule N.measure_space_vimage) | |
| 1023 | show "sigma_algebra ?E" | |
| 1024 | by (rule M.sigma_algebra_cong) (auto simp: M) | |
| 1025 | show "T \<in> measure_preserving N ?E" | |
| 1026 | using `T \<in> measurable N M` by (auto simp: M measurable_def measure_preserving_def) | |
| 1027 | qed | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1028 | show "\<And>i. M.\<mu> (A i) \<noteq> \<infinity>" by fact | 
| 41831 | 1029 | qed | 
| 1030 | ||
| 1031 | lemma (in measure_space) measure_preserving_Int_stable: | |
| 1032 | fixes A :: "nat \<Rightarrow> 'a set" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1033 | assumes E: "Int_stable E" "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure E (A i) \<noteq> \<infinity>" | 
| 41831 | 1034 | and N: "measure_space (sigma E)" | 
| 1035 | and T: "T \<in> measure_preserving M E" | |
| 1036 | shows "T \<in> measure_preserving M (sigma E)" | |
| 1037 | proof | |
| 1038 | interpret E: measure_space "sigma E" by fact | |
| 1039 | show "T \<in> measurable M (sigma E)" | |
| 1040 | using T E.sets_into_space | |
| 1041 | by (intro measurable_sigma) (auto simp: measure_preserving_def measurable_def) | |
| 1042 | fix X assume "X \<in> sets (sigma E)" | |
| 1043 | show "\<mu> (T -` X \<inter> space M) = E.\<mu> X" | |
| 1044 | proof (rule measure_unique_Int_stable_vimage[symmetric]) | |
| 1045 | show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1046 | "\<And>i. E.\<mu> (A i) \<noteq> \<infinity>" using E by auto | 
| 41831 | 1047 | show "measure_space M" by default | 
| 1048 | next | |
| 1049 | fix X assume "X \<in> sets E" then show "E.\<mu> X = \<mu> (T -` X \<inter> space M)" | |
| 1050 | using T unfolding measure_preserving_def by auto | |
| 1051 | qed fact+ | |
| 1052 | qed | |
| 1053 | ||
| 38656 | 1054 | section "Real measure values" | 
| 1055 | ||
| 1056 | lemma (in measure_space) real_measure_Union: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1057 | assumes finite: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>" | 
| 38656 | 1058 |   and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
 | 
| 1059 | shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)" | |
| 1060 | unfolding measure_additive[symmetric, OF measurable] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1061 | using measurable(1,2)[THEN positive_measure] | 
| 43920 | 1062 | using finite by (cases rule: ereal2_cases[of "\<mu> A" "\<mu> B"]) auto | 
| 38656 | 1063 | |
| 1064 | lemma (in measure_space) real_measure_finite_Union: | |
| 1065 | assumes measurable: | |
| 1066 | "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1067 | assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>" | 
| 38656 | 1068 | shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1069 | using finite measurable(2)[THEN positive_measure] | 
| 43920 | 1070 | by (force intro!: setsum_real_of_ereal[symmetric] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1071 | simp: measure_setsum[OF measurable, symmetric]) | 
| 38656 | 1072 | |
| 1073 | lemma (in measure_space) real_measure_Diff: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1074 | assumes finite: "\<mu> A \<noteq> \<infinity>" | 
| 38656 | 1075 | and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A" | 
| 1076 | shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)" | |
| 1077 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1078 | have "\<mu> (A - B) \<le> \<mu> A" "\<mu> B \<le> \<mu> A" | 
| 38656 | 1079 | using measurable by (auto intro!: measure_mono) | 
| 1080 | hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)" | |
| 1081 | using measurable finite by (rule_tac real_measure_Union) auto | |
| 1082 | thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2) | |
| 1083 | qed | |
| 1084 | ||
| 1085 | lemma (in measure_space) real_measure_UNION: | |
| 1086 | assumes measurable: "range A \<subseteq> sets M" "disjoint_family A" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1087 | assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>" | 
| 38656 | 1088 | shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))" | 
| 1089 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1090 | have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto | 
| 43920 | 1091 | with summable_sums[OF summable_ereal_pos, of "\<lambda>i. \<mu> (A i)"] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1092 | measure_countably_additive[OF measurable] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1093 | have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1094 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1095 |   { fix i
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1096 | have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1097 | using measurable by (auto intro!: measure_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1098 | moreover have "0 \<le> \<mu> (A i)" using measurable by auto | 
| 43920 | 1099 | ultimately have "\<mu> (A i) = ereal (real (\<mu> (A i)))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1100 | using finite by (cases "\<mu> (A i)") auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1101 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1102 | have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto | 
| 43920 | 1103 | then have "\<mu> (\<Union>i. A i) = ereal (real (\<mu> (\<Union>i. A i)))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1104 | using finite by (cases "\<mu> (\<Union>i. A i)") auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1105 | ultimately show ?thesis | 
| 43920 | 1106 | unfolding sums_ereal[symmetric] by simp | 
| 38656 | 1107 | qed | 
| 1108 | ||
| 1109 | lemma (in measure_space) real_measure_subadditive: | |
| 1110 | assumes measurable: "A \<in> sets M" "B \<in> sets M" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1111 | and fin: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>" | 
| 38656 | 1112 | shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)" | 
| 1113 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1114 | have "0 \<le> \<mu> (A \<union> B)" using measurable by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1115 | then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1116 | using measure_subadditive[OF measurable] fin | 
| 43920 | 1117 | by (cases rule: ereal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto | 
| 38656 | 1118 | qed | 
| 1119 | ||
| 1120 | lemma (in measure_space) real_measure_setsum_singleton: | |
| 1121 |   assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1122 |   and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
 | 
| 38656 | 1123 |   shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1124 | using measure_finite_singleton[OF S] fin | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1125 | using positive_measure[OF S(2)] | 
| 43920 | 1126 | by (force intro!: setsum_real_of_ereal[symmetric]) | 
| 38656 | 1127 | |
| 1128 | lemma (in measure_space) real_continuity_from_below: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1129 | assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>" | 
| 38656 | 1130 | shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1131 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1132 | have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto | 
| 43920 | 1133 | then have "ereal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)" | 
| 1134 | using fin by (auto intro: ereal_real') | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1135 | then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1136 | using continuity_from_below_Lim[OF A] | 
| 43920 | 1137 | by (intro lim_real_of_ereal) simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1138 | qed | 
| 38656 | 1139 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1140 | lemma (in measure_space) continuity_from_above_Lim: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1141 | assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1142 | shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)" | 
| 43920 | 1143 | using LIMSEQ_ereal_INFI[OF measure_decseq, OF A] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1144 | using continuity_from_above[OF A fin] by simp | 
| 38656 | 1145 | |
| 1146 | lemma (in measure_space) real_continuity_from_above: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1147 | assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" | 
| 38656 | 1148 | shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1149 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1150 | have "0 \<le> \<mu> (\<Inter>i. A i)" using A by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1151 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1152 | have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1153 | using A by (auto intro!: measure_mono) | 
| 43920 | 1154 | ultimately have "ereal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)" | 
| 1155 | using fin by (auto intro: ereal_real') | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1156 | then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1157 | using continuity_from_above_Lim[OF A fin] | 
| 43920 | 1158 | by (intro lim_real_of_ereal) simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1159 | qed | 
| 38656 | 1160 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1161 | lemma (in measure_space) real_measure_countably_subadditive: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1162 | assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. \<mu> (A i)) \<noteq> \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1163 | shows "real (\<mu> (\<Union>i. A i)) \<le> (\<Sum>i. real (\<mu> (A i)))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1164 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1165 |   { fix i
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1166 | have "0 \<le> \<mu> (A i)" using A by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1167 | moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1168 | ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1169 | moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto | 
| 43920 | 1170 | ultimately have "ereal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. ereal (real (\<mu> (A i))))" | 
| 1171 | using measure_countably_subadditive[OF A] by (auto simp: ereal_real) | |
| 1172 | also have "\<dots> = ereal (\<Sum>i. real (\<mu> (A i)))" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1173 | using A | 
| 43920 | 1174 | by (auto intro!: sums_unique[symmetric] sums_ereal[THEN iffD2] summable_sums summable_real_of_ereal fin) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1175 | finally show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1176 | qed | 
| 38656 | 1177 | |
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1178 | locale finite_measure = sigma_finite_measure + | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1179 | assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<infinity>" | 
| 38656 | 1180 | |
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1181 | lemma finite_measureI[Pure.intro!]: | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1182 | assumes "measure_space M" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1183 | assumes *: "measure M (space M) \<noteq> \<infinity>" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1184 | shows "finite_measure M" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1185 | proof - | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1186 | interpret measure_space M by fact | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1187 | show "finite_measure M" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1188 | proof | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1189 | show "measure M (space M) \<noteq> \<infinity>" by fact | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1190 | show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1191 | using * by (auto intro!: exI[of _ "\<lambda>x. space M"]) | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1192 | qed | 
| 38656 | 1193 | qed | 
| 1194 | ||
| 39092 | 1195 | lemma (in finite_measure) finite_measure[simp, intro]: | 
| 38656 | 1196 | assumes "A \<in> sets M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1197 | shows "\<mu> A \<noteq> \<infinity>" | 
| 38656 | 1198 | proof - | 
| 1199 | from `A \<in> sets M` have "A \<subseteq> space M" | |
| 1200 | using sets_into_space by blast | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1201 | then have "\<mu> A \<le> \<mu> (space M)" | 
| 38656 | 1202 | using assms top by (rule measure_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1203 | then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1204 | using finite_measure_of_space by auto | 
| 38656 | 1205 | qed | 
| 1206 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1207 | definition (in finite_measure) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1208 | "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1209 | |
| 43920 | 1210 | lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = ereal (\<mu>' A)" | 
| 1211 | by (auto simp: \<mu>'_def ereal_real) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1212 | |
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1213 | lemma (in finite_measure) positive_measure'[simp, intro]: "0 \<le> \<mu>' A" | 
| 43920 | 1214 | unfolding \<mu>'_def by (auto simp: real_of_ereal_pos) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1215 | |
| 43556 
0d78c8d31d0d
move conditional expectation to its own theory file
 hoelzl parents: 
42991diff
changeset | 1216 | lemma (in finite_measure) real_measure: | 
| 43920 | 1217 | assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = ereal r" | 
| 43556 
0d78c8d31d0d
move conditional expectation to its own theory file
 hoelzl parents: 
42991diff
changeset | 1218 | using finite_measure[OF A] positive_measure[OF A] by (cases "\<mu> A") auto | 
| 
0d78c8d31d0d
move conditional expectation to its own theory file
 hoelzl parents: 
42991diff
changeset | 1219 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1220 | lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1221 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1222 | assume "A \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1223 | moreover then have "\<mu> A \<le> \<mu> (space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1224 | using sets_into_space by (auto intro!: measure_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1225 | ultimately show ?thesis | 
| 43920 | 1226 | by (auto simp: \<mu>'_def intro!: real_of_ereal_positive_mono) | 
| 1227 | qed (simp add: \<mu>'_def real_of_ereal_pos) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1228 | |
| 38656 | 1229 | lemma (in finite_measure) restricted_finite_measure: | 
| 1230 | assumes "S \<in> sets M" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1231 | shows "finite_measure (restricted_space S)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1232 | (is "finite_measure ?r") | 
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1233 | 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 | 1234 | show "measure_space ?r" using restricted_measure_space[OF assms] . | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1235 | show "measure ?r (space ?r) \<noteq> \<infinity>" using finite_measure[OF `S \<in> sets M`] by auto | 
| 38656 | 1236 | qed | 
| 1237 | ||
| 40859 | 1238 | lemma (in measure_space) restricted_to_finite_measure: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1239 | assumes "S \<in> sets M" "\<mu> S \<noteq> \<infinity>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 1240 | shows "finite_measure (restricted_space S)" | 
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1241 | proof | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1242 | show "measure_space (restricted_space S)" | 
| 40859 | 1243 | using `S \<in> sets M` by (rule restricted_measure_space) | 
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1244 | show "measure (restricted_space S) (space (restricted_space S)) \<noteq> \<infinity>" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1245 | by simp fact | 
| 40859 | 1246 | qed | 
| 1247 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1248 | lemma (in finite_measure) finite_measure_Diff: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1249 | assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1250 | shows "\<mu>' (A - B) = \<mu>' A - \<mu>' B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1251 | using sets[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1252 | using Diff[OF sets, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1253 | using measure_Diff[OF _ assms] by simp | 
| 38656 | 1254 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1255 | lemma (in finite_measure) finite_measure_Union: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1256 |   assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1257 | shows "\<mu>' (A \<union> B) = \<mu>' A + \<mu>' B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1258 | using measure_additive[OF assms] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1259 | using sets[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1260 | using Un[OF sets, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1261 | by simp | 
| 38656 | 1262 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1263 | lemma (in finite_measure) finite_measure_finite_Union: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1264 | assumes S: "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1265 | and dis: "disjoint_family_on A S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1266 | shows "\<mu>' (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. \<mu>' (A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1267 | using measure_setsum[OF assms] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1268 | using finite_UN[of S A, OF S, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1269 | using S(2)[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1270 | by simp | 
| 38656 | 1271 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1272 | lemma (in finite_measure) finite_measure_UNION: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1273 | assumes A: "range A \<subseteq> sets M" "disjoint_family A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1274 | shows "(\<lambda>i. \<mu>' (A i)) sums (\<mu>' (\<Union>i. A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1275 | using real_measure_UNION[OF A] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1276 | using countable_UN[OF A(1), THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1277 | using A(1)[THEN subsetD, THEN finite_measure_eq] | 
| 39092 | 1278 | by auto | 
| 1279 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1280 | lemma (in finite_measure) finite_measure_mono: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1281 | assumes B: "B \<in> sets M" and "A \<subseteq> B" shows "\<mu>' A \<le> \<mu>' B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1282 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1283 | assume "A \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1284 | from this[THEN finite_measure_eq] B[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1285 | show ?thesis using measure_mono[OF `A \<subseteq> B` `A \<in> sets M` `B \<in> sets M`] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1286 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1287 | assume "A \<notin> sets M" then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1288 | using positive_measure'[of B] unfolding \<mu>'_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1289 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1290 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1291 | lemma (in finite_measure) finite_measure_subadditive: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1292 | assumes m: "A \<in> sets M" "B \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1293 | shows "\<mu>' (A \<union> B) \<le> \<mu>' A + \<mu>' B" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1294 | using measure_subadditive[OF m] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1295 | using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1296 | |
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1297 | lemma (in finite_measure) finite_measure_subadditive_finite: | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1298 | assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1299 | shows "\<mu>' (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu>' (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1300 | using measure_subadditive_finite[OF assms] assms | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1301 | by (simp add: finite_measure_eq finite_UN) | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1302 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1303 | lemma (in finite_measure) finite_measure_countably_subadditive: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1304 | assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. \<mu>' (A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1305 | shows "\<mu>' (\<Union>i. A i) \<le> (\<Sum>i. \<mu>' (A i))" | 
| 38656 | 1306 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1307 | note A[THEN subsetD, THEN finite_measure_eq, simp] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1308 | note countable_UN[OF A, THEN finite_measure_eq, simp] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1309 | from `summable (\<lambda>i. \<mu>' (A i))` | 
| 43920 | 1310 | have "(\<lambda>i. ereal (\<mu>' (A i))) sums ereal (\<Sum>i. \<mu>' (A i))" | 
| 1311 | by (simp add: sums_ereal) (rule summable_sums) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1312 | from sums_unique[OF this, symmetric] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1313 | measure_countably_subadditive[OF A] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1314 | show ?thesis by simp | 
| 38656 | 1315 | qed | 
| 1316 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1317 | lemma (in finite_measure) finite_measure_finite_singleton: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1318 |   assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1319 |   shows "\<mu>' S = (\<Sum>x\<in>S. \<mu>' {x})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1320 | using real_measure_setsum_singleton[OF assms] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1321 | using *[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1322 |   using finite_UN[of S "\<lambda>x. {x}", OF assms, THEN finite_measure_eq]
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1323 | by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1324 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1325 | lemma (in finite_measure) finite_continuity_from_below: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1326 | assumes A: "range A \<subseteq> sets M" and "incseq A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1327 | shows "(\<lambda>i. \<mu>' (A i)) ----> \<mu>' (\<Union>i. A i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1328 | using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1329 | using A[THEN subsetD, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1330 | using countable_UN[OF A, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1331 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1332 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1333 | lemma (in finite_measure) finite_continuity_from_above: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1334 | assumes A: "range A \<subseteq> sets M" and "decseq A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1335 | shows "(\<lambda>n. \<mu>' (A n)) ----> \<mu>' (\<Inter>i. A i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1336 | using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1337 | using A[THEN subsetD, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1338 | using countable_INT[OF A, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1339 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1340 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1341 | lemma (in finite_measure) finite_measure_compl: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1342 | assumes S: "S \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1343 | shows "\<mu>' (space M - S) = \<mu>' (space M) - \<mu>' S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1344 | using measure_compl[OF S, OF finite_measure, OF S] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1345 | using S[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1346 | using compl_sets[OF S, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1347 | using top[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1348 | by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1349 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1350 | lemma (in finite_measure) finite_measure_inter_full_set: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1351 | assumes S: "S \<in> sets M" "T \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1352 | assumes T: "\<mu>' T = \<mu>' (space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1353 | shows "\<mu>' (S \<inter> T) = \<mu>' S" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1354 | using measure_inter_full_set[OF S finite_measure] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1355 | using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1356 | using Int[OF S, THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1357 | using S[THEN finite_measure_eq] top[THEN finite_measure_eq] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1358 | by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1359 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1360 | lemma (in finite_measure) empty_measure'[simp]: "\<mu>' {} = 0"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1361 | unfolding \<mu>'_def by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1362 | |
| 38656 | 1363 | section "Finite spaces" | 
| 1364 | ||
| 45777 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1365 | locale finite_measure_space = finite_measure + finite_sigma_algebra | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1366 | |
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1367 | lemma finite_measure_spaceI[Pure.intro!]: | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1368 | assumes "finite (space M)" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1369 | assumes sets_Pow: "sets M = Pow (space M)" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1370 | and space: "measure M (space M) \<noteq> \<infinity>" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1371 |     and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> measure M {x}"
 | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1372 |     and add: "\<And>A. A \<subseteq> space M \<Longrightarrow> measure M A = (\<Sum>x\<in>A. measure M {x})"
 | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1373 | shows "finite_measure_space M" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1374 | proof - | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1375 | interpret finite_sigma_algebra M | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1376 | proof | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1377 | show "finite (space M)" by fact | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1378 | qed (auto simp: sets_Pow) | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1379 | interpret measure_space M | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1380 | proof (rule finite_additivity_sufficient) | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1381 | show "sigma_algebra M" by default | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1382 | show "finite (space M)" by fact | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1383 | show "positive M (measure M)" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1384 | by (auto simp: add positive_def intro!: setsum_nonneg pos) | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1385 | show "additive M (measure M)" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1386 | using `finite (space M)` | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1387 | by (auto simp add: additive_def add | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1388 | intro!: setsum_Un_disjoint dest: finite_subset) | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1389 | qed | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1390 | interpret finite_measure M | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1391 | proof | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1392 | show "\<mu> (space M) \<noteq> \<infinity>" by fact | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1393 | qed default | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1394 | show "finite_measure_space M" | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1395 | by default | 
| 
c36637603821
remove unnecessary sublocale instantiations in HOL-Probability (for clarity and speedup); remove Infinite_Product_Measure.product_prob_space which was a duplicate of Probability_Measure.product_prob_space
 hoelzl parents: 
45342diff
changeset | 1396 | qed | 
| 39092 | 1397 | |
| 38656 | 1398 | lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1399 |   using measure_setsum[of "space M" "\<lambda>i. {i}"]
 | 
| 36624 | 1400 | by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) | 
| 1401 | ||
| 42892 | 1402 | lemma (in finite_measure_space) finite_measure_singleton: | 
| 1403 |   assumes A: "A \<subseteq> space M" shows "\<mu>' A = (\<Sum>x\<in>A. \<mu>' {x})"
 | |
| 1404 | using A finite_subset[OF A finite_space] | |
| 1405 | by (intro finite_measure_finite_singleton) auto | |
| 1406 | ||
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1407 | lemma (in finite_measure_space) finite_measure_subadditive_setsum: | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1408 | assumes "finite I" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1409 | shows "\<mu>' (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu>' (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1410 | proof cases | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1411 | assume "(\<Union>i\<in>I. A i) \<subseteq> space M" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1412 | then have "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M" by auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1413 | from finite_measure_subadditive_finite[OF `finite I` this] | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1414 | show ?thesis by auto | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1415 | next | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1416 | assume "\<not> (\<Union>i\<in>I. A i) \<subseteq> space M" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1417 | then have "\<mu>' (\<Union>i\<in>I. A i) = 0" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1418 | by (simp add: \<mu>'_def) | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1419 | also have "0 \<le> (\<Sum>i\<in>I. \<mu>' (A i))" | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1420 | by (auto intro!: setsum_nonneg) | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1421 | finally show ?thesis . | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1422 | qed | 
| 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1423 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1424 | lemma suminf_cmult_indicator: | 
| 43920 | 1425 | fixes f :: "nat \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1426 | assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1427 | shows "(\<Sum>n. f n * indicator (A n) x) = f i" | 
| 38656 | 1428 | proof - | 
| 43920 | 1429 | have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)" | 
| 38656 | 1430 | using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto | 
| 43920 | 1431 | then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)" | 
| 38656 | 1432 | by (auto simp: setsum_cases) | 
| 43920 | 1433 | moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)" | 
| 1434 | proof (rule ereal_SUPI) | |
| 1435 | fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y" | |
| 38656 | 1436 | from this[of "Suc i"] show "f i \<le> y" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1437 | qed (insert assms, simp) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1438 | ultimately show ?thesis using assms | 
| 43920 | 1439 | by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def) | 
| 38656 | 1440 | qed | 
| 1441 | ||
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1442 | lemma suminf_indicator: | 
| 38656 | 1443 | assumes "disjoint_family A" | 
| 43920 | 1444 | shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x" | 
| 38656 | 1445 | proof cases | 
| 1446 | assume *: "x \<in> (\<Union>i. A i)" | |
| 1447 | then obtain i where "x \<in> A i" by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1448 | from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"] | 
| 38656 | 1449 | show ?thesis using * by simp | 
| 1450 | qed simp | |
| 1451 | ||
| 35582 | 1452 | end |