| author | blanchet | 
| Thu, 20 Sep 2012 02:42:48 +0200 | |
| changeset 49454 | cca4390e8071 | 
| parent 49394 | 52e636ace94e | 
| child 49773 | 16907431e477 | 
| permissions | -rw-r--r-- | 
| 42067 | 1 | (* Title: HOL/Probability/Caratheodory.thy | 
| 2 | Author: Lawrence C Paulson | |
| 3 | Author: Johannes Hölzl, TU München | |
| 4 | *) | |
| 5 | ||
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 6 | header {*Caratheodory Extension Theorem*}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 7 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 8 | theory Caratheodory | 
| 47694 | 9 | imports Measure_Space | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 10 | begin | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 11 | |
| 42950 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 12 | lemma sums_def2: | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 13 | "f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x" | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 14 | unfolding sums_def | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 15 | apply (subst LIMSEQ_Suc_iff[symmetric]) | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 16 | unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .. | 
| 
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
 hoelzl parents: 
42145diff
changeset | 17 | |
| 42067 | 18 | text {*
 | 
| 19 | Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson. | |
| 20 | *} | |
| 21 | ||
| 43920 | 22 | lemma suminf_ereal_2dimen: | 
| 23 | fixes f:: "nat \<times> nat \<Rightarrow> ereal" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 24 | assumes pos: "\<And>p. 0 \<le> f p" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 25 | assumes "\<And>m. g m = (\<Sum>n. f (m,n))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 26 | shows "(\<Sum>i. f (prod_decode i)) = suminf g" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 27 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 28 | have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 29 | using assms by (simp add: fun_eq_iff) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 30 | have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 31 | by (simp add: setsum_reindex[OF inj_prod_decode] comp_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 32 |   { fix n
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 33 |     let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 34 |     { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 35 | then have "a < ?M fst" "b < ?M snd" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 36 | by (auto intro!: Max_ge le_imp_less_Suc image_eqI) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 37 |     then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 38 | by (auto intro!: setsum_mono3 simp: pos) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 39 |     then have "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto }
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 40 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 41 |   { fix a b
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 42 |     let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 43 |     { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 44 | by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 45 |     then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 46 | by (auto intro!: setsum_mono3 simp: pos) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 47 | ultimately | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 48 | show ?thesis unfolding g_def using pos | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44918diff
changeset | 49 | by (auto intro!: SUPR_eq simp: setsum_cartesian_product reindex SUP_upper2 | 
| 43920 | 50 | setsum_nonneg suminf_ereal_eq_SUPR SUPR_pair | 
| 51 | SUPR_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 52 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 53 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 54 | subsection {* Measure Spaces *}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 55 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 56 | definition subadditive where "subadditive M f \<longleftrightarrow> | 
| 47694 | 57 |   (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 58 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 59 | definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow> | 
| 47694 | 60 | (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 61 | (f (\<Union>i. A i) \<le> (\<Sum>i. f (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: 
41023diff
changeset | 62 | |
| 47694 | 63 | definition lambda_system where "lambda_system \<Omega> M f = {l \<in> M.
 | 
| 64 | \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 65 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 66 | definition outer_measure_space where "outer_measure_space M f \<longleftrightarrow> | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 67 | positive M f \<and> increasing M f \<and> countably_subadditive M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 68 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 69 | definition measure_set where "measure_set M f X = {r.
 | 
| 47694 | 70 | \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 71 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 72 | lemma subadditiveD: | 
| 47694 | 73 |   "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 74 | by (auto simp add: subadditive_def) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 75 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 76 | subsection {* Lambda Systems *}
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 77 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 78 | lemma (in algebra) lambda_system_eq: | 
| 47694 | 79 |   shows "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 80 | proof - | 
| 47694 | 81 | have [simp]: "!!l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l" | 
| 37032 | 82 | by (metis Int_Diff Int_absorb1 Int_commute sets_into_space) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 83 | show ?thesis | 
| 37032 | 84 | by (auto simp add: lambda_system_def) (metis Int_commute)+ | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 85 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 86 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 87 | lemma (in algebra) lambda_system_empty: | 
| 47694 | 88 |   "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
 | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 89 | by (auto simp add: positive_def lambda_system_eq) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 90 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 91 | lemma lambda_system_sets: | 
| 47694 | 92 | "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 93 | by (simp add: lambda_system_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 94 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 95 | lemma (in algebra) lambda_system_Compl: | 
| 43920 | 96 | fixes f:: "'a set \<Rightarrow> ereal" | 
| 47694 | 97 | assumes x: "x \<in> lambda_system \<Omega> M f" | 
| 98 | shows "\<Omega> - x \<in> lambda_system \<Omega> M f" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 99 | proof - | 
| 47694 | 100 | have "x \<subseteq> \<Omega>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 101 | by (metis sets_into_space lambda_system_sets x) | 
| 47694 | 102 | hence "\<Omega> - (\<Omega> - x) = x" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 103 | by (metis double_diff equalityE) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 104 | with x show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 105 | by (force simp add: lambda_system_def ac_simps) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 106 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 107 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 108 | lemma (in algebra) lambda_system_Int: | 
| 43920 | 109 | fixes f:: "'a set \<Rightarrow> ereal" | 
| 47694 | 110 | assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" | 
| 111 | shows "x \<inter> y \<in> lambda_system \<Omega> M f" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 112 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 113 | from xl yl show ?thesis | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 114 | proof (auto simp add: positive_def lambda_system_eq Int) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 115 | fix u | 
| 47694 | 116 | assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M" | 
| 117 | and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z" | |
| 118 | and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z" | |
| 119 | have "u - x \<inter> y \<in> M" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 120 | by (metis Diff Diff_Int Un u x y) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 121 | moreover | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 122 | have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 123 | moreover | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 124 | have "u - x \<inter> y - y = u - y" by blast | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 125 | ultimately | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 126 | have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 127 | by force | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 128 | have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 129 | = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 130 | by (simp add: ey ac_simps) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 131 | also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 132 | by (simp add: Int_ac) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 133 | also have "... = f (u \<inter> y) + f (u - y)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 134 | using fx [THEN bspec, of "u \<inter> y"] Int y u | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 135 | by force | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 136 | also have "... = f u" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 137 | by (metis fy u) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 138 | finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" . | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 139 | qed | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 140 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 141 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 142 | lemma (in algebra) lambda_system_Un: | 
| 43920 | 143 | fixes f:: "'a set \<Rightarrow> ereal" | 
| 47694 | 144 | assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" | 
| 145 | shows "x \<union> y \<in> lambda_system \<Omega> M f" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 146 | proof - | 
| 47694 | 147 | have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M" | 
| 38656 | 148 | by (metis Diff_Un Un compl_sets lambda_system_sets xl yl) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 149 | moreover | 
| 47694 | 150 | have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))" | 
| 46731 | 151 | by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+ | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 152 | ultimately show ?thesis | 
| 38656 | 153 | by (metis lambda_system_Compl lambda_system_Int xl yl) | 
| 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 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 156 | lemma (in algebra) lambda_system_algebra: | 
| 47694 | 157 | "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)" | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 158 | apply (auto simp add: algebra_iff_Un) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 159 | apply (metis lambda_system_sets set_mp sets_into_space) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 160 | apply (metis lambda_system_empty) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 161 | apply (metis lambda_system_Compl) | 
| 38656 | 162 | apply (metis lambda_system_Un) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 163 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 164 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 165 | lemma (in algebra) lambda_system_strong_additive: | 
| 47694 | 166 |   assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
 | 
| 167 | and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 168 | shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 169 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 170 | have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 171 | moreover | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 172 | have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 173 | moreover | 
| 47694 | 174 | have "(z \<inter> (x \<union> y)) \<in> M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 175 | by (metis Int Un lambda_system_sets xl yl z) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 176 | ultimately show ?thesis using xl yl | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 177 | by (simp add: lambda_system_eq) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 178 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 179 | |
| 47694 | 180 | lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 181 | proof (auto simp add: additive_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 182 | fix x and y | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 183 |   assume disj: "x \<inter> y = {}"
 | 
| 47694 | 184 | and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f" | 
| 185 | hence "x \<in> M" "y \<in> M" by (blast intro: lambda_system_sets)+ | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 186 | thus "f (x \<union> y) = f x + f y" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 187 | using lambda_system_strong_additive [OF top disj xl yl] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 188 | by (simp add: Un) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 189 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 190 | |
| 42145 | 191 | lemma (in ring_of_sets) countably_subadditive_subadditive: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 192 | assumes f: "positive M f" and cs: "countably_subadditive M f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 193 | shows "subadditive M f" | 
| 35582 | 194 | proof (auto simp add: subadditive_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 195 | fix x y | 
| 47694 | 196 |   assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 197 | hence "disjoint_family (binaryset x y)" | 
| 35582 | 198 | by (auto simp add: disjoint_family_on_def binaryset_def) | 
| 47694 | 199 | hence "range (binaryset x y) \<subseteq> M \<longrightarrow> | 
| 200 | (\<Union>i. binaryset x y i) \<in> M \<longrightarrow> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 201 | f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 202 | using cs by (auto simp add: countably_subadditive_def) | 
| 47694 | 203 |   hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 204 | f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 205 | by (simp add: range_binaryset_eq UN_binaryset_eq) | 
| 38656 | 206 | thus "f (x \<union> y) \<le> f x + f y" using f x y | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 207 | by (auto simp add: Un o_def suminf_binaryset_eq positive_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 208 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 209 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 210 | lemma lambda_system_increasing: | 
| 47694 | 211 | "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f" | 
| 38656 | 212 | by (simp add: increasing_def lambda_system_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 213 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 214 | lemma lambda_system_positive: | 
| 47694 | 215 | "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 216 | by (simp add: positive_def lambda_system_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 217 | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 218 | lemma (in algebra) lambda_system_strong_sum: | 
| 43920 | 219 | fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal" | 
| 47694 | 220 | assumes f: "positive M f" and a: "a \<in> M" | 
| 221 | and A: "range A \<subseteq> lambda_system \<Omega> M f" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 222 | and disj: "disjoint_family A" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 223 |   shows  "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 224 | proof (induct n) | 
| 38656 | 225 | case 0 show ?case using f by (simp add: positive_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 226 | next | 
| 38656 | 227 | case (Suc n) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 228 |   have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
 | 
| 38656 | 229 | by (force simp add: disjoint_family_on_def neq_iff) | 
| 47694 | 230 | have 3: "A n \<in> lambda_system \<Omega> M f" using A | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 231 | by blast | 
| 47694 | 232 | interpret l: algebra \<Omega> "lambda_system \<Omega> M f" | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 233 | using f by (rule lambda_system_algebra) | 
| 47694 | 234 |   have 4: "UNION {0..<n} A \<in> lambda_system \<Omega> M f"
 | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 235 | using A l.UNION_in_sets by simp | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 236 | from Suc.hyps show ?case | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 237 | by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4]) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 238 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 239 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 240 | lemma (in sigma_algebra) lambda_system_caratheodory: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 241 | assumes oms: "outer_measure_space M f" | 
| 47694 | 242 | and A: "range A \<subseteq> lambda_system \<Omega> M f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 243 | and disj: "disjoint_family A" | 
| 47694 | 244 | shows "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 245 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 246 | have pos: "positive M f" and inc: "increasing M f" | 
| 38656 | 247 | and csa: "countably_subadditive M f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 248 | by (metis oms outer_measure_space_def)+ | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 249 | have sa: "subadditive M f" | 
| 38656 | 250 | by (metis countably_subadditive_subadditive csa pos) | 
| 47694 | 251 | have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A | 
| 252 | by auto | |
| 253 | interpret ls: algebra \<Omega> "lambda_system \<Omega> M f" | |
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 254 | using pos by (rule lambda_system_algebra) | 
| 47694 | 255 | have A'': "range A \<subseteq> M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 256 | by (metis A image_subset_iff lambda_system_sets) | 
| 38656 | 257 | |
| 47694 | 258 | have U_in: "(\<Union>i. A i) \<in> M" | 
| 37032 | 259 | by (metis A'' countable_UN) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 260 | have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (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: 
41023diff
changeset | 261 | proof (rule antisym) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 262 | show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 263 | using csa[unfolded countably_subadditive_def] A'' disj U_in by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 264 | have *: "\<And>i. 0 \<le> f (A i)" using pos A'' unfolding positive_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 265 |     have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 266 | show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)" | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 267 | using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 268 | using A'' | 
| 47694 | 269 | by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] countable_UN) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 270 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 271 |   {
 | 
| 38656 | 272 | fix a | 
| 47694 | 273 | assume a [iff]: "a \<in> M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 274 | have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 275 | proof - | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 276 | show ?thesis | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 277 | proof (rule antisym) | 
| 47694 | 278 | have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A'' | 
| 33536 | 279 | by blast | 
| 38656 | 280 | moreover | 
| 33536 | 281 | have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj | 
| 38656 | 282 | by (auto simp add: disjoint_family_on_def) | 
| 283 | moreover | |
| 47694 | 284 | have "a \<inter> (\<Union>i. A i) \<in> M" | 
| 33536 | 285 | by (metis Int U_in a) | 
| 38656 | 286 | ultimately | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 287 | have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 288 | using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"] | 
| 38656 | 289 | by (simp add: o_def) | 
| 290 | hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 291 | (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))" | 
| 38656 | 292 | by (rule add_right_mono) | 
| 293 | moreover | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 294 | have "(\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 295 | proof (intro suminf_bound_add allI) | 
| 33536 | 296 | fix n | 
| 47694 | 297 |             have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
 | 
| 38656 | 298 | by (metis A'' UNION_in_sets) | 
| 33536 | 299 |             have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
 | 
| 37032 | 300 | by (blast intro: increasingD [OF inc] A'' UNION_in_sets) | 
| 47694 | 301 |             have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f"
 | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 302 | using ls.UNION_in_sets by (simp add: A) | 
| 38656 | 303 |             hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
 | 
| 37032 | 304 | by (simp add: lambda_system_eq UNION_in) | 
| 33536 | 305 |             have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
 | 
| 44106 | 306 | by (blast intro: increasingD [OF inc] UNION_in U_in) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 307 | thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a" | 
| 38656 | 308 | by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 309 | next | 
| 47694 | 310 | have "\<And>i. a \<inter> A i \<in> M" using A'' by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 311 | then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto | 
| 47694 | 312 | have "\<And>i. a - (\<Union>i. A i) \<in> M" using A'' by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 313 | then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 314 | then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto | 
| 33536 | 315 | qed | 
| 38656 | 316 | ultimately show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a" | 
| 317 | by (rule order_trans) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 318 | next | 
| 38656 | 319 | have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))" | 
| 37032 | 320 | by (blast intro: increasingD [OF inc] U_in) | 
| 33536 | 321 | also have "... \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" | 
| 37032 | 322 | by (blast intro: subadditiveD [OF sa] U_in) | 
| 33536 | 323 | finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" . | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 324 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 325 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 326 | } | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 327 | thus ?thesis | 
| 38656 | 328 | by (simp add: lambda_system_eq sums_iff U_eq U_in) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 329 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 330 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 331 | lemma (in sigma_algebra) caratheodory_lemma: | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 332 | assumes oms: "outer_measure_space M f" | 
| 47694 | 333 | defines "L \<equiv> lambda_system \<Omega> M f" | 
| 334 | shows "measure_space \<Omega> L f" | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 335 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 336 | have pos: "positive M f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 337 | by (metis oms outer_measure_space_def) | 
| 47694 | 338 | have alg: "algebra \<Omega> L" | 
| 38656 | 339 | using lambda_system_algebra [of f, OF pos] | 
| 47694 | 340 | by (simp add: algebra_iff_Un L_def) | 
| 42065 
2b98b4c2e2f1
add ring_of_sets and subset_class as basis for algebra
 hoelzl parents: 
41981diff
changeset | 341 | then | 
| 47694 | 342 | have "sigma_algebra \<Omega> L" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 343 | using lambda_system_caratheodory [OF oms] | 
| 47694 | 344 | by (simp add: sigma_algebra_disjoint_iff L_def) | 
| 38656 | 345 | moreover | 
| 47694 | 346 | have "countably_additive L f" "positive L f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 347 | using pos lambda_system_caratheodory [OF oms] | 
| 47694 | 348 | by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def) | 
| 38656 | 349 | ultimately | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 350 | show ?thesis | 
| 47694 | 351 | using pos by (simp add: measure_space_def) | 
| 38656 | 352 | qed | 
| 353 | ||
| 39096 | 354 | lemma inf_measure_nonempty: | 
| 47694 | 355 |   assumes f: "positive M f" and b: "b \<in> M" and a: "a \<subseteq> b" "{} \<in> M"
 | 
| 39096 | 356 | shows "f b \<in> measure_set M f a" | 
| 357 | proof - | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 358 |   let ?A = "\<lambda>i::nat. (if i = 0 then b else {})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 359 | have "(\<Sum>i. f (?A i)) = (\<Sum>i<1::nat. f (?A i))" | 
| 47761 | 360 | by (rule suminf_finite) (simp_all add: f[unfolded positive_def]) | 
| 39096 | 361 | also have "... = f b" | 
| 362 | by simp | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 363 | finally show ?thesis using assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 364 | by (auto intro!: exI [of _ ?A] | 
| 39096 | 365 | simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) | 
| 366 | qed | |
| 367 | ||
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 368 | lemma (in ring_of_sets) inf_measure_agrees: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 369 | assumes posf: "positive M f" and ca: "countably_additive M f" | 
| 47694 | 370 | and s: "s \<in> M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 371 | shows "Inf (measure_set M f s) = f s" | 
| 43920 | 372 | unfolding Inf_ereal_def | 
| 38656 | 373 | proof (safe intro!: Greatest_equality) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 374 | fix z | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 375 | assume z: "z \<in> measure_set M f s" | 
| 38656 | 376 | from this obtain A where | 
| 47694 | 377 | A: "range A \<subseteq> M" and disj: "disjoint_family A" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 378 | and "s \<subseteq> (\<Union>x. A x)" and si: "(\<Sum>i. f (A i)) = z" | 
| 38656 | 379 | by (auto simp add: measure_set_def comp_def) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 380 | hence seq: "s = (\<Union>i. A i \<inter> s)" by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 381 | have inc: "increasing M f" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 382 | by (metis additive_increasing ca countably_additive_additive posf) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 383 | have sums: "(\<Sum>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 384 | proof (rule ca[unfolded countably_additive_def, rule_format]) | 
| 47694 | 385 | show "range (\<lambda>n. A n \<inter> s) \<subseteq> M" using A s | 
| 33536 | 386 | by blast | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 387 | show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj | 
| 35582 | 388 | by (auto simp add: disjoint_family_on_def) | 
| 47694 | 389 | show "(\<Union>i. A i \<inter> s) \<in> M" using A s | 
| 33536 | 390 | by (metis UN_extend_simps(4) s seq) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 391 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 392 | hence "f s = (\<Sum>i. f (A i \<inter> s))" | 
| 37032 | 393 | using seq [symmetric] by (simp add: sums_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 394 | also have "... \<le> (\<Sum>i. f (A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 395 | proof (rule suminf_le_pos) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 396 | fix n show "f (A n \<inter> s) \<le> f (A n)" using A s | 
| 38656 | 397 | by (force intro: increasingD [OF inc]) | 
| 47694 | 398 | fix N have "A N \<inter> s \<in> M" using A s by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 399 | then show "0 \<le> f (A N \<inter> s)" using posf unfolding positive_def by auto | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 400 | qed | 
| 38656 | 401 | also have "... = z" by (rule si) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 402 | finally show "f s \<le> z" . | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 403 | next | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 404 | fix y | 
| 38656 | 405 | assume y: "\<forall>u \<in> measure_set M f s. y \<le> u" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 406 | thus "y \<le> f s" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 407 | by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 408 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 409 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 410 | lemma measure_set_pos: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 411 | assumes posf: "positive M f" "r \<in> measure_set M f X" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 412 | shows "0 \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 413 | proof - | 
| 47694 | 414 | obtain A where "range A \<subseteq> M" and r: "r = (\<Sum>i. f (A i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 415 | using `r \<in> measure_set M f X` unfolding measure_set_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 416 | then show "0 \<le> r" using posf unfolding r positive_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 417 | by (intro suminf_0_le) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 418 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 419 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 420 | lemma inf_measure_pos: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 421 | assumes posf: "positive M f" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 422 | shows "0 \<le> Inf (measure_set M f X)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 423 | proof (rule complete_lattice_class.Inf_greatest) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 424 | fix r assume "r \<in> measure_set M f X" with posf show "0 \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 425 | by (rule measure_set_pos) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 426 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 427 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 428 | lemma inf_measure_empty: | 
| 47694 | 429 |   assumes posf: "positive M f" and "{} \<in> M"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 430 |   shows "Inf (measure_set M f {}) = 0"
 | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 431 | proof (rule antisym) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 432 |   show "Inf (measure_set M f {}) \<le> 0"
 | 
| 47694 | 433 |     by (metis complete_lattice_class.Inf_lower `{} \<in> M`
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 434 | inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 435 | qed (rule inf_measure_pos[OF posf]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 436 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 437 | lemma (in ring_of_sets) inf_measure_positive: | 
| 47694 | 438 |   assumes p: "positive M f" and "{} \<in> M"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 439 | shows "positive M (\<lambda>x. Inf (measure_set M f x))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 440 | proof (unfold positive_def, intro conjI ballI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 441 |   show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
 | 
| 47694 | 442 | fix A assume "A \<in> M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 443 | qed (rule inf_measure_pos[OF p]) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 444 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 445 | lemma (in ring_of_sets) inf_measure_increasing: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 446 | assumes posf: "positive M f" | 
| 47694 | 447 | shows "increasing (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))" | 
| 44918 | 448 | apply (clarsimp simp add: increasing_def) | 
| 38656 | 449 | apply (rule complete_lattice_class.Inf_greatest) | 
| 450 | apply (rule complete_lattice_class.Inf_lower) | |
| 37032 | 451 | apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 452 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 453 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 454 | lemma (in ring_of_sets) inf_measure_le: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 455 | assumes posf: "positive M f" and inc: "increasing M f" | 
| 47694 | 456 |       and x: "x \<in> {r . \<exists>A. range A \<subseteq> M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
 | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 457 | shows "Inf (measure_set M f s) \<le> x" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 458 | proof - | 
| 47694 | 459 | obtain A where A: "range A \<subseteq> M" and ss: "s \<subseteq> (\<Union>i. A i)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 460 | and xeq: "(\<Sum>i. f (A i)) = x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 461 | using x by auto | 
| 47694 | 462 | have dA: "range (disjointed A) \<subseteq> M" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 463 | by (metis A range_disjointed_sets) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 464 | have "\<forall>n. f (disjointed A n) \<le> f (A n)" | 
| 38656 | 465 | by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 466 | moreover have "\<forall>i. 0 \<le> f (disjointed A i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 467 | using posf dA unfolding positive_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 468 | ultimately have sda: "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 469 | by (blast intro!: suminf_le_pos) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 470 | hence ley: "(\<Sum>i. f (disjointed A i)) \<le> x" | 
| 38656 | 471 | by (metis xeq) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 472 | hence y: "(\<Sum>i. f (disjointed A i)) \<in> measure_set M f s" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 473 | apply (auto simp add: measure_set_def) | 
| 38656 | 474 | apply (rule_tac x="disjointed A" in exI) | 
| 475 | apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 476 | done | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 477 | show ?thesis | 
| 38656 | 478 | by (blast intro: y order_trans [OF _ ley] posf complete_lattice_class.Inf_lower) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 479 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 480 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 481 | lemma (in ring_of_sets) inf_measure_close: | 
| 43920 | 482 | fixes e :: ereal | 
| 47694 | 483 | assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (\<Omega>)" and "Inf (measure_set M f s) \<noteq> \<infinity>" | 
| 484 | shows "\<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and> | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 485 | (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 486 | proof - | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 487 | from `Inf (measure_set M f s) \<noteq> \<infinity>` have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 488 | using inf_measure_pos[OF posf, of s] by auto | 
| 38656 | 489 | obtain l where "l \<in> measure_set M f s" "l \<le> Inf (measure_set M f s) + e" | 
| 43920 | 490 | using Inf_ereal_close[OF fin e] by auto | 
| 38656 | 491 | thus ?thesis | 
| 492 | by (auto intro!: exI[of _ l] simp: measure_set_def comp_def) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 493 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 494 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 495 | lemma (in ring_of_sets) inf_measure_countably_subadditive: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 496 | assumes posf: "positive M f" and inc: "increasing M f" | 
| 47694 | 497 | shows "countably_subadditive (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 498 | proof (simp add: countably_subadditive_def, safe) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 499 | fix A :: "nat \<Rightarrow> 'a set" | 
| 46731 | 500 | let ?outer = "\<lambda>B. Inf (measure_set M f B)" | 
| 47694 | 501 | assume A: "range A \<subseteq> Pow (\<Omega>)" | 
| 38656 | 502 | and disj: "disjoint_family A" | 
| 47694 | 503 | and sb: "(\<Union>i. A i) \<subseteq> \<Omega>" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 504 | |
| 43920 | 505 |   { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
 | 
| 47694 | 506 | hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> M \<and> disjoint_family (BB n) \<and> | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 507 | A n \<subseteq> (\<Union>i. BB n i) \<and> (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 508 | apply (safe intro!: choice inf_measure_close [of f, OF posf]) | 
| 43920 | 509 | using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 510 | then obtain BB | 
| 47694 | 511 | where BB: "\<And>n. (range (BB n) \<subseteq> M)" | 
| 38656 | 512 | and disjBB: "\<And>n. disjoint_family (BB n)" | 
| 513 | and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)" | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 514 | and BBle: "\<And>n. (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 515 | by auto blast | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 516 | have sll: "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n)) + e" | 
| 38656 | 517 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 518 | have sum_eq_1: "(\<Sum>n. e*(1/2) ^ Suc n) = e" | 
| 43920 | 519 | using suminf_half_series_ereal e | 
| 520 | by (simp add: ereal_zero_le_0_iff zero_le_divide_ereal suminf_cmult_ereal) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 521 | have "\<And>n i. 0 \<le> f (BB n i)" using posf[unfolded positive_def] BB by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 522 | then have "\<And>n. 0 \<le> (\<Sum>i. f (BB n i))" by (rule suminf_0_le) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 523 | then have "(\<Sum>n. \<Sum>i. (f (BB n i))) \<le> (\<Sum>n. ?outer (A n) + e*(1/2) ^ Suc n)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 524 | by (rule suminf_le_pos[OF BBle]) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 525 | also have "... = (\<Sum>n. ?outer (A n)) + e" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 526 | using sum_eq_1 inf_measure_pos[OF posf] e | 
| 43920 | 527 | by (subst suminf_add_ereal) (auto simp add: ereal_zero_le_0_iff) | 
| 38656 | 528 | finally show ?thesis . | 
| 529 | qed | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 530 | def C \<equiv> "(split BB) o prod_decode" | 
| 47694 | 531 | have C: "!!n. C n \<in> M" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 532 | apply (rule_tac p="prod_decode n" in PairE) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 533 | apply (simp add: C_def) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 534 | apply (metis BB subsetD rangeI) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 535 | done | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 536 | have sbC: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)" | 
| 38656 | 537 | proof (auto simp add: C_def) | 
| 538 | fix x i | |
| 539 | assume x: "x \<in> A i" | |
| 540 | with sbBB [of i] obtain j where "x \<in> BB i j" | |
| 541 | by blast | |
| 542 | thus "\<exists>i. x \<in> split BB (prod_decode i)" | |
| 543 | by (metis prod_encode_inverse prod.cases) | |
| 544 | qed | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 545 | have "(f \<circ> C) = (f \<circ> (\<lambda>(x, y). BB x y)) \<circ> prod_decode" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 546 | by (rule ext) (auto simp add: C_def) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 547 | moreover have "suminf ... = (\<Sum>n. \<Sum>i. f (BB n i))" using BBle | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 548 | using BB posf[unfolded positive_def] | 
| 43920 | 549 | by (force intro!: suminf_ereal_2dimen simp: o_def) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 550 | ultimately have Csums: "(\<Sum>i. f (C i)) = (\<Sum>n. \<Sum>i. f (BB n i))" by (simp add: o_def) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 551 | have "?outer (\<Union>i. A i) \<le> (\<Sum>n. \<Sum>i. f (BB n i))" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 552 | apply (rule inf_measure_le [OF posf(1) inc], auto) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 553 | apply (rule_tac x="C" in exI) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 554 | apply (auto simp add: C sbC Csums) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 555 | done | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 556 | also have "... \<le> (\<Sum>n. ?outer (A n)) + e" using sll | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 557 | by blast | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 558 | finally have "?outer (\<Union>i. A i) \<le> (\<Sum>n. ?outer (A n)) + e" . } | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 559 | note for_finite_Inf = this | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 560 | |
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 561 | show "?outer (\<Union>i. A i) \<le> (\<Sum>n. ?outer (A n))" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 562 | proof cases | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 563 | assume "\<forall>i. ?outer (A i) \<noteq> \<infinity>" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 564 | with for_finite_Inf show ?thesis | 
| 43920 | 565 | by (intro ereal_le_epsilon) auto | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 566 | next | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 567 | assume "\<not> (\<forall>i. ?outer (A i) \<noteq> \<infinity>)" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 568 | then have "\<exists>i. ?outer (A i) = \<infinity>" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 569 | by auto | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 570 | then have "(\<Sum>n. ?outer (A n)) = \<infinity>" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 571 | using suminf_PInfty[OF inf_measure_pos, OF posf] | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 572 | by metis | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 573 | then show ?thesis by simp | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 574 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 575 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 576 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 577 | lemma (in ring_of_sets) inf_measure_outer: | 
| 47694 | 578 | "\<lbrakk> positive M f ; increasing M f \<rbrakk> \<Longrightarrow> | 
| 579 | outer_measure_space (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 580 | using inf_measure_pos[of M f] | 
| 38656 | 581 | by (simp add: outer_measure_space_def inf_measure_empty | 
| 582 | inf_measure_increasing inf_measure_countably_subadditive positive_def) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 583 | |
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 584 | lemma (in ring_of_sets) algebra_subset_lambda_system: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 585 | assumes posf: "positive M f" and inc: "increasing M f" | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 586 | and add: "additive M f" | 
| 47694 | 587 | shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))" | 
| 38656 | 588 | proof (auto dest: sets_into_space | 
| 589 | simp add: algebra.lambda_system_eq [OF algebra_Pow]) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 590 | fix x s | 
| 47694 | 591 | assume x: "x \<in> M" | 
| 592 | and s: "s \<subseteq> \<Omega>" | |
| 593 | have [simp]: "!!x. x \<in> M \<Longrightarrow> s \<inter> (\<Omega> - x) = s-x" using s | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 594 | by blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 595 | have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x)) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 596 | \<le> Inf (measure_set M f s)" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 597 | proof cases | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 598 | assume "Inf (measure_set M f s) = \<infinity>" then show ?thesis by simp | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 599 | next | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 600 | assume fin: "Inf (measure_set M f s) \<noteq> \<infinity>" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 601 |     then have "measure_set M f s \<noteq> {}"
 | 
| 43920 | 602 | by (auto simp: top_ereal_def) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 603 | show ?thesis | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 604 | proof (rule complete_lattice_class.Inf_greatest) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 605 | fix r assume "r \<in> measure_set M f s" | 
| 47694 | 606 | then obtain A where A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)" | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 607 | and r: "r = (\<Sum>i. f (A i))" unfolding measure_set_def by auto | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 608 | have "Inf (measure_set M f (s \<inter> x)) \<le> (\<Sum>i. f (A i \<inter> x))" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 609 | unfolding measure_set_def | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 610 | proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\<lambda>i. A i \<inter> x"]) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 611 | from A(1) show "disjoint_family (\<lambda>i. A i \<inter> x)" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 612 | by (rule disjoint_family_on_bisimulation) auto | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 613 | qed (insert x A, auto) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 614 | moreover | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 615 | have "Inf (measure_set M f (s - x)) \<le> (\<Sum>i. f (A i - x))" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 616 | unfolding measure_set_def | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 617 | proof (safe intro!: complete_lattice_class.Inf_lower exI[of _ "\<lambda>i. A i - x"]) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 618 | from A(1) show "disjoint_family (\<lambda>i. A i - x)" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 619 | by (rule disjoint_family_on_bisimulation) auto | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 620 | qed (insert x A, auto) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 621 | ultimately have "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le> | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 622 | (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 623 | also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))" | 
| 43920 | 624 | using A(2) x posf by (subst suminf_add_ereal) (auto simp: positive_def) | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 625 | also have "\<dots> = (\<Sum>i. f (A i))" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 626 | using A x | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 627 | by (subst add[THEN additiveD, symmetric]) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 628 | (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f]) | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 629 | finally show "Inf (measure_set M f (s \<inter> x)) + Inf (measure_set M f (s - x)) \<le> r" | 
| 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 630 | using r by simp | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 631 | qed | 
| 42066 
6db76c88907a
generalized Caratheodory from algebra to ring_of_sets
 hoelzl parents: 
42065diff
changeset | 632 | qed | 
| 38656 | 633 | moreover | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 634 | have "Inf (measure_set M f s) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 635 | \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))" | 
| 42145 | 636 | proof - | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 637 | have "Inf (measure_set M f s) = Inf (measure_set M f ((s\<inter>x) \<union> (s-x)))" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 638 | by (metis Un_Diff_Int Un_commute) | 
| 38656 | 639 | also have "... \<le> Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))" | 
| 640 | apply (rule subadditiveD) | |
| 42145 | 641 | apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41689diff
changeset | 642 | apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf]) | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 643 | apply (rule inf_measure_countably_subadditive) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 644 | using s by (auto intro!: posf inc) | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 645 | finally show ?thesis . | 
| 42145 | 646 | qed | 
| 38656 | 647 | ultimately | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 648 | show "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x)) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 649 | = Inf (measure_set M f s)" | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 650 | by (rule order_antisym) | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 651 | qed | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 652 | |
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 653 | lemma measure_down: | 
| 47694 | 654 | "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>" | 
| 655 | by (simp add: measure_space_def positive_def countably_additive_def) | |
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 656 | blast | 
| 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 657 | |
| 47762 | 658 | theorem (in ring_of_sets) caratheodory': | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 659 | assumes posf: "positive M f" and ca: "countably_additive M f" | 
| 47694 | 660 | shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 661 | proof - | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 662 | have inc: "increasing M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 663 | by (metis additive_increasing ca countably_additive_additive posf) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 664 | let ?infm = "(\<lambda>x. Inf (measure_set M f x))" | 
| 47694 | 665 | def ls \<equiv> "lambda_system \<Omega> (Pow \<Omega>) ?infm" | 
| 666 | have mls: "measure_space \<Omega> ls ?infm" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 667 | using sigma_algebra.caratheodory_lemma | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 668 | [OF sigma_algebra_Pow inf_measure_outer [OF posf inc]] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 669 | by (simp add: ls_def) | 
| 47694 | 670 | hence sls: "sigma_algebra \<Omega> ls" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 671 | by (simp add: measure_space_def) | 
| 47694 | 672 | have "M \<subseteq> ls" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 673 | by (simp add: ls_def) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 674 | (metis ca posf inc countably_additive_additive algebra_subset_lambda_system) | 
| 47694 | 675 | hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls" | 
| 676 | using sigma_algebra.sigma_sets_subset [OF sls, of "M"] | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 677 | by simp | 
| 47694 | 678 | have "measure_space \<Omega> (sigma_sets \<Omega> M) ?infm" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 679 | by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 680 | (simp_all add: sgs_sb space_closed) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 681 | thus ?thesis using inf_measure_agrees [OF posf ca] | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 682 | by (intro exI[of _ ?infm]) auto | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41023diff
changeset | 683 | qed | 
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 684 | |
| 42145 | 685 | subsubsection {*Alternative instances of caratheodory*}
 | 
| 686 | ||
| 687 | lemma (in ring_of_sets) countably_additive_iff_continuous_from_below: | |
| 688 | assumes f: "positive M f" "additive M f" | |
| 689 | shows "countably_additive M f \<longleftrightarrow> | |
| 47694 | 690 | (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))" | 
| 42145 | 691 | unfolding countably_additive_def | 
| 692 | proof safe | |
| 47694 | 693 | assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)" | 
| 694 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M" | |
| 695 | then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets) | |
| 42145 | 696 | with count_sum[THEN spec, of "disjointed A"] A(3) | 
| 697 | have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)" | |
| 698 | by (auto simp: UN_disjointed_eq disjoint_family_disjointed) | |
| 699 | moreover have "(\<lambda>n. (\<Sum>i=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))" | |
| 700 | using f(1)[unfolded positive_def] dA | |
| 43920 | 701 | by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos) | 
| 42145 | 702 | from LIMSEQ_Suc[OF this] | 
| 703 | have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))" | |
| 704 | unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost . | |
| 705 | moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)" | |
| 706 | using disjointed_additive[OF f A(1,2)] . | |
| 707 | ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp | |
| 708 | next | |
| 47694 | 709 | assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" | 
| 710 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M" | |
| 42145 | 711 | have *: "(\<Union>n. (\<Union>i\<le>n. A i)) = (\<Union>i. A i)" by auto | 
| 712 | have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)" | |
| 713 | proof (unfold *[symmetric], intro cont[rule_format]) | |
| 47694 | 714 | show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> M" | 
| 42145 | 715 | using A * by auto | 
| 716 | qed (force intro!: incseq_SucI) | |
| 717 | moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>n. f (A i))" | |
| 718 | using A | |
| 719 | by (intro additive_sum[OF f, of _ A, symmetric]) | |
| 720 | (auto intro: disjoint_family_on_mono[where B=UNIV]) | |
| 721 | ultimately | |
| 722 | have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)" | |
| 723 | unfolding sums_def2 by simp | |
| 724 | from sums_unique[OF this] | |
| 725 | show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp | |
| 726 | qed | |
| 727 | ||
| 728 | lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous: | |
| 729 | assumes f: "positive M f" "additive M f" | |
| 47694 | 730 | shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)) | 
| 731 |      \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
 | |
| 42145 | 732 | proof safe | 
| 47694 | 733 | assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))" | 
| 734 |   fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
 | |
| 42145 | 735 | with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0" | 
| 736 | using `positive M f`[unfolded positive_def] by auto | |
| 737 | next | |
| 47694 | 738 |   assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
 | 
| 739 | fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>" | |
| 42145 | 740 | |
| 47694 | 741 | have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b" | 
| 42145 | 742 | using additive_increasing[OF f] unfolding increasing_def by simp | 
| 743 | ||
| 744 | have decseq_fA: "decseq (\<lambda>i. f (A i))" | |
| 745 | using A by (auto simp: decseq_def intro!: f_mono) | |
| 746 | have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))" | |
| 747 | using A by (auto simp: decseq_def) | |
| 748 | then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))" | |
| 749 | using A unfolding decseq_def by (auto intro!: f_mono Diff) | |
| 750 | have "f (\<Inter>x. A x) \<le> f (A 0)" | |
| 751 | using A by (auto intro!: f_mono) | |
| 752 | then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>" | |
| 753 | using A by auto | |
| 754 |   { fix i
 | |
| 755 | have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono) | |
| 756 | then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>" | |
| 757 | using A by auto } | |
| 758 | note f_fin = this | |
| 759 | have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 0" | |
| 760 | proof (intro cont[rule_format, OF _ decseq _ f_fin]) | |
| 47694 | 761 |     show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
 | 
| 42145 | 762 | using A by auto | 
| 763 | qed | |
| 43920 | 764 | from INF_Lim_ereal[OF decseq_f this] | 
| 42145 | 765 | have "(INF n. f (A n - (\<Inter>i. A i))) = 0" . | 
| 766 | moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)" | |
| 767 | by auto | |
| 768 | ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)" | |
| 769 | using A(4) f_fin f_Int_fin | |
| 43920 | 770 | by (subst INFI_ereal_add) (auto simp: decseq_f) | 
| 42145 | 771 |   moreover {
 | 
| 772 | fix n | |
| 773 | have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))" | |
| 774 | using A by (subst f(2)[THEN additiveD]) auto | |
| 775 | also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n" | |
| 776 | by auto | |
| 777 | finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . } | |
| 778 | ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)" | |
| 779 | by simp | |
| 43920 | 780 | with LIMSEQ_ereal_INFI[OF decseq_fA] | 
| 42145 | 781 | show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp | 
| 782 | qed | |
| 783 | ||
| 784 | lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
 | |
| 47694 | 785 | lemma positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def) | 
| 42145 | 786 | |
| 787 | lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below: | |
| 47694 | 788 | assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>" | 
| 789 |   assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
 | |
| 790 | assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M" | |
| 42145 | 791 | shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" | 
| 792 | proof - | |
| 47694 | 793 | have "\<forall>A\<in>M. \<exists>x. f A = ereal x" | 
| 42145 | 794 | proof | 
| 47694 | 795 | fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x" | 
| 42145 | 796 | unfolding positive_def by (cases "f A") auto | 
| 797 | qed | |
| 798 | from bchoice[OF this] guess f' .. note f' = this[rule_format] | |
| 799 | from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0" | |
| 800 | by (intro cont[rule_format]) (auto simp: decseq_def incseq_def) | |
| 801 | moreover | |
| 802 |   { fix i
 | |
| 803 | have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)" | |
| 804 | using A by (intro f(2)[THEN additiveD, symmetric]) auto | |
| 805 | also have "(\<Union>i. A i) - A i \<union> A i = (\<Union>i. A i)" | |
| 806 | by auto | |
| 807 | finally have "f' (\<Union>i. A i) - f' (A i) = f' ((\<Union>i. A i) - A i)" | |
| 808 | using A by (subst (asm) (1 2 3) f') auto | |
| 43920 | 809 | then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))" | 
| 42145 | 810 | using A f' by auto } | 
| 811 | ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0" | |
| 43920 | 812 | by (simp add: zero_ereal_def) | 
| 42145 | 813 | then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)" | 
| 44568 
e6f291cb5810
discontinue many legacy theorems about LIM and LIMSEQ, in favor of tendsto theorems
 huffman parents: 
44106diff
changeset | 814 | by (rule LIMSEQ_diff_approach_zero2[OF tendsto_const]) | 
| 42145 | 815 | then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" | 
| 816 | using A by (subst (1 2) f') auto | |
| 817 | qed | |
| 818 | ||
| 819 | lemma (in ring_of_sets) empty_continuous_imp_countably_additive: | |
| 47694 | 820 | assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. f A \<noteq> \<infinity>" | 
| 821 |   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
 | |
| 42145 | 822 | shows "countably_additive M f" | 
| 823 | using countably_additive_iff_continuous_from_below[OF f] | |
| 824 | using empty_continuous_imp_continuous_from_below[OF f fin] cont | |
| 825 | by blast | |
| 826 | ||
| 827 | lemma (in ring_of_sets) caratheodory_empty_continuous: | |
| 47694 | 828 | assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>" | 
| 829 |   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
 | |
| 830 | shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>" | |
| 47762 | 831 | proof (intro caratheodory' empty_continuous_imp_countably_additive f) | 
| 47694 | 832 | show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto | 
| 42145 | 833 | qed (rule cont) | 
| 834 | ||
| 47762 | 835 | section {* Volumes *}
 | 
| 836 | ||
| 837 | definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
 | |
| 838 | "volume M f \<longleftrightarrow> | |
| 839 |   (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
 | |
| 840 | (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))" | |
| 841 | ||
| 842 | lemma volumeI: | |
| 843 |   assumes "f {} = 0"
 | |
| 844 | assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a" | |
| 845 | assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)" | |
| 846 | shows "volume M f" | |
| 847 | using assms by (auto simp: volume_def) | |
| 848 | ||
| 849 | lemma volume_positive: | |
| 850 | "volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a" | |
| 851 | by (auto simp: volume_def) | |
| 852 | ||
| 853 | lemma volume_empty: | |
| 854 |   "volume M f \<Longrightarrow> f {} = 0"
 | |
| 855 | by (auto simp: volume_def) | |
| 856 | ||
| 857 | lemma volume_finite_additive: | |
| 858 | assumes "volume M f" | |
| 859 | assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M" | |
| 860 | shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))" | |
| 861 | proof - | |
| 862 | have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>A`I \<in> M" | |
| 863 | using A unfolding SUP_def by (auto simp: disjoint_family_on_disjoint_image) | |
| 864 | with `volume M f` have "f (\<Union>A`I) = (\<Sum>a\<in>A`I. f a)" | |
| 865 | unfolding volume_def by blast | |
| 866 | also have "\<dots> = (\<Sum>i\<in>I. f (A i))" | |
| 867 | proof (subst setsum_reindex_nonzero) | |
| 868 | fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j" | |
| 869 |     with `disjoint_family_on A I` have "A i = {}"
 | |
| 870 | by (auto simp: disjoint_family_on_def) | |
| 871 | then show "f (A i) = 0" | |
| 872 | using volume_empty[OF `volume M f`] by simp | |
| 873 | qed (auto intro: `finite I`) | |
| 874 | finally show "f (UNION I A) = (\<Sum>i\<in>I. f (A i))" | |
| 875 | by simp | |
| 876 | qed | |
| 877 | ||
| 878 | lemma (in ring_of_sets) volume_additiveI: | |
| 879 | assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a" | |
| 880 |   assumes [simp]: "\<mu> {} = 0"
 | |
| 881 |   assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
 | |
| 882 | shows "volume M \<mu>" | |
| 883 | proof (unfold volume_def, safe) | |
| 884 | fix C assume "finite C" "C \<subseteq> M" "disjoint C" | |
| 885 | then show "\<mu> (\<Union>C) = setsum \<mu> C" | |
| 886 | proof (induct C) | |
| 887 | case (insert c C) | |
| 888 | from insert(1,2,4,5) have "\<mu> (\<Union>insert c C) = \<mu> c + \<mu> (\<Union>C)" | |
| 889 | by (auto intro!: add simp: disjoint_def) | |
| 890 | with insert show ?case | |
| 891 | by (simp add: disjoint_def) | |
| 892 | qed simp | |
| 893 | qed fact+ | |
| 894 | ||
| 895 | lemma (in semiring_of_sets) extend_volume: | |
| 896 | assumes "volume M \<mu>" | |
| 897 | shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)" | |
| 898 | proof - | |
| 899 | let ?R = generated_ring | |
| 900 | have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)" | |
| 901 | by (auto simp: generated_ring_def) | |
| 902 | from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this | |
| 903 | ||
| 904 |   { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
 | |
| 905 | fix D assume D: "D \<subseteq> M" "finite D" "disjoint D" | |
| 906 | assume "\<Union>C = \<Union>D" | |
| 907 | have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))" | |
| 908 | proof (intro setsum_cong refl) | |
| 909 | fix d assume "d \<in> D" | |
| 910 | have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d" | |
| 911 | using `d \<in> D` `\<Union>C = \<Union>D` by auto | |
| 912 | moreover have "\<mu> (\<Union>c\<in>C. c \<inter> d) = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" | |
| 913 | proof (rule volume_finite_additive) | |
| 914 |         { fix c assume "c \<in> C" then show "c \<inter> d \<in> M"
 | |
| 915 | using C D `d \<in> D` by auto } | |
| 916 | show "(\<Union>a\<in>C. a \<inter> d) \<in> M" | |
| 917 | unfolding Un_eq_d using `d \<in> D` D by auto | |
| 918 | show "disjoint_family_on (\<lambda>a. a \<inter> d) C" | |
| 919 | using `disjoint C` by (auto simp: disjoint_family_on_def disjoint_def) | |
| 920 | qed fact+ | |
| 921 | ultimately show "\<mu> d = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" by simp | |
| 922 | qed } | |
| 923 | note split_sum = this | |
| 924 | ||
| 925 |   { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
 | |
| 926 | fix D assume D: "D \<subseteq> M" "finite D" "disjoint D" | |
| 927 | assume "\<Union>C = \<Union>D" | |
| 928 | with split_sum[OF C D] split_sum[OF D C] | |
| 929 | have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)" | |
| 930 | by (simp, subst setsum_commute, simp add: ac_simps) } | |
| 931 | note sum_eq = this | |
| 932 | ||
| 933 |   { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
 | |
| 934 | then have "\<Union>C \<in> ?R" by (auto simp: generated_ring_def) | |
| 935 | with \<mu>'_spec[THEN bspec, of "\<Union>C"] | |
| 936 | obtain D where | |
| 937 | D: "D \<subseteq> M" "finite D" "disjoint D" "\<Union>C = \<Union>D" and "\<mu>' (\<Union>C) = (\<Sum>d\<in>D. \<mu> d)" | |
| 938 | by blast | |
| 939 | with sum_eq[OF C D] have "\<mu>' (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" by simp } | |
| 940 | note \<mu>' = this | |
| 941 | ||
| 942 | show ?thesis | |
| 943 | proof (intro exI conjI ring_of_sets.volume_additiveI[OF generating_ring] ballI) | |
| 944 |     fix a assume "a \<in> M" with \<mu>'[of "{a}"] show "\<mu>' a = \<mu> a"
 | |
| 945 | by (simp add: disjoint_def) | |
| 946 | next | |
| 947 | fix a assume "a \<in> ?R" then guess Ca .. note Ca = this | |
| 948 | with \<mu>'[of Ca] `volume M \<mu>`[THEN volume_positive] | |
| 949 | show "0 \<le> \<mu>' a" | |
| 950 | by (auto intro!: setsum_nonneg) | |
| 951 | next | |
| 952 |     show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto
 | |
| 953 | next | |
| 954 | fix a assume "a \<in> ?R" then guess Ca .. note Ca = this | |
| 955 | fix b assume "b \<in> ?R" then guess Cb .. note Cb = this | |
| 956 |     assume "a \<inter> b = {}"
 | |
| 957 |     with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto
 | |
| 958 |     then have C_Int_cases: "Ca \<inter> Cb = {{}} \<or> Ca \<inter> Cb = {}" by auto
 | |
| 959 | ||
| 960 |     from `a \<inter> b = {}` have "\<mu>' (\<Union> (Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
 | |
| 961 | using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union) | |
| 962 | also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)" | |
| 963 | using C_Int_cases volume_empty[OF `volume M \<mu>`] by (elim disjE) simp_all | |
| 964 | also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)" | |
| 965 | using Ca Cb by (simp add: setsum_Un_Int) | |
| 966 | also have "\<dots> = \<mu>' a + \<mu>' b" | |
| 967 | using Ca Cb by (simp add: \<mu>') | |
| 968 | finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b" | |
| 969 | using Ca Cb by simp | |
| 970 | qed | |
| 971 | qed | |
| 972 | ||
| 973 | section {* Caratheodory on semirings *}
 | |
| 974 | ||
| 975 | theorem (in semiring_of_sets) caratheodory: | |
| 976 | assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>" | |
| 977 | shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'" | |
| 978 | proof - | |
| 979 | have "volume M \<mu>" | |
| 980 | proof (rule volumeI) | |
| 981 |     { fix a assume "a \<in> M" then show "0 \<le> \<mu> a"
 | |
| 982 | using pos unfolding positive_def by auto } | |
| 983 | note p = this | |
| 984 | ||
| 985 | fix C assume sets_C: "C \<subseteq> M" "\<Union>C \<in> M" and "disjoint C" "finite C" | |
| 986 |     have "\<exists>F'. bij_betw F' {..<card C} C"
 | |
| 987 | by (rule finite_same_card_bij[OF _ `finite C`]) auto | |
| 988 | then guess F' .. note F' = this | |
| 989 |     then have F': "C = F' ` {..< card C}" "inj_on F' {..< card C}"
 | |
| 990 | by (auto simp: bij_betw_def) | |
| 991 |     { fix i j assume *: "i < card C" "j < card C" "i \<noteq> j"
 | |
| 992 | with F' have "F' i \<in> C" "F' j \<in> C" "F' i \<noteq> F' j" | |
| 993 | unfolding inj_on_def by auto | |
| 994 | with `disjoint C`[THEN disjointD] | |
| 995 |       have "F' i \<inter> F' j = {}"
 | |
| 996 | by auto } | |
| 997 | note F'_disj = this | |
| 998 |     def F \<equiv> "\<lambda>i. if i < card C then F' i else {}"
 | |
| 999 | then have "disjoint_family F" | |
| 1000 | using F'_disj by (auto simp: disjoint_family_on_def) | |
| 1001 | moreover from F' have "(\<Union>i. F i) = \<Union>C" | |
| 1002 | by (auto simp: F_def set_eq_iff split: split_if_asm) | |
| 1003 | moreover have sets_F: "\<And>i. F i \<in> M" | |
| 1004 | using F' sets_C by (auto simp: F_def) | |
| 1005 | moreover note sets_C | |
| 1006 | ultimately have "\<mu> (\<Union>C) = (\<Sum>i. \<mu> (F i))" | |
| 1007 | using ca[unfolded countably_additive_def, THEN spec, of F] by auto | |
| 1008 | also have "\<dots> = (\<Sum>i<card C. \<mu> (F' i))" | |
| 1009 | proof - | |
| 1010 |       have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) sums (\<Sum>i<card C. \<mu> (F' i))"
 | |
| 1011 | by (rule sums_If_finite_set) auto | |
| 1012 |       also have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) = (\<lambda>i. \<mu> (F i))"
 | |
| 1013 | using pos by (auto simp: positive_def F_def) | |
| 1014 | finally show "(\<Sum>i. \<mu> (F i)) = (\<Sum>i<card C. \<mu> (F' i))" | |
| 1015 | by (simp add: sums_iff) | |
| 1016 | qed | |
| 1017 | also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)" | |
| 1018 | using F'(2) by (subst (2) F') (simp add: setsum_reindex) | |
| 1019 | finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" . | |
| 1020 | next | |
| 1021 |     show "\<mu> {} = 0"
 | |
| 1022 | using `positive M \<mu>` by (rule positiveD1) | |
| 1023 | qed | |
| 1024 | from extend_volume[OF this] obtain \<mu>_r where | |
| 1025 | V: "volume generated_ring \<mu>_r" "\<And>a. a \<in> M \<Longrightarrow> \<mu> a = \<mu>_r a" | |
| 1026 | by auto | |
| 1027 | ||
| 1028 | interpret G: ring_of_sets \<Omega> generated_ring | |
| 1029 | by (rule generating_ring) | |
| 1030 | ||
| 1031 | have pos: "positive generated_ring \<mu>_r" | |
| 1032 | using V unfolding positive_def by (auto simp: positive_def intro!: volume_positive volume_empty) | |
| 1033 | ||
| 1034 | have "countably_additive generated_ring \<mu>_r" | |
| 1035 | proof (rule countably_additiveI) | |
| 1036 | fix A' :: "nat \<Rightarrow> 'a set" assume A': "range A' \<subseteq> generated_ring" "disjoint_family A'" | |
| 1037 | and Un_A: "(\<Union>i. A' i) \<in> generated_ring" | |
| 1038 | ||
| 1039 | from generated_ringE[OF Un_A] guess C' . note C' = this | |
| 1040 | ||
| 1041 |     { fix c assume "c \<in> C'"
 | |
| 1042 | moreover def A \<equiv> "\<lambda>i. A' i \<inter> c" | |
| 1043 | ultimately have A: "range A \<subseteq> generated_ring" "disjoint_family A" | |
| 1044 | and Un_A: "(\<Union>i. A i) \<in> generated_ring" | |
| 1045 | using A' C' | |
| 1046 | by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def) | |
| 1047 | from A C' `c \<in> C'` have UN_eq: "(\<Union>i. A i) = c" | |
| 1048 | by (auto simp: A_def) | |
| 1049 | ||
| 1050 | have "\<forall>i::nat. \<exists>f::nat \<Rightarrow> 'a set. \<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j)) \<and> disjoint_family f \<and> \<Union>range f = A i \<and> (\<forall>j. f j \<in> M)" | |
| 1051 | (is "\<forall>i. ?P i") | |
| 1052 | proof | |
| 1053 | fix i | |
| 1054 | from A have Ai: "A i \<in> generated_ring" by auto | |
| 1055 | from generated_ringE[OF this] guess C . note C = this | |
| 1056 | ||
| 1057 |         have "\<exists>F'. bij_betw F' {..<card C} C"
 | |
| 1058 | by (rule finite_same_card_bij[OF _ `finite C`]) auto | |
| 1059 | then guess F .. note F = this | |
| 1060 |         def f \<equiv> "\<lambda>i. if i < card C then F i else {}"
 | |
| 1061 |         then have f: "bij_betw f {..< card C} C"
 | |
| 1062 | by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto | |
| 1063 | with C have "\<forall>j. f j \<in> M" | |
| 1064 | by (auto simp: Pi_iff f_def dest!: bij_betw_imp_funcset) | |
| 1065 | moreover | |
| 1066 |         from f C have d_f: "disjoint_family_on f {..<card C}"
 | |
| 1067 | by (intro disjoint_image_disjoint_family_on) (auto simp: bij_betw_def) | |
| 1068 | then have "disjoint_family f" | |
| 1069 | by (auto simp: disjoint_family_on_def f_def) | |
| 1070 | moreover | |
| 1071 | have Ai_eq: "A i = (\<Union> x<card C. f x)" | |
| 1072 | using f C Ai unfolding bij_betw_def by (simp add: Union_image_eq[symmetric]) | |
| 1073 | then have "\<Union>range f = A i" | |
| 1074 | using f C Ai unfolding bij_betw_def by (auto simp: f_def) | |
| 1075 | moreover | |
| 1076 |         { have "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_r (f j) else 0)"
 | |
| 1077 | using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def) | |
| 1078 | also have "\<dots> = (\<Sum>j<card C. \<mu>_r (f j))" | |
| 1079 | by (rule sums_If_finite_set[THEN sums_unique, symmetric]) simp | |
| 1080 | also have "\<dots> = \<mu>_r (A i)" | |
| 1081 | using C f[THEN bij_betw_imp_funcset] unfolding Ai_eq | |
| 1082 | by (intro volume_finite_additive[OF V(1) _ d_f, symmetric]) | |
| 1083 | (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic) | |
| 1084 | finally have "\<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j))" .. } | |
| 1085 | ultimately show "?P i" | |
| 1086 | by blast | |
| 1087 | qed | |
| 1088 | from choice[OF this] guess f .. note f = this | |
| 1089 | then have UN_f_eq: "(\<Union>i. split f (prod_decode i)) = (\<Union>i. A i)" | |
| 1090 | unfolding UN_extend_simps surj_prod_decode by (auto simp: set_eq_iff) | |
| 1091 | ||
| 1092 | have d: "disjoint_family (\<lambda>i. split f (prod_decode i))" | |
| 1093 | unfolding disjoint_family_on_def | |
| 1094 | proof (intro ballI impI) | |
| 1095 | fix m n :: nat assume "m \<noteq> n" | |
| 1096 | then have neq: "prod_decode m \<noteq> prod_decode n" | |
| 1097 | using inj_prod_decode[of UNIV] by (auto simp: inj_on_def) | |
| 1098 |         show "split f (prod_decode m) \<inter> split f (prod_decode n) = {}"
 | |
| 1099 | proof cases | |
| 1100 | assume "fst (prod_decode m) = fst (prod_decode n)" | |
| 1101 | then show ?thesis | |
| 1102 | using neq f by (fastforce simp: disjoint_family_on_def) | |
| 1103 | next | |
| 1104 | assume neq: "fst (prod_decode m) \<noteq> fst (prod_decode n)" | |
| 1105 | have "split f (prod_decode m) \<subseteq> A (fst (prod_decode m))" | |
| 1106 | "split f (prod_decode n) \<subseteq> A (fst (prod_decode n))" | |
| 1107 | using f[THEN spec, of "fst (prod_decode m)"] | |
| 1108 | using f[THEN spec, of "fst (prod_decode n)"] | |
| 1109 | by (auto simp: set_eq_iff) | |
| 1110 | with f A neq show ?thesis | |
| 1111 | by (fastforce simp: disjoint_family_on_def subset_eq set_eq_iff) | |
| 1112 | qed | |
| 1113 | qed | |
| 1114 | from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (split f (prod_decode n)))" | |
| 1115 | by (intro suminf_ereal_2dimen[symmetric] positiveD2[OF pos] generated_ringI_Basic) | |
| 1116 | (auto split: prod.split) | |
| 1117 | also have "\<dots> = (\<Sum>n. \<mu> (split f (prod_decode n)))" | |
| 1118 | using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split) | |
| 1119 | also have "\<dots> = \<mu> (\<Union>i. split f (prod_decode i))" | |
| 1120 | using f `c \<in> C'` C' | |
| 1121 | by (intro ca[unfolded countably_additive_def, rule_format]) | |
| 1122 | (auto split: prod.split simp: UN_f_eq d UN_eq) | |
| 1123 | finally have "(\<Sum>n. \<mu>_r (A' n \<inter> c)) = \<mu> c" | |
| 1124 | using UN_f_eq UN_eq by (simp add: A_def) } | |
| 1125 | note eq = this | |
| 1126 | ||
| 1127 | have "(\<Sum>n. \<mu>_r (A' n)) = (\<Sum>n. \<Sum>c\<in>C'. \<mu>_r (A' n \<inter> c))" | |
| 49394 
52e636ace94e
removing find_theorems commands that were left in the developments accidently
 bulwahn parents: 
47762diff
changeset | 1128 | using C' A' | 
| 47762 | 1129 | by (subst volume_finite_additive[symmetric, OF V(1)]) | 
| 1130 | (auto simp: disjoint_def disjoint_family_on_def Union_image_eq[symmetric] simp del: Union_image_eq | |
| 1131 | intro!: G.Int G.finite_Union arg_cong[where f="\<lambda>X. suminf (\<lambda>i. \<mu>_r (X i))"] ext | |
| 1132 | intro: generated_ringI_Basic) | |
| 1133 | also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))" | |
| 1134 | using C' A' | |
| 1135 | by (intro suminf_setsum_ereal positiveD2[OF pos] G.Int G.finite_Union) | |
| 1136 | (auto intro: generated_ringI_Basic) | |
| 1137 | also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)" | |
| 1138 | using eq V C' by (auto intro!: setsum_cong) | |
| 1139 | also have "\<dots> = \<mu>_r (\<Union>C')" | |
| 1140 | using C' Un_A | |
| 1141 | by (subst volume_finite_additive[symmetric, OF V(1)]) | |
| 1142 | (auto simp: disjoint_family_on_def disjoint_def Union_image_eq[symmetric] simp del: Union_image_eq | |
| 1143 | intro: generated_ringI_Basic) | |
| 1144 | finally show "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)" | |
| 1145 | using C' by simp | |
| 1146 | qed | |
| 1147 | from G.caratheodory'[OF `positive generated_ring \<mu>_r` `countably_additive generated_ring \<mu>_r`] | |
| 1148 | guess \<mu>' .. | |
| 1149 | with V show ?thesis | |
| 1150 | unfolding sigma_sets_generated_ring_eq | |
| 1151 | by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic) | |
| 1152 | qed | |
| 1153 | ||
| 33271 
7be66dee1a5a
New theory Probability, which contains a development of measure theory
 paulson parents: diff
changeset | 1154 | end |